author | haftmann |
Wed, 25 Nov 2009 11:16:57 +0100 | |
changeset 33959 | 2afc55e8ed27 |
parent 30738 | 0842e906300c |
child 37757 | dc78d2d9e90a |
permissions | -rw-r--r-- |
22803 | 1 |
(* Title: HOL/Library/While_Combinator.thy |
10251 | 2 |
Author: Tobias Nipkow |
3 |
Copyright 2000 TU Muenchen |
|
4 |
*) |
|
5 |
||
14706 | 6 |
header {* A general ``while'' combinator *} |
10251 | 7 |
|
15131 | 8 |
theory While_Combinator |
30738 | 9 |
imports Main |
15131 | 10 |
begin |
10251 | 11 |
|
23821 | 12 |
text {* |
13 |
We define the while combinator as the "mother of all tail recursive functions". |
|
10251 | 14 |
*} |
15 |
||
23821 | 16 |
function (tailrec) while :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" |
17 |
where |
|
18 |
while_unfold[simp del]: "while b c s = (if b s then while b c (c s) else s)" |
|
19 |
by auto |
|
10251 | 20 |
|
23821 | 21 |
declare while_unfold[code] |
10984 | 22 |
|
18372 | 23 |
lemma def_while_unfold: |
24 |
assumes fdef: "f == while test do" |
|
25 |
shows "f x = (if test x then f(do x) else x)" |
|
14300 | 26 |
proof - |
27 |
have "f x = while test do x" using fdef by simp |
|
28 |
also have "\<dots> = (if test x then while test do (do x) else x)" |
|
29 |
by(rule while_unfold) |
|
30 |
also have "\<dots> = (if test x then f(do x) else x)" by(simp add:fdef[symmetric]) |
|
31 |
finally show ?thesis . |
|
32 |
qed |
|
33 |
||
34 |
||
10251 | 35 |
text {* |
36 |
The proof rule for @{term while}, where @{term P} is the invariant. |
|
37 |
*} |
|
38 |
||
18372 | 39 |
theorem while_rule_lemma: |
40 |
assumes invariant: "!!s. P s ==> b s ==> P (c s)" |
|
41 |
and terminate: "!!s. P s ==> \<not> b s ==> Q s" |
|
42 |
and wf: "wf {(t, s). P s \<and> b s \<and> t = c s}" |
|
43 |
shows "P s \<Longrightarrow> Q (while b c s)" |
|
19736 | 44 |
using wf |
45 |
apply (induct s) |
|
18372 | 46 |
apply simp |
47 |
apply (subst while_unfold) |
|
48 |
apply (simp add: invariant terminate) |
|
49 |
done |
|
10251 | 50 |
|
10653 | 51 |
theorem while_rule: |
10984 | 52 |
"[| P s; |
53 |
!!s. [| P s; b s |] ==> P (c s); |
|
54 |
!!s. [| P s; \<not> b s |] ==> Q s; |
|
10997 | 55 |
wf r; |
10984 | 56 |
!!s. [| P s; b s |] ==> (c s, s) \<in> r |] ==> |
57 |
Q (while b c s)" |
|
19736 | 58 |
apply (rule while_rule_lemma) |
59 |
prefer 4 apply assumption |
|
60 |
apply blast |
|
61 |
apply blast |
|
62 |
apply (erule wf_subset) |
|
63 |
apply blast |
|
64 |
done |
|
10653 | 65 |
|
10984 | 66 |
text {* |
67 |
\medskip An application: computation of the @{term lfp} on finite |
|
68 |
sets via iteration. |
|
69 |
*} |
|
70 |
||
71 |
theorem lfp_conv_while: |
|
72 |
"[| mono f; finite U; f U = U |] ==> |
|
73 |
lfp f = fst (while (\<lambda>(A, fA). A \<noteq> fA) (\<lambda>(A, fA). (fA, f fA)) ({}, f {}))" |
|
74 |
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 | 75 |
r = "((Pow U \<times> UNIV) \<times> (Pow U \<times> UNIV)) \<inter> |
10984 | 76 |
inv_image finite_psubset (op - U o fst)" in while_rule) |
77 |
apply (subst lfp_unfold) |
|
78 |
apply assumption |
|
79 |
apply (simp add: monoD) |
|
80 |
apply (subst lfp_unfold) |
|
81 |
apply assumption |
|
82 |
apply clarsimp |
|
83 |
apply (blast dest: monoD) |
|
84 |
apply (fastsimp intro!: lfp_lowerbound) |
|
85 |
apply (blast intro: wf_finite_psubset Int_lower2 [THEN [2] wf_subset]) |
|
19769
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents:
19736
diff
changeset
|
86 |
apply (clarsimp simp add: finite_psubset_def order_less_le) |
10984 | 87 |
apply (blast intro!: finite_Diff dest: monoD) |
88 |
done |
|
89 |
||
90 |
||
91 |
text {* |
|
14589 | 92 |
An example of using the @{term while} combinator. |
10984 | 93 |
*} |
94 |
||
15197 | 95 |
text{* Cannot use @{thm[source]set_eq_subset} because it leads to |
96 |
looping because the antisymmetry simproc turns the subset relationship |
|
97 |
back into equality. *} |
|
98 |
||
14589 | 99 |
theorem "P (lfp (\<lambda>N::int set. {0} \<union> {(n + 2) mod 6 | n. n \<in> N})) = |
100 |
P {0, 4, 2}" |
|
10997 | 101 |
proof - |
19736 | 102 |
have seteq: "!!A B. (A = B) = ((!a : A. a:B) & (!b:B. b:A))" |
103 |
by blast |
|
10997 | 104 |
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 | 105 |
apply blast |
10997 | 106 |
done |
107 |
show ?thesis |
|
11914 | 108 |
apply (subst lfp_conv_while [where ?U = "{0, 1, 2, 3, 4, 5}"]) |
10997 | 109 |
apply (rule monoI) |
110 |
apply blast |
|
111 |
apply simp |
|
112 |
apply (simp add: aux set_eq_subset) |
|
113 |
txt {* The fixpoint computation is performed purely by rewriting: *} |
|
15197 | 114 |
apply (simp add: while_unfold aux seteq del: subset_empty) |
10997 | 115 |
done |
116 |
qed |
|
10251 | 117 |
|
118 |
end |