| author | blanchet | 
| Wed, 21 Dec 2011 15:04:28 +0100 | |
| changeset 45945 | aa8100cc02dc | 
| parent 45834 | 9c232d370244 | 
| child 46365 | 547d1a1dcaf6 | 
| permissions | -rw-r--r-- | 
| 22803 | 1  | 
(* Title: HOL/Library/While_Combinator.thy  | 
| 10251 | 2  | 
Author: Tobias Nipkow  | 
| 
37757
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
3  | 
Author: Alexander Krauss  | 
| 10251 | 4  | 
Copyright 2000 TU Muenchen  | 
5  | 
*)  | 
|
6  | 
||
| 14706 | 7  | 
header {* A general ``while'' combinator *}
 | 
| 10251 | 8  | 
|
| 15131 | 9  | 
theory While_Combinator  | 
| 30738 | 10  | 
imports Main  | 
| 15131 | 11  | 
begin  | 
| 10251 | 12  | 
|
| 37760 | 13  | 
subsection {* Partial version *}
 | 
| 
37757
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
14  | 
|
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
15  | 
definition while_option :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where
 | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
16  | 
"while_option b c s = (if (\<exists>k. ~ b ((c ^^ k) s))  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
17  | 
then Some ((c ^^ (LEAST k. ~ b ((c ^^ k) s))) s)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
18  | 
else None)"  | 
| 10251 | 19  | 
|
| 
37757
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
20  | 
theorem while_option_unfold[code]:  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
21  | 
"while_option b c s = (if b s then while_option b c (c s) else Some s)"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
22  | 
proof cases  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
23  | 
assume "b s"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
24  | 
show ?thesis  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
25  | 
proof (cases "\<exists>k. ~ b ((c ^^ k) s)")  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
26  | 
case True  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
27  | 
then obtain k where 1: "~ b ((c ^^ k) s)" ..  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
28  | 
with `b s` obtain l where "k = Suc l" by (cases k) auto  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
29  | 
with 1 have "~ b ((c ^^ l) (c s))" by (auto simp: funpow_swap1)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
30  | 
then have 2: "\<exists>l. ~ b ((c ^^ l) (c s))" ..  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
31  | 
from 1  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
32  | 
have "(LEAST k. ~ b ((c ^^ k) s)) = Suc (LEAST l. ~ b ((c ^^ Suc l) s))"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
33  | 
by (rule Least_Suc) (simp add: `b s`)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
34  | 
also have "... = Suc (LEAST l. ~ b ((c ^^ l) (c s)))"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
35  | 
by (simp add: funpow_swap1)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
36  | 
finally  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
37  | 
show ?thesis  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
38  | 
using True 2 `b s` by (simp add: funpow_swap1 while_option_def)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
39  | 
next  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
40  | 
case False  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
41  | 
then have "~ (\<exists>l. ~ b ((c ^^ Suc l) s))" by blast  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
42  | 
then have "~ (\<exists>l. ~ b ((c ^^ l) (c s)))"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
43  | 
by (simp add: funpow_swap1)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
44  | 
with False `b s` show ?thesis by (simp add: while_option_def)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
45  | 
qed  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
46  | 
next  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
47  | 
assume [simp]: "~ b s"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
48  | 
have least: "(LEAST k. ~ b ((c ^^ k) s)) = 0"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
49  | 
by (rule Least_equality) auto  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
50  | 
moreover  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
51  | 
have "\<exists>k. ~ b ((c ^^ k) s)" by (rule exI[of _ "0::nat"]) auto  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
52  | 
ultimately show ?thesis unfolding while_option_def by auto  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
53  | 
qed  | 
| 10251 | 54  | 
|
| 45834 | 55  | 
lemma while_option_stop2:  | 
56  | 
"while_option b c s = Some t \<Longrightarrow> EX k. t = (c^^k) s \<and> \<not> b t"  | 
|
57  | 
apply(simp add: while_option_def split: if_splits)  | 
|
58  | 
by (metis (lam_lifting) LeastI_ex)  | 
|
59  | 
||
60  | 
lemma while_option_stop: "while_option b c s = Some t \<Longrightarrow> ~ b t"  | 
|
61  | 
by(metis while_option_stop2)  | 
|
| 
37757
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
62  | 
|
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
63  | 
theorem while_option_rule:  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
64  | 
assumes step: "!!s. P s ==> b s ==> P (c s)"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
65  | 
and result: "while_option b c s = Some t"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
66  | 
and init: "P s"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
67  | 
shows "P t"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
68  | 
proof -  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
69  | 
def k == "LEAST k. ~ b ((c ^^ k) s)"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
70  | 
from assms have t: "t = (c ^^ k) s"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
71  | 
by (simp add: while_option_def k_def split: if_splits)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
72  | 
have 1: "ALL i<k. b ((c ^^ i) s)"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
73  | 
by (auto simp: k_def dest: not_less_Least)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
74  | 
|
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
75  | 
  { fix i assume "i <= k" then have "P ((c ^^ i) s)"
 | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
76  | 
by (induct i) (auto simp: init step 1) }  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
77  | 
thus "P t" by (auto simp: t)  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
78  | 
qed  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
79  | 
|
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
80  | 
|
| 37760 | 81  | 
subsection {* Total version *}
 | 
