| author | wenzelm | 
| Thu, 06 Jun 2024 22:03:20 +0200 | |
| changeset 80271 | 198fc882ec0f | 
| parent 74974 | 7733c794cfea | 
| child 81347 | 31f9e5ada550 | 
| 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: 
30738diff
changeset | 3 | Author: Alexander Krauss | 
| 10251 | 4 | *) | 
| 5 | ||
| 60500 | 6 | section \<open>A general ``while'' combinator\<close> | 
| 10251 | 7 | |
| 15131 | 8 | theory While_Combinator | 
| 30738 | 9 | imports Main | 
| 15131 | 10 | begin | 
| 10251 | 11 | |
| 60500 | 12 | subsection \<open>Partial version\<close> | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 13 | |
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 14 | definition while_option :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where
 | 
| 67091 | 15 | "while_option b c s = (if (\<exists>k. \<not> b ((c ^^ k) s)) | 
| 16 | then Some ((c ^^ (LEAST k. \<not> b ((c ^^ k) s))) s) | |
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 17 | else None)" | 
| 10251 | 18 | |
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 19 | theorem while_option_unfold[code]: | 
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 20 | "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: 
30738diff
changeset | 21 | proof cases | 
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 22 | assume "b s" | 
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 23 | show ?thesis | 
| 67091 | 24 | proof (cases "\<exists>k. \<not> b ((c ^^ k) s)") | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 25 | case True | 
| 67091 | 26 | then obtain k where 1: "\<not> b ((c ^^ k) s)" .. | 
| 60500 | 27 | with \<open>b s\<close> obtain l where "k = Suc l" by (cases k) auto | 
| 67091 | 28 | with 1 have "\<not> b ((c ^^ l) (c s))" by (auto simp: funpow_swap1) | 
| 29 | then have 2: "\<exists>l. \<not> b ((c ^^ l) (c s))" .. | |
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 30 | from 1 | 
| 67091 | 31 | have "(LEAST k. \<not> b ((c ^^ k) s)) = Suc (LEAST l. \<not> b ((c ^^ Suc l) s))" | 
| 60500 | 32 | by (rule Least_Suc) (simp add: \<open>b s\<close>) | 
| 67091 | 33 | also have "... = Suc (LEAST l. \<not> b ((c ^^ l) (c s)))" | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 34 | by (simp add: funpow_swap1) | 
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 35 | finally | 
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 36 | show ?thesis | 
| 60500 | 37 | using True 2 \<open>b s\<close> by (simp add: funpow_swap1 while_option_def) | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 38 | next | 
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 39 | case False | 
| 67091 | 40 | then have "\<not> (\<exists>l. \<not> b ((c ^^ Suc l) s))" by blast | 
| 41 | then have "\<not> (\<exists>l. \<not> b ((c ^^ l) (c s)))" | |
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 42 | by (simp add: funpow_swap1) | 
| 60500 | 43 | with False \<open>b s\<close> show ?thesis by (simp add: while_option_def) | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 44 | qed | 
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 45 | next | 
| 67091 | 46 | assume [simp]: "\<not> b s" | 
| 47 | have least: "(LEAST k. \<not> b ((c ^^ k) s)) = 0" | |
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 48 | by (rule Least_equality) auto | 
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 49 | moreover | 
| 67091 | 50 | have "\<exists>k. \<not> b ((c ^^ k) s)" by (rule exI[of _ "0::nat"]) auto | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 51 | 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: 
30738diff
changeset | 52 | qed | 
| 10251 | 53 | |
| 45834 | 54 | lemma while_option_stop2: | 
| 67091 | 55 | "while_option b c s = Some t \<Longrightarrow> \<exists>k. t = (c^^k) s \<and> \<not> b t" | 
| 45834 | 56 | apply(simp add: while_option_def split: if_splits) | 
| 46365 | 57 | by (metis (lifting) LeastI_ex) | 
| 45834 | 58 | |
| 67091 | 59 | lemma while_option_stop: "while_option b c s = Some t \<Longrightarrow> \<not> b t" | 
| 45834 | 60 | by(metis while_option_stop2) | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 61 | |
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 62 | theorem while_option_rule: | 
| 63040 | 63 | assumes step: "!!s. P s ==> b s ==> P (c s)" | 
| 64 | and result: "while_option b c s = Some t" | |
| 65 | and init: "P s" | |
| 66 | shows "P t" | |
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 67 | proof - | 
| 67091 | 68 | define k where "k = (LEAST k. \<not> b ((c ^^ k) s))" | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 69 | 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: 
30738diff
changeset | 70 | by (simp add: while_option_def k_def split: if_splits) | 
| 67717 | 71 | have 1: "\<forall>i<k. b ((c ^^ i) s)" | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 72 | 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: 
30738diff
changeset | 73 | |
| 67091 | 74 |   { fix i assume "i \<le> k" then have "P ((c ^^ i) s)"
 | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 75 | 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: 
30738diff
changeset | 76 | 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: 
30738diff
changeset | 77 | qed | 
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 78 | |
| 50577 | 79 | lemma funpow_commute: | 
| 80 | "\<lbrakk>\<forall>k' < k. f (c ((c^^k') s)) = c' (f ((c^^k') s))\<rbrakk> \<Longrightarrow> f ((c^^k) s) = (c'^^k) (f s)" | |
| 81 | by (induct k arbitrary: s) auto | |
| 82 | ||
| 54050 | 83 | lemma while_option_commute_invariant: | 
| 84 | assumes Invariant: "\<And>s. P s \<Longrightarrow> b s \<Longrightarrow> P (c s)" | |
| 85 | assumes TestCommute: "\<And>s. P s \<Longrightarrow> b s = b' (f s)" | |
| 86 | assumes BodyCommute: "\<And>s. P s \<Longrightarrow> b s \<Longrightarrow> f (c s) = c' (f s)" | |
| 87 | assumes Initial: "P s" | |
| 55466 | 88 | shows "map_option f (while_option b c s) = while_option b' c' (f s)" | 
| 50577 | 89 | unfolding while_option_def | 
| 90 | proof (rule trans[OF if_distrib if_cong], safe, unfold option.inject) | |
| 54050 | 91 | fix k | 
| 92 | assume "\<not> b ((c ^^ k) s)" | |
| 93 | with Initial show "\<exists>k. \<not> b' ((c' ^^ k) (f s))" | |
| 50577 | 94 | proof (induction k arbitrary: s) | 
| 54050 | 95 | case 0 thus ?case by (auto simp: TestCommute intro: exI[of _ 0]) | 
| 50577 | 96 | next | 
| 54050 | 97 | case (Suc k) thus ?case | 
| 98 | proof (cases "b s") | |
| 99 | assume "b s" | |
| 100 | with Suc.IH[of "c s"] Suc.prems show ?thesis | |
| 101 | by (metis BodyCommute Invariant comp_apply funpow.simps(2) funpow_swap1) | |
| 102 | next | |
| 103 | assume "\<not> b s" | |
| 104 | with Suc show ?thesis by (auto simp: TestCommute intro: exI [of _ 0]) | |
| 105 | qed | |
| 50577 | 106 | qed | 
| 107 | next | |
| 54050 | 108 | fix k | 
| 109 | assume "\<not> b' ((c' ^^ k) (f s))" | |
| 110 | with Initial show "\<exists>k. \<not> b ((c ^^ k) s)" | |
| 50577 | 111 | proof (induction k arbitrary: s) | 
| 54050 | 112 | case 0 thus ?case by (auto simp: TestCommute intro: exI[of _ 0]) | 
| 50577 | 113 | next | 
| 54050 | 114 | case (Suc k) thus ?case | 
| 50577 | 115 | proof (cases "b s") | 
| 54050 | 116 | assume "b s" | 
| 117 | with Suc.IH[of "c s"] Suc.prems show ?thesis | |
| 118 | by (metis BodyCommute Invariant comp_apply funpow.simps(2) funpow_swap1) | |
| 119 | next | |
| 120 | assume "\<not> b s" | |
| 121 | with Suc show ?thesis by (auto simp: TestCommute intro: exI [of _ 0]) | |
| 122 | qed | |
| 50577 | 123 | qed | 
| 124 | next | |
| 54050 | 125 | fix k | 
| 126 | assume k: "\<not> b' ((c' ^^ k) (f s))" | |
| 127 | have *: "(LEAST k. \<not> b' ((c' ^^ k) (f s))) = (LEAST k. \<not> b ((c ^^ k) s))" | |
| 128 | (is "?k' = ?k") | |
| 50577 | 129 | proof (cases ?k') | 
| 130 | case 0 | |
| 54050 | 131 | have "\<not> b' ((c' ^^ 0) (f s))" | 
| 132 | unfolding 0[symmetric] by (rule LeastI[of _ k]) (rule k) | |
| 133 | hence "\<not> b s" by (auto simp: TestCommute Initial) | |
| 50577 | 134 | hence "?k = 0" by (intro Least_equality) auto | 
| 135 | with 0 show ?thesis by auto | |
| 136 | next | |
| 137 | case (Suc k') | |
| 54050 | 138 | have "\<not> b' ((c' ^^ Suc k') (f s))" | 
| 139 | unfolding Suc[symmetric] by (rule LeastI) (rule k) | |
| 50577 | 140 | moreover | 
| 141 |     { fix k assume "k \<le> k'"
 | |
| 142 | hence "k < ?k'" unfolding Suc by simp | |
| 143 | hence "b' ((c' ^^ k) (f s))" by (rule iffD1[OF not_not, OF not_less_Least]) | |
| 54050 | 144 | } | 
| 145 | note b' = this | |
| 50577 | 146 |     { fix k assume "k \<le> k'"
 | 
| 54050 | 147 | hence "f ((c ^^ k) s) = (c' ^^ k) (f s)" | 
| 148 | and "b ((c ^^ k) s) = b' ((c' ^^ k) (f s))" | |
| 149 | and "P ((c ^^ k) s)" | |
| 150 | by (induct k) (auto simp: b' assms) | |
| 60500 | 151 | with \<open>k \<le> k'\<close> | 
| 54050 | 152 | have "b ((c ^^ k) s)" | 
| 153 | and "f ((c ^^ k) s) = (c' ^^ k) (f s)" | |
| 154 | and "P ((c ^^ k) s)" | |
| 155 | by (auto simp: b') | |
| 156 | } | |
| 157 | note b = this(1) and body = this(2) and inv = this(3) | |
| 158 | hence k': "f ((c ^^ k') s) = (c' ^^ k') (f s)" by auto | |
| 50577 | 159 | ultimately show ?thesis unfolding Suc using b | 
| 61166 
5976fe402824
renamed method "goals" to "goal_cases" to emphasize its meaning;
 wenzelm parents: 
61115diff
changeset | 160 | proof (intro Least_equality[symmetric], goal_cases) | 
| 60580 | 161 | case 1 | 
| 54050 | 162 | hence Test: "\<not> b' (f ((c ^^ Suc k') s))" | 
| 163 | by (auto simp: BodyCommute inv b) | |
| 164 | have "P ((c ^^ Suc k') s)" by (auto simp: Invariant inv b) | |
| 165 | with Test show ?case by (auto simp: TestCommute) | |
| 166 | next | |
| 60580 | 167 | case 2 | 
| 168 | thus ?case by (metis not_less_eq_eq) | |
| 54050 | 169 | qed | 
| 50577 | 170 | qed | 
| 171 | have "f ((c ^^ ?k) s) = (c' ^^ ?k') (f s)" unfolding * | |
| 54050 | 172 | proof (rule funpow_commute, clarify) | 
| 173 | fix k assume "k < ?k" | |
| 174 | hence TestTrue: "b ((c ^^ k) s)" by (auto dest: not_less_Least) | |
| 60500 | 175 | from \<open>k < ?k\<close> have "P ((c ^^ k) s)" | 
| 54050 | 176 | proof (induct k) | 
| 177 | case 0 thus ?case by (auto simp: assms) | |
| 178 | next | |
| 179 | case (Suc h) | |
| 180 | hence "P ((c ^^ h) s)" by auto | |
| 181 | with Suc show ?case | |
| 182 | by (auto, metis (lifting, no_types) Invariant Suc_lessD not_less_Least) | |
| 183 | qed | |
| 184 | with TestTrue show "f (c ((c ^^ k) s)) = c' (f ((c ^^ k) s))" | |
| 185 | by (metis BodyCommute) | |
| 186 | qed | |
| 50577 | 187 | thus "\<exists>z. (c ^^ ?k) s = z \<and> f z = (c' ^^ ?k') (f s)" by blast | 
| 188 | qed | |
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 189 | |
| 54050 | 190 | lemma while_option_commute: | 
| 191 | assumes "\<And>s. b s = b' (f s)" "\<And>s. \<lbrakk>b s\<rbrakk> \<Longrightarrow> f (c s) = c' (f s)" | |
| 55466 | 192 | shows "map_option f (while_option b c s) = while_option b' c' (f s)" | 
| 54050 | 193 | by(rule while_option_commute_invariant[where P = "\<lambda>_. True"]) | 
| 194 | (auto simp add: assms) | |
| 195 | ||
| 60500 | 196 | subsection \<open>Total version\<close> | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 197 | |
| 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 198 | 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: 
30738diff
changeset | 199 | 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: 
30738diff
changeset | 200 | |
| 50008 | 201 | lemma while_unfold [code]: | 
| 37757 
dc78d2d9e90a
added "while_option", which needs no well-foundedness; defined "while" in terms of "while_option"
 krauss parents: 
30738diff
changeset | 202 | "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: 
30738diff
changeset | 203 | unfolding while_def by (subst while_option_unfold) simp | 
| 10984 | 204 | |
| 18372 | 205 | lemma def_while_unfold: | 
| 206 | assumes fdef: "f == while test do" | |
| 207 | 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: 
30738diff
changeset | 208 | unfolding fdef by (fact while_unfold) | 
| 14300 | 209 | |
| 210 | ||
| 60500 | 211 | text \<open> | 
| 69593 | 212 | The proof rule for \<^term>\<open>while\<close>, where \<^term>\<open>P\<close> is the invariant. | 
| 60500 | 213 | \<close> | 
| 10251 | 214 | |
| 18372 | 215 | theorem while_rule_lemma: | 
| 216 | assumes invariant: "!!s. P s ==> b s ==> P (c s)" | |
| 217 | and terminate: "!!s. P s ==> \<not> b s ==> Q s" | |
| 218 |     and wf: "wf {(t, s). P s \<and> b s \<and> t = c s}"
 | |
| 219 | shows "P s \<Longrightarrow> Q (while b c s)" | |
| 19736 | 220 | using wf | 
| 221 | apply (induct s) | |
| 18372 | 222 | apply simp | 
| 223 | apply (subst while_unfold) | |
| 224 | apply (simp add: invariant terminate) | |
| 225 | done | |
| 10251 | 226 | |
| 10653 | 227 | theorem while_rule: | 
| 10984 | 228 | "[| P s; | 
| 229 | !!s. [| P s; b s |] ==> P (c s); | |
| 230 | !!s. [| P s; \<not> b s |] ==> Q s; | |
| 10997 | 231 | wf r; | 
| 10984 | 232 | !!s. [| P s; b s |] ==> (c s, s) \<in> r |] ==> | 
| 233 | Q (while b c s)" | |
| 19736 | 234 | apply (rule while_rule_lemma) | 
| 235 | prefer 4 apply assumption | |
| 236 | apply blast | |
| 237 | apply blast | |
| 238 | apply (erule wf_subset) | |
| 239 | apply blast | |
| 240 | done | |
| 10653 | 241 | |
| 74974 | 242 | text \<open>Combine invariant preservation and variant decrease in one goal:\<close> | 
| 243 | theorem while_rule2: | |
| 244 | "[| P s; | |
| 245 | !!s. [| P s; b s |] ==> P (c s) \<and> (c s, s) \<in> r; | |
| 246 | !!s. [| P s; \<not> b s |] ==> Q s; | |
| 247 | wf r |] ==> | |
| 248 | Q (while b c s)" | |
| 249 | using while_rule[of P] by metis | |
| 250 | ||
| 60500 | 251 | text\<open>Proving termination:\<close> | 
| 41720 | 252 | |
| 253 | theorem wf_while_option_Some: | |
| 41764 | 254 |   assumes "wf {(t, s). (P s \<and> b s) \<and> t = c s}"
 | 
| 67091 | 255 | and "\<And>s. P s \<Longrightarrow> b s \<Longrightarrow> P(c s)" and "P s" | 
| 256 | shows "\<exists>t. while_option b c s = Some t" | |
| 41764 | 257 | using assms(1,3) | 
| 54050 | 258 | proof (induction s) | 
| 259 | case less thus ?case using assms(2) | |
| 260 | by (subst while_option_unfold) simp | |
| 261 | qed | |
| 262 | ||
| 263 | lemma wf_rel_while_option_Some: | |
| 264 | assumes wf: "wf R" | |
| 265 | assumes smaller: "\<And>s. P s \<and> b s \<Longrightarrow> (c s, s) \<in> R" | |
| 266 | assumes inv: "\<And>s. P s \<and> b s \<Longrightarrow> P(c s)" | |
| 267 | assumes init: "P s" | |
| 268 | shows "\<exists>t. while_option b c s = Some t" | |
| 269 | proof - | |
| 270 |  from smaller have "{(t,s). P s \<and> b s \<and> t = c s} \<subseteq> R" by auto
 | |
| 271 |  with wf have "wf {(t,s). P s \<and> b s \<and> t = c s}" by (auto simp: wf_subset)
 | |
| 272 | with inv init show ?thesis by (auto simp: wf_while_option_Some) | |
| 273 | qed | |
| 41720 | 274 | |
| 275 | theorem measure_while_option_Some: fixes f :: "'s \<Rightarrow> nat" | |
| 67091 | 276 | shows "(\<And>s. P s \<Longrightarrow> b s \<Longrightarrow> P(c s) \<and> f(c s) < f s) | 
| 277 | \<Longrightarrow> P s \<Longrightarrow> \<exists>t. while_option b c s = Some t" | |
| 41764 | 278 | by(blast intro: wf_while_option_Some[OF wf_if_measure, of P b f]) | 
| 10251 | 279 | |
| 60500 | 280 | text\<open>Kleene iteration starting from the empty set and assuming some finite | 
| 281 | bounding set:\<close> | |
| 45834 | 282 | |
| 283 | lemma while_option_finite_subset_Some: fixes C :: "'a set" | |
| 284 | assumes "mono f" and "!!X. X \<subseteq> C \<Longrightarrow> f X \<subseteq> C" and "finite C" | |
| 285 |   shows "\<exists>P. while_option (\<lambda>A. f A \<noteq> A) f {} = Some P"
 | |
| 286 | proof(rule measure_while_option_Some[where | |
| 287 |     f= "%A::'a set. card C - card A" and P= "%A. A \<subseteq> C \<and> A \<subseteq> f A" and s= "{}"])
 | |
| 288 | fix A assume A: "A \<subseteq> C \<and> A \<subseteq> f A" "f A \<noteq> A" | |
| 289 | show "(f A \<subseteq> C \<and> f A \<subseteq> f (f A)) \<and> card C - card (f A) < card C - card A" | |
| 290 | (is "?L \<and> ?R") | |
| 291 | proof | |
| 60500 | 292 | show ?L by(metis A(1) assms(2) monoD[OF \<open>mono f\<close>]) | 
| 45834 | 293 | show ?R by (metis A assms(2,3) card_seteq diff_less_mono2 equalityI linorder_le_less_linear rev_finite_subset) | 
| 294 | qed | |
| 295 | qed simp | |
| 296 | ||
| 297 | lemma lfp_the_while_option: | |
| 298 | assumes "mono f" and "!!X. X \<subseteq> C \<Longrightarrow> f X \<subseteq> C" and "finite C" | |
| 299 |   shows "lfp f = the(while_option (\<lambda>A. f A \<noteq> A) f {})"
 | |
| 300 | proof- | |
| 301 |   obtain P where "while_option (\<lambda>A. f A \<noteq> A) f {} = Some P"
 | |
| 302 | using while_option_finite_subset_Some[OF assms] by blast | |
| 303 | with while_option_stop2[OF this] lfp_Kleene_iter[OF assms(1)] | |
| 304 | show ?thesis by auto | |
| 305 | qed | |
| 306 | ||
| 50180 | 307 | lemma lfp_while: | 
| 308 | assumes "mono f" and "!!X. X \<subseteq> C \<Longrightarrow> f X \<subseteq> C" and "finite C" | |
| 309 |   shows "lfp f = while (\<lambda>A. f A \<noteq> A) f {}"
 | |
| 310 | unfolding while_def using assms by (rule lfp_the_while_option) blast | |
| 311 | ||
| 63561 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 312 | lemma wf_finite_less: | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 313 | assumes "finite (C :: 'a::order set)" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 314 |   shows "wf {(x, y). {x, y} \<subseteq> C \<and> x < y}"
 | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 315 | by (rule wf_measure[where f="\<lambda>b. card {a. a \<in> C \<and> a < b}", THEN wf_subset])
 | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 316 | (fastforce simp: less_eq assms intro: psubset_card_mono) | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 317 | |
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 318 | lemma wf_finite_greater: | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 319 | assumes "finite (C :: 'a::order set)" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 320 |   shows "wf {(x, y). {x, y} \<subseteq> C \<and> y < x}"
 | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 321 | by (rule wf_measure[where f="\<lambda>b. card {a. a \<in> C \<and> b < a}", THEN wf_subset])
 | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 322 | (fastforce simp: less_eq assms intro: psubset_card_mono) | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 323 | |
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 324 | lemma while_option_finite_increasing_Some: | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 325 | fixes f :: "'a::order \<Rightarrow> 'a" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 326 | assumes "mono f" and "finite (UNIV :: 'a set)" and "s \<le> f s" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 327 | shows "\<exists>P. while_option (\<lambda>A. f A \<noteq> A) f s = Some P" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 328 | by (rule wf_rel_while_option_Some[where R="{(x, y). y < x}" and P="\<lambda>A. A \<le> f A" and s="s"])
 | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 329 | (auto simp: assms monoD intro: wf_finite_greater[where C="UNIV::'a set", simplified]) | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 330 | |
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 331 | lemma lfp_the_while_option_lattice: | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 332 | fixes f :: "'a::complete_lattice \<Rightarrow> 'a" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 333 | assumes "mono f" and "finite (UNIV :: 'a set)" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 334 | shows "lfp f = the (while_option (\<lambda>A. f A \<noteq> A) f bot)" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 335 | proof - | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 336 | obtain P where "while_option (\<lambda>A. f A \<noteq> A) f bot = Some P" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 337 | using while_option_finite_increasing_Some[OF assms, where s=bot] by simp blast | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 338 | with while_option_stop2[OF this] lfp_Kleene_iter[OF assms(1)] | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 339 | show ?thesis by auto | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 340 | qed | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 341 | |
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 342 | lemma lfp_while_lattice: | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 343 | fixes f :: "'a::complete_lattice \<Rightarrow> 'a" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 344 | assumes "mono f" and "finite (UNIV :: 'a set)" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 345 | shows "lfp f = while (\<lambda>A. f A \<noteq> A) f bot" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 346 | unfolding while_def using assms by (rule lfp_the_while_option_lattice) | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 347 | |
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 348 | lemma while_option_finite_decreasing_Some: | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 349 | fixes f :: "'a::order \<Rightarrow> 'a" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 350 | assumes "mono f" and "finite (UNIV :: 'a set)" and "f s \<le> s" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 351 | shows "\<exists>P. while_option (\<lambda>A. f A \<noteq> A) f s = Some P" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 352 | by (rule wf_rel_while_option_Some[where R="{(x, y). x < y}" and P="\<lambda>A. f A \<le> A" and s="s"])
 | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 353 | (auto simp add: assms monoD intro: wf_finite_less[where C="UNIV::'a set", simplified]) | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 354 | |
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 355 | lemma gfp_the_while_option_lattice: | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 356 | fixes f :: "'a::complete_lattice \<Rightarrow> 'a" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 357 | assumes "mono f" and "finite (UNIV :: 'a set)" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 358 | shows "gfp f = the(while_option (\<lambda>A. f A \<noteq> A) f top)" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 359 | proof - | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 360 | obtain P where "while_option (\<lambda>A. f A \<noteq> A) f top = Some P" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 361 | using while_option_finite_decreasing_Some[OF assms, where s=top] by simp blast | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 362 | with while_option_stop2[OF this] gfp_Kleene_iter[OF assms(1)] | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 363 | show ?thesis by auto | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 364 | qed | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 365 | |
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 366 | lemma gfp_while_lattice: | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 367 | fixes f :: "'a::complete_lattice \<Rightarrow> 'a" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 368 | assumes "mono f" and "finite (UNIV :: 'a set)" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 369 | shows "gfp f = while (\<lambda>A. f A \<noteq> A) f top" | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63040diff
changeset | 370 | unfolding while_def using assms by (rule gfp_the_while_option_lattice) | 
| 53217 | 371 | |
| 60500 | 372 | text\<open>Computing the reflexive, transitive closure by iterating a successor | 
| 53217 | 373 | function. Stops when an element is found that dos not satisfy the test. | 
| 374 | ||
| 375 | More refined (and hence more efficient) versions can be found in ITP 2011 paper | |
| 376 | by Nipkow (the theories are in the AFP entry Flyspeck by Nipkow) | |
| 60500 | 377 | and the AFP article Executable Transitive Closures by René Thiemann.\<close> | 
| 53217 | 378 | |
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 379 | context | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 380 | fixes p :: "'a \<Rightarrow> bool" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 381 | and f :: "'a \<Rightarrow> 'a list" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 382 | and x :: 'a | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 383 | begin | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 384 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60580diff
changeset | 385 | qualified fun rtrancl_while_test :: "'a list \<times> 'a set \<Rightarrow> bool" | 
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 386 | where "rtrancl_while_test (ws,_) = (ws \<noteq> [] \<and> p(hd ws))" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 387 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60580diff
changeset | 388 | qualified fun rtrancl_while_step :: "'a list \<times> 'a set \<Rightarrow> 'a list \<times> 'a set" | 
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 389 | where "rtrancl_while_step (ws, Z) = | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 390 | (let x = hd ws; new = remdups (filter (\<lambda>y. y \<notin> Z) (f x)) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 391 | in (new @ tl ws, set new \<union> Z))" | 
| 53217 | 392 | |
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 393 | definition rtrancl_while :: "('a list * 'a set) option"
 | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 394 | where "rtrancl_while = while_option rtrancl_while_test rtrancl_while_step ([x],{x})"
 | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 395 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60580diff
changeset | 396 | qualified fun rtrancl_while_invariant :: "'a list \<times> 'a set \<Rightarrow> bool" | 
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 397 | where "rtrancl_while_invariant (ws, Z) = | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 398 |    (x \<in> Z \<and> set ws \<subseteq> Z \<and> distinct ws \<and> {(x,y). y \<in> set(f x)} `` (Z - set ws) \<subseteq> Z \<and>
 | 
| 67613 | 399 |     Z \<subseteq> {(x,y). y \<in> set(f x)}\<^sup>* `` {x} \<and> (\<forall>z\<in>Z - set ws. p z))"
 | 
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 400 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60580diff
changeset | 401 | qualified lemma rtrancl_while_invariant: | 
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 402 | assumes inv: "rtrancl_while_invariant st" and test: "rtrancl_while_test st" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 403 | shows "rtrancl_while_invariant (rtrancl_while_step st)" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 404 | proof (cases st) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 405 | fix ws Z assume st: "st = (ws, Z)" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 406 | with test obtain h t where "ws = h # t" "p h" by (cases ws) auto | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 407 | with inv st show ?thesis by (auto intro: rtrancl.rtrancl_into_rtrancl) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 408 | qed | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 409 | |
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 410 | lemma rtrancl_while_Some: assumes "rtrancl_while = Some(ws,Z)" | 
| 53217 | 411 | shows "if ws = [] | 
| 67613 | 412 |        then Z = {(x,y). y \<in> set(f x)}\<^sup>* `` {x} \<and> (\<forall>z\<in>Z. p z)
 | 
