author | wenzelm |
Sat, 27 May 2006 17:42:02 +0200 | |
changeset 19736 | d8d0f8f51d69 |
parent 18372 | 2bffdf62fe7f |
child 19769 | c40ce2de2020 |
permissions | -rw-r--r-- |
10251 | 1 |
(* Title: HOL/Library/While.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 2000 TU Muenchen |
|
5 |
*) |
|
6 |
||
14706 | 7 |
header {* A general ``while'' combinator *} |
10251 | 8 |
|
15131 | 9 |
theory While_Combinator |
15140 | 10 |
imports Main |
15131 | 11 |
begin |
10251 | 12 |
|
13 |
text {* |
|
14 |
We define a while-combinator @{term while} and prove: (a) an |
|
15 |
unrestricted unfolding law (even if while diverges!) (I got this |
|
16 |
idea from Wolfgang Goerigk), and (b) the invariant rule for reasoning |
|
17 |
about @{term while}. |
|
18 |
*} |
|
19 |
||
20 |
consts while_aux :: "('a => bool) \<times> ('a => 'a) \<times> 'a => 'a" |
|
11626 | 21 |
recdef (permissive) while_aux |
10251 | 22 |
"same_fst (\<lambda>b. True) (\<lambda>b. same_fst (\<lambda>c. True) (\<lambda>c. |
23 |
{(t, s). b s \<and> c s = t \<and> |
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11626
diff
changeset
|
24 |
\<not> (\<exists>f. f (0::nat) = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1)))}))" |
10251 | 25 |
"while_aux (b, c, s) = |
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11626
diff
changeset
|
26 |
(if (\<exists>f. f (0::nat) = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1))) |
10251 | 27 |
then arbitrary |
28 |
else if b s then while_aux (b, c, c s) |
|
29 |
else s)" |
|
30 |
||
10774 | 31 |
recdef_tc while_aux_tc: while_aux |
32 |
apply (rule wf_same_fst) |
|
33 |
apply (rule wf_same_fst) |
|
34 |
apply (simp add: wf_iff_no_infinite_down_chain) |
|
35 |
apply blast |
|
36 |
done |
|
37 |
||
19736 | 38 |
definition |
10251 | 39 |
while :: "('a => bool) => ('a => 'a) => 'a => 'a" |
40 |
"while b c s == while_aux (b, c, s)" |
|
41 |
||
42 |
lemma while_aux_unfold: |
|
43 |
"while_aux (b, c, s) = |
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11626
diff
changeset
|
44 |
(if \<exists>f. f (0::nat) = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1)) |
10251 | 45 |
then arbitrary |
46 |
else if b s then while_aux (b, c, c s) |
|
47 |
else s)" |
|
48 |
apply (rule while_aux_tc [THEN while_aux.simps [THEN trans]]) |
|
49 |
apply (rule refl) |
|
50 |
done |
|
51 |
||
52 |
text {* |
|
53 |
The recursion equation for @{term while}: directly executable! |
|
54 |
*} |
|
55 |
||
12791
ccc0f45ad2c4
registered directly executable version with the code generator
kleing
parents:
11914
diff
changeset
|
56 |
theorem while_unfold [code]: |
10251 | 57 |
"while b c s = (if b s then while b c (c s) else s)" |
58 |
apply (unfold while_def) |
|
59 |
apply (rule while_aux_unfold [THEN trans]) |
|
60 |
apply auto |
|
61 |
apply (subst while_aux_unfold) |
|
62 |
apply simp |
|
63 |
apply clarify |
|
64 |
apply (erule_tac x = "\<lambda>i. f (Suc i)" in allE) |
|
65 |
apply blast |
|
66 |
done |
|
67 |
||
10984 | 68 |
hide const while_aux |
69 |
||
18372 | 70 |
lemma def_while_unfold: |
71 |
assumes fdef: "f == while test do" |
|
72 |
shows "f x = (if test x then f(do x) else x)" |
|
14300 | 73 |
proof - |
74 |
have "f x = while test do x" using fdef by simp |
|
75 |
also have "\<dots> = (if test x then while test do (do x) else x)" |
|
76 |
by(rule while_unfold) |
|
77 |
also have "\<dots> = (if test x then f(do x) else x)" by(simp add:fdef[symmetric]) |
|
78 |
finally show ?thesis . |
|
79 |
qed |
|
80 |
||
81 |
||
10251 | 82 |
text {* |
83 |
The proof rule for @{term while}, where @{term P} is the invariant. |
|
84 |
*} |
|
85 |
||
18372 | 86 |
theorem while_rule_lemma: |
87 |
assumes invariant: "!!s. P s ==> b s ==> P (c s)" |
|
88 |
and terminate: "!!s. P s ==> \<not> b s ==> Q s" |
|
89 |
and wf: "wf {(t, s). P s \<and> b s \<and> t = c s}" |
|
90 |
shows "P s \<Longrightarrow> Q (while b c s)" |
|
19736 | 91 |
using wf |
92 |
apply (induct s) |
|
18372 | 93 |
apply simp |
94 |
apply (subst while_unfold) |
|
95 |
apply (simp add: invariant terminate) |
|
96 |
done |
|
10251 | 97 |
|
10653 | 98 |
theorem while_rule: |
10984 | 99 |
"[| P s; |
100 |
!!s. [| P s; b s |] ==> P (c s); |
|
101 |
!!s. [| P s; \<not> b s |] ==> Q s; |
|
10997 | 102 |
wf r; |
10984 | 103 |
!!s. [| P s; b s |] ==> (c s, s) \<in> r |] ==> |
104 |
Q (while b c s)" |
|
19736 | 105 |
apply (rule while_rule_lemma) |
106 |
prefer 4 apply assumption |
|
107 |
apply blast |
|
108 |
apply blast |
|
109 |
apply (erule wf_subset) |
|
110 |
apply blast |
|
111 |
done |
|
10653 | 112 |
|
10984 | 113 |
text {* |
114 |
\medskip An application: computation of the @{term lfp} on finite |
|
115 |
sets via iteration. |
|
116 |
*} |
|
117 |
||
118 |
theorem lfp_conv_while: |
|
119 |
"[| mono f; finite U; f U = U |] ==> |
|
120 |
lfp f = fst (while (\<lambda>(A, fA). A \<noteq> fA) (\<lambda>(A, fA). (fA, f fA)) ({}, f {}))" |
|
121 |
apply (rule_tac P = "\<lambda>(A, B). (A \<subseteq> U \<and> B = f A \<and> A \<subseteq> B \<and> B \<subseteq> lfp f)" and |
|
11047 | 122 |
r = "((Pow U \<times> UNIV) \<times> (Pow U \<times> UNIV)) \<inter> |
10984 | 123 |
inv_image finite_psubset (op - U o fst)" in while_rule) |
124 |
apply (subst lfp_unfold) |
|
125 |
apply assumption |
|
126 |
apply (simp add: monoD) |
|
127 |
apply (subst lfp_unfold) |
|
128 |
apply assumption |
|
129 |
apply clarsimp |
|
130 |
apply (blast dest: monoD) |
|
131 |
apply (fastsimp intro!: lfp_lowerbound) |
|
132 |
apply (blast intro: wf_finite_psubset Int_lower2 [THEN [2] wf_subset]) |
|
133 |
apply (clarsimp simp add: inv_image_def finite_psubset_def order_less_le) |
|
134 |
apply (blast intro!: finite_Diff dest: monoD) |
|
135 |
done |
|
136 |
||
137 |
||
138 |
text {* |
|
14589 | 139 |
An example of using the @{term while} combinator. |
10984 | 140 |
*} |
141 |
||
15197 | 142 |
text{* Cannot use @{thm[source]set_eq_subset} because it leads to |
143 |
looping because the antisymmetry simproc turns the subset relationship |
|
144 |
back into equality. *} |
|
145 |
||
14589 | 146 |
theorem "P (lfp (\<lambda>N::int set. {0} \<union> {(n + 2) mod 6 | n. n \<in> N})) = |
147 |
P {0, 4, 2}" |
|
10997 | 148 |
proof - |
19736 | 149 |
have seteq: "!!A B. (A = B) = ((!a : A. a:B) & (!b:B. b:A))" |
150 |
by blast |
|
10997 | 151 |
have aux: "!!f A B. {f n | n. A n \<or> B n} = {f n | n. A n} \<union> {f n | n. B n}" |
10984 | 152 |
apply blast |
10997 | 153 |
done |
154 |
show ?thesis |
|
11914 | 155 |
apply (subst lfp_conv_while [where ?U = "{0, 1, 2, 3, 4, 5}"]) |
10997 | 156 |
apply (rule monoI) |
157 |
apply blast |
|
158 |
apply simp |
|
159 |
apply (simp add: aux set_eq_subset) |
|
160 |
txt {* The fixpoint computation is performed purely by rewriting: *} |
|
15197 | 161 |
apply (simp add: while_unfold aux seteq del: subset_empty) |
10997 | 162 |
done |
163 |
qed |
|
10251 | 164 |
|
165 |
end |