| 
37757
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
82  | 
|
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
83  | 
definition while :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
84  | 
where "while b c s = the (while_option b c s)"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
85  | 
|
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
86  | 
lemma while_unfold:  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
87  | 
"while b c s = (if b s then while b c (c s) else s)"  | 
| 
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
88  | 
unfolding while_def by (subst while_option_unfold) simp  | 
| 10984 | 89  | 
|
| 18372 | 90  | 
lemma def_while_unfold:  | 
91  | 
assumes fdef: "f == while test do"  | 
|
92  | 
shows "f x = (if test x then f(do x) else x)"  | 
|
| 
37757
 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 
krauss 
parents: 
30738 
diff
changeset
 | 
93  | 
unfolding fdef by (fact while_unfold)  | 
| 14300 | 94  | 
|
95  | 
||
| 10251 | 96  | 
text {*
 | 
97  | 
 The proof rule for @{term while}, where @{term P} is the invariant.
 | 
|
98  | 
*}  | 
|
99  | 
||
| 18372 | 100  | 
theorem while_rule_lemma:  | 
101  | 
assumes invariant: "!!s. P s ==> b s ==> P (c s)"  | 
|
102  | 
and terminate: "!!s. P s ==> \<not> b s ==> Q s"  | 
|
103  | 
    and wf: "wf {(t, s). P s \<and> b s \<and> t = c s}"
 | 
|
104  | 
shows "P s \<Longrightarrow> Q (while b c s)"  | 
|
| 19736 | 105  | 
using wf  | 
106  | 
apply (induct s)  | 
|
| 18372 | 107  | 
apply simp  | 
108  | 
apply (subst while_unfold)  | 
|
109  | 
apply (simp add: invariant terminate)  | 
|
110  | 
done  | 
|
| 10251 | 111  | 
|
| 10653 | 112  | 
theorem while_rule:  | 
| 10984 | 113  | 
"[| P s;  | 
114  | 
!!s. [| P s; b s |] ==> P (c s);  | 
|
115  | 
!!s. [| P s; \<not> b s |] ==> Q s;  | 
|
| 10997 | 116  | 
wf r;  | 
| 10984 | 117  | 
!!s. [| P s; b s |] ==> (c s, s) \<in> r |] ==>  | 
118  | 
Q (while b c s)"  | 
|
| 19736 | 119  | 
apply (rule while_rule_lemma)  | 
120  | 
prefer 4 apply assumption  | 
|
121  | 
apply blast  | 
|
122  | 
apply blast  | 
|
123  | 
apply (erule wf_subset)  | 
|
124  | 
apply blast  | 
|
125  | 
done  | 
|
| 10653 | 126  | 
|
| 41720 | 127  | 
text{* Proving termination: *}
 | 