| 413 |        else \<not>p(hd ws) \<and> hd ws \<in> {(x,y). y \<in> set(f x)}\<^sup>* `` {x}"
 | |
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 414 | proof - | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 415 |   have "rtrancl_while_invariant ([x],{x})" by simp
 | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 416 | with rtrancl_while_invariant have I: "rtrancl_while_invariant (ws,Z)" | 
| 53217 | 417 | by (rule while_option_rule[OF _ assms[unfolded rtrancl_while_def]]) | 
| 418 |   { assume "ws = []"
 | |
| 419 | hence ?thesis using I | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61166diff
changeset | 420 | by (auto simp del:Image_Collect_case_prod dest: Image_closed_trancl) | 
| 53217 | 421 | } moreover | 
| 422 |   { assume "ws \<noteq> []"
 | |
| 423 | hence ?thesis using I while_option_stop[OF assms[unfolded rtrancl_while_def]] | |
| 424 | by (simp add: subset_iff) | |
| 425 | } | |
| 426 | ultimately show ?thesis by simp | |
| 427 | qed | |
| 428 | ||
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 429 | lemma rtrancl_while_finite_Some: | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 430 |   assumes "finite ({(x, y). y \<in> set (f x)}\<^sup>* `` {x})" (is "finite ?Cl")
 | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 431 | shows "\<exists>y. rtrancl_while = Some y" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 432 | proof - | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 433 |   let ?R = "(\<lambda>(_, Z). card (?Cl - Z)) <*mlex*> (\<lambda>(ws, _). length ws) <*mlex*> {}"
 | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 434 | have "wf ?R" by (blast intro: wf_mlex) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 435 | then show ?thesis unfolding rtrancl_while_def | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 436 | proof (rule wf_rel_while_option_Some[of ?R rtrancl_while_invariant]) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 437 | fix st assume *: "rtrancl_while_invariant st \<and> rtrancl_while_test st" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 438 | hence I: "rtrancl_while_invariant (rtrancl_while_step st)" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 439 | by (blast intro: rtrancl_while_invariant) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 440 | show "(rtrancl_while_step st, st) \<in> ?R" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 441 | proof (cases st) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 442 | fix ws Z let ?ws = "fst (rtrancl_while_step st)" and ?Z = "snd (rtrancl_while_step st)" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 443 | assume st: "st = (ws, Z)" | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 444 | with * obtain h t where ws: "ws = h # t" "p h" by (cases ws) auto | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 445 |       { assume "remdups (filter (\<lambda>y. y \<notin> Z) (f h)) \<noteq> []"
 | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 446 | then obtain z where "z \<in> set (remdups (filter (\<lambda>y. y \<notin> Z) (f h)))" by fastforce | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 447 | with st ws I have "Z \<subset> ?Z" "Z \<subseteq> ?Cl" "?Z \<subseteq> ?Cl" by auto | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 448 | with assms have "card (?Cl - ?Z) < card (?Cl - Z)" by (blast intro: psubset_card_mono) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 449 | with st ws have ?thesis unfolding mlex_prod_def by simp | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 450 | } | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 451 | moreover | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 452 |       { assume "remdups (filter (\<lambda>y. y \<notin> Z) (f h)) = []"
 | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 453 | with st ws have "?Z = Z" "?ws = t" by (auto simp: filter_empty_conv) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 454 | with st ws have ?thesis unfolding mlex_prod_def by simp | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 455 | } | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 456 | ultimately show ?thesis by blast | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 457 | qed | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 458 | qed (simp_all add: rtrancl_while_invariant) | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 459 | qed | 
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 460 | |
| 10251 | 461 | end | 
| 54196 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 462 | |
| 
0c188a3c671a
refactored rtrancl_while; prove termination for finite rtrancl
 traytel parents: 
54050diff
changeset | 463 | end |