128  | 
||
129  | 
theorem wf_while_option_Some:  | 
|
| 41764 | 130  | 
  assumes "wf {(t, s). (P s \<and> b s) \<and> t = c s}"
 | 
131  | 
and "!!s. P s \<Longrightarrow> b s \<Longrightarrow> P(c s)" and "P s"  | 
|
| 41720 | 132  | 
shows "EX t. while_option b c s = Some t"  | 
| 41764 | 133  | 
using assms(1,3)  | 
| 41720 | 134  | 
apply (induct s)  | 
| 41764 | 135  | 
using assms(2)  | 
| 41720 | 136  | 
apply (subst while_option_unfold)  | 
137  | 
apply simp  | 
|
138  | 
done  | 
|
139  | 
||
140  | 
theorem measure_while_option_Some: fixes f :: "'s \<Rightarrow> nat"  | 
|
| 41764 | 141  | 
shows "(!!s. P s \<Longrightarrow> b s \<Longrightarrow> P(c s) \<and> f(c s) < f s)  | 
142  | 
\<Longrightarrow> P s \<Longrightarrow> EX t. while_option b c s = Some t"  | 
|
143  | 
by(blast intro: wf_while_option_Some[OF wf_if_measure, of P b f])  | 
|
| 10251 | 144  | 
|
| 45834 | 145  | 
text{* Kleene iteration starting from the empty set and assuming some finite
 | 
146  | 
bounding set: *}  | 
|
147  | 
||
148  | 
lemma while_option_finite_subset_Some: fixes C :: "'a set"  | 
|
149  | 
assumes "mono f" and "!!X. X \<subseteq> C \<Longrightarrow> f X \<subseteq> C" and "finite C"  | 
|
150  | 
  shows "\<exists>P. while_option (\<lambda>A. f A \<noteq> A) f {} = Some P"
 | 
|
151  | 
proof(rule measure_while_option_Some[where  | 
|
152  | 
    f= "%A::'a set. card C - card A" and P= "%A. A \<subseteq> C \<and> A \<subseteq> f A" and s= "{}"])
 | 
|
153  | 
fix A assume A: "A \<subseteq> C \<and> A \<subseteq> f A" "f A \<noteq> A"  | 
|
154  | 
show "(f A \<subseteq> C \<and> f A \<subseteq> f (f A)) \<and> card C - card (f A) < card C - card A"  | 
|
155  | 
(is "?L \<and> ?R")  | 
|
156  | 
proof  | 
|
157  | 
show ?L by(metis A(1) assms(2) monoD[OF `mono f`])  | 
|
158  | 
show ?R by (metis A assms(2,3) card_seteq diff_less_mono2 equalityI linorder_le_less_linear rev_finite_subset)  | 
|
159  | 
qed  | 
|
160  | 
qed simp  | 
|
161  | 
||
162  | 
lemma lfp_the_while_option:  | 
|
163  | 
assumes "mono f" and "!!X. X \<subseteq> C \<Longrightarrow> f X \<subseteq> C" and "finite C"  | 
|
164  | 
  shows "lfp f = the(while_option (\<lambda>A. f A \<noteq> A) f {})"
 | 
|
165  | 
proof-  | 
|
166  | 
  obtain P where "while_option (\<lambda>A. f A \<noteq> A) f {} = Some P"
 | 
|
167  | 
using while_option_finite_subset_Some[OF assms] by blast  | 
|
168  | 
with while_option_stop2[OF this] lfp_Kleene_iter[OF assms(1)]  | 
|
169  | 
show ?thesis by auto  | 
|
170  | 
qed  | 
|
171  | 
||
| 10251 | 172  | 
end  |