| author | aspinall | 
| Fri, 29 Dec 2006 16:46:39 +0100 | |
| changeset 21929 | fb0cd849bc60 | 
| parent 18372 | 2bffdf62fe7f | 
| child 22271 | 51a80e238b29 | 
| permissions | -rw-r--r-- | 
| 12516 | 1 | (* Title: HOL/MicroJava/BV/Kildall.thy | 
| 10496 | 2 | ID: $Id$ | 
| 12516 | 3 | Author: Tobias Nipkow, Gerwin Klein | 
| 10496 | 4 | Copyright 2000 TUM | 
| 5 | ||
| 6 | Kildall's algorithm | |
| 7 | *) | |
| 8 | ||
| 12911 | 9 | header {* \isaheader{Kildall's Algorithm}\label{sec:Kildall} *}
 | 
| 10496 | 10 | |
| 16417 | 11 | theory Kildall imports SemilatAlg While_Combinator begin | 
| 12516 | 12 | |
| 10496 | 13 | |
| 14 | consts | |
| 12516 | 15 | iter :: "'s binop \<Rightarrow> 's step_type \<Rightarrow> | 
| 11298 | 16 | 's list \<Rightarrow> nat set \<Rightarrow> 's list \<times> nat set" | 
| 13006 | 17 | propa :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list * nat set" | 
| 10496 | 18 | |
| 19 | primrec | |
| 12516 | 20 | "propa f [] ss w = (ss,w)" | 
| 21 | "propa f (q'#qs) ss w = (let (q,t) = q'; | |
| 22 | u = t +_f ss!q; | |
| 23 | w' = (if u = ss!q then w else insert q w) | |
| 24 | in propa f qs (ss[q := u]) w')" | |
| 10496 | 25 | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 26 | defs iter_def: | 
| 12516 | 27 | "iter f step ss w == | 
| 11298 | 28 |  while (%(ss,w). w \<noteq> {})
 | 
| 12516 | 29 | (%(ss,w). let p = SOME p. p \<in> w | 
| 30 |                  in propa f (step p (ss!p)) ss (w-{p}))
 | |
| 11298 | 31 | (ss,w)" | 
| 10496 | 32 | |
| 33 | constdefs | |
| 13006 | 34 | unstables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat set" | 
| 12516 | 35 | "unstables r step ss == {p. p < size ss \<and> \<not>stable r step ss p}"
 | 
| 10496 | 36 | |
| 13006 | 37 | kildall :: "'s ord \<Rightarrow> 's binop \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> 's list" | 
| 12516 | 38 | "kildall r f step ss == fst(iter f step ss (unstables r step ss))" | 
| 10496 | 39 | |
| 13006 | 40 | consts merges :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> 's list" | 
| 10496 | 41 | primrec | 
| 12516 | 42 | "merges f [] ss = ss" | 
| 43 | "merges f (p'#ps) ss = (let (p,s) = p' in merges f ps (ss[p := s +_f ss!p]))" | |
| 10496 | 44 | |
| 45 | ||
| 13074 | 46 | lemmas [simp] = Let_def semilat.le_iff_plus_unchanged [symmetric] | 
| 10496 | 47 | |
| 48 | ||
| 13074 | 49 | lemma (in semilat) nth_merges: | 
| 50 | "\<And>ss. \<lbrakk>p < length ss; ss \<in> list n A; \<forall>(p,t)\<in>set ps. p<n \<and> t\<in>A \<rbrakk> \<Longrightarrow> | |
| 12516 | 51 | (merges f ps ss)!p = map snd [(p',t') \<in> ps. p'=p] ++_f ss!p" | 
| 13074 | 52 | (is "\<And>ss. \<lbrakk>_; _; ?steptype ps\<rbrakk> \<Longrightarrow> ?P ss ps") | 
| 12516 | 53 | proof (induct ps) | 
| 54 | show "\<And>ss. ?P ss []" by simp | |
| 55 | ||
| 56 | fix ss p' ps' | |
| 57 | assume ss: "ss \<in> list n A" | |
| 58 | assume l: "p < length ss" | |
| 59 | assume "?steptype (p'#ps')" | |
| 60 | then obtain a b where | |
| 61 | p': "p'=(a,b)" and ab: "a<n" "b\<in>A" and "?steptype ps'" | |
| 62 | by (cases p', auto) | |
| 13074 | 63 | assume "\<And>ss. p< length ss \<Longrightarrow> ss \<in> list n A \<Longrightarrow> ?steptype ps' \<Longrightarrow> ?P ss ps'" | 
| 12516 | 64 | hence IH: "\<And>ss. ss \<in> list n A \<Longrightarrow> p < length ss \<Longrightarrow> ?P ss ps'" . | 
| 65 | ||
| 13074 | 66 | from ss ab | 
| 12516 | 67 | have "ss[a := b +_f ss!a] \<in> list n A" by (simp add: closedD) | 
| 68 | moreover | |
| 69 | from calculation | |
| 70 | have "p < length (ss[a := b +_f ss!a])" by simp | |
| 71 | ultimately | |
| 72 | have "?P (ss[a := b +_f ss!a]) ps'" by (rule IH) | |
| 13074 | 73 | with p' l | 
| 12516 | 74 | show "?P ss (p'#ps')" by simp | 
| 75 | qed | |
| 76 | ||
| 77 | ||
| 10496 | 78 | (** merges **) | 
| 79 | ||
| 80 | lemma length_merges [rule_format, simp]: | |
| 12516 | 81 | "\<forall>ss. size(merges f ps ss) = size ss" | 
| 10496 | 82 | by (induct_tac ps, auto) | 
| 83 | ||
| 12516 | 84 | |
| 13074 | 85 | lemma (in semilat) merges_preserves_type_lemma: | 
| 86 | shows "\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x) \<in> set ps. p<n \<and> x\<in>A) | |
| 87 | \<longrightarrow> merges f ps xs \<in> list n A" | |
| 88 | apply (insert closedI) | |
| 89 | apply (unfold closed_def) | |
| 90 | apply (induct_tac ps) | |
| 91 | apply simp | |
| 92 | apply clarsimp | |
| 93 | done | |
| 10496 | 94 | |
| 13074 | 95 | lemma (in semilat) merges_preserves_type [simp]: | 
| 96 | "\<lbrakk> xs \<in> list n A; \<forall>(p,x) \<in> set ps. p<n \<and> x\<in>A \<rbrakk> | |
| 13006 | 97 | \<Longrightarrow> merges f ps xs \<in> list n A" | 
| 13074 | 98 | by (simp add: merges_preserves_type_lemma) | 
| 99 | ||
| 100 | lemma (in semilat) merges_incr_lemma: | |
| 101 | "\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x)\<in>set ps. p<size xs \<and> x \<in> A) \<longrightarrow> xs <=[r] merges f ps xs" | |
| 102 | apply (induct_tac ps) | |
| 103 | apply simp | |
| 104 | apply simp | |
| 105 | apply clarify | |
| 106 | apply (rule order_trans) | |
| 10496 | 107 | apply simp | 
| 13074 | 108 | apply (erule list_update_incr) | 
| 109 | apply simp | |
| 110 | apply simp | |
| 111 | apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in]) | |
| 112 | done | |
| 10496 | 113 | |
| 13074 | 114 | lemma (in semilat) merges_incr: | 
| 115 | "\<lbrakk> xs \<in> list n A; \<forall>(p,x)\<in>set ps. p<size xs \<and> x \<in> A \<rbrakk> | |
| 13006 | 116 | \<Longrightarrow> xs <=[r] merges f ps xs" | 
| 12516 | 117 | by (simp add: merges_incr_lemma) | 
| 118 | ||
| 119 | ||
| 13074 | 120 | lemma (in semilat) merges_same_conv [rule_format]: | 
| 121 | "(\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x)\<in>set ps. p<size xs \<and> x\<in>A) \<longrightarrow> | |
| 12516 | 122 | (merges f ps xs = xs) = (\<forall>(p,x)\<in>set ps. x <=_r xs!p))" | 
| 10496 | 123 | apply (induct_tac ps) | 
| 124 | apply simp | |
| 125 | apply clarsimp | |
| 12516 | 126 | apply (rename_tac p x ps xs) | 
| 10496 | 127 | apply (rule iffI) | 
| 128 | apply (rule context_conjI) | |
| 129 | apply (subgoal_tac "xs[p := x +_f xs!p] <=[r] xs") | |
| 17537 
3825229092f0
fixed proof script of lemma merges_same_conv (Why did it stop working?);
 wenzelm parents: 
16417diff
changeset | 130 | apply (drule_tac p = p in le_listD) | 
| 
3825229092f0
fixed proof script of lemma merges_same_conv (Why did it stop working?);
 wenzelm parents: 
16417diff
changeset | 131 | apply simp | 
| 
3825229092f0
fixed proof script of lemma merges_same_conv (Why did it stop working?);
 wenzelm parents: 
16417diff
changeset | 132 | apply simp | 
| 10496 | 133 | apply (erule subst, rule merges_incr) | 
| 134 | apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in]) | |
| 12516 | 135 | apply clarify | 
| 136 | apply (rule conjI) | |
| 137 | apply simp | |
| 138 | apply (blast dest: boundedD) | |
| 139 | apply blast | |
| 10496 | 140 | apply clarify | 
| 141 | apply (erule allE) | |
| 142 | apply (erule impE) | |
| 143 | apply assumption | |
| 144 | apply (drule bspec) | |
| 145 | apply assumption | |
| 146 | apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2]) | |
| 12516 | 147 | apply blast | 
| 10496 | 148 | apply clarify | 
| 149 | apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2]) | |
| 150 | done | |
| 151 | ||
| 152 | ||
| 13074 | 153 | lemma (in semilat) list_update_le_listI [rule_format]: | 
| 13006 | 154 | "set xs <= A \<longrightarrow> set ys <= A \<longrightarrow> xs <=[r] ys \<longrightarrow> p < size xs \<longrightarrow> | 
| 13074 | 155 | x <=_r ys!p \<longrightarrow> x\<in>A \<longrightarrow> xs[p := x +_f xs!p] <=[r] ys" | 
| 156 | apply(insert semilat) | |
| 10496 | 157 | apply (unfold Listn.le_def lesub_def semilat_def) | 
| 158 | apply (simp add: list_all2_conv_all_nth nth_list_update) | |
| 159 | done | |
| 160 | ||
| 13074 | 161 | lemma (in semilat) merges_pres_le_ub: | 
| 18372 | 162 | assumes "set ts <= A" and "set ss <= A" | 
| 163 | and "\<forall>(p,t)\<in>set ps. t <=_r ts!p \<and> t \<in> A \<and> p < size ts" and "ss <=[r] ts" | |
| 164 | shows "merges f ps ss <=[r] ts" | |
| 10496 | 165 | proof - | 
| 13074 | 166 |   { fix t ts ps
 | 
| 10496 | 167 | have | 
| 13074 | 168 | "\<And>qs. \<lbrakk>set ts <= A; \<forall>(p,t)\<in>set ps. t <=_r ts!p \<and> t \<in> A \<and> p< size ts \<rbrakk> \<Longrightarrow> | 
| 13006 | 169 | set qs <= set ps \<longrightarrow> | 
| 170 | (\<forall>ss. set ss <= A \<longrightarrow> ss <=[r] ts \<longrightarrow> merges f qs ss <=[r] ts)" | |
| 10496 | 171 | apply (induct_tac qs) | 
| 172 | apply simp | |
| 173 | apply (simp (no_asm_simp)) | |
| 174 | apply clarify | |
| 175 | apply (rotate_tac -2) | |
| 176 | apply simp | |
| 177 | apply (erule allE, erule impE, erule_tac [2] mp) | |
| 12516 | 178 | apply (drule bspec, assumption) | 
| 179 | apply (simp add: closedD) | |
| 180 | apply (drule bspec, assumption) | |
| 10496 | 181 | apply (simp add: list_update_le_listI) | 
| 12516 | 182 | done | 
| 10496 | 183 | } note this [dest] | 
| 184 | ||
| 18372 | 185 | from prems show ?thesis by blast | 
| 10496 | 186 | qed | 
| 187 | ||
| 188 | ||
| 12516 | 189 | (** propa **) | 
| 10496 | 190 | |
| 191 | ||
| 12516 | 192 | lemma decomp_propa: | 
| 13006 | 193 | "\<And>ss w. (\<forall>(q,t)\<in>set qs. q < size ss) \<Longrightarrow> | 
| 12516 | 194 | propa f qs ss w = | 
| 195 |    (merges f qs ss, {q. \<exists>t. (q,t)\<in>set qs \<and> t +_f ss!q \<noteq> ss!q} Un w)"
 | |
| 196 | apply (induct qs) | |
| 197 | apply simp | |
| 10496 | 198 | apply (simp (no_asm)) | 
| 12516 | 199 | apply clarify | 
| 200 | apply simp | |
| 201 | apply (rule conjI) | |
| 10496 | 202 | apply blast | 
| 203 | apply (simp add: nth_list_update) | |
| 204 | apply blast | |
| 205 | done | |
| 206 | ||
| 207 | (** iter **) | |
| 208 | ||
| 13074 | 209 | lemma (in semilat) stable_pres_lemma: | 
| 210 | shows "\<lbrakk>pres_type step n A; bounded step n; | |
| 12516 | 211 | ss \<in> list n A; p \<in> w; \<forall>q\<in>w. q < n; | 
| 212 | \<forall>q. q < n \<longrightarrow> q \<notin> w \<longrightarrow> stable r step ss q; q < n; | |
| 213 | \<forall>s'. (q,s') \<in> set (step p (ss ! p)) \<longrightarrow> s' +_f ss ! q = ss ! q; | |
| 13006 | 214 | q \<notin> w \<or> q = p \<rbrakk> | 
| 215 | \<Longrightarrow> stable r step (merges f (step p (ss!p)) ss) q" | |
| 10496 | 216 | apply (unfold stable_def) | 
| 12516 | 217 | apply (subgoal_tac "\<forall>s'. (q,s') \<in> set (step p (ss!p)) \<longrightarrow> s' : A") | 
| 218 | prefer 2 | |
| 219 | apply clarify | |
| 220 | apply (erule pres_typeD) | |
| 221 | prefer 3 apply assumption | |
| 222 | apply (rule listE_nth_in) | |
| 223 | apply assumption | |
| 224 | apply simp | |
| 225 | apply simp | |
| 10496 | 226 | apply simp | 
| 227 | apply clarify | |
| 13074 | 228 | apply (subst nth_merges) | 
| 12516 | 229 | apply simp | 
| 230 | apply (blast dest: boundedD) | |
| 10496 | 231 | apply assumption | 
| 12516 | 232 | apply clarify | 
| 233 | apply (rule conjI) | |
| 234 | apply (blast dest: boundedD) | |
| 235 | apply (erule pres_typeD) | |
| 236 | prefer 3 apply assumption | |
| 237 | apply simp | |
| 13074 | 238 | apply simp | 
| 239 | apply(subgoal_tac "q < length ss") | |
| 240 | prefer 2 apply simp | |
| 241 | apply (frule nth_merges [of q _ _ "step p (ss!p)"]) (* fixme: why does method subst not work?? *) | |
| 242 | apply assumption | |
| 12516 | 243 | apply clarify | 
| 10496 | 244 | apply (rule conjI) | 
| 12516 | 245 | apply (blast dest: boundedD) | 
| 246 | apply (erule pres_typeD) | |
| 247 | prefer 3 apply assumption | |
| 248 | apply simp | |
| 249 | apply simp | |
| 250 | apply (drule_tac P = "\<lambda>x. (a, b) \<in> set (step q x)" in subst) | |
| 251 | apply assumption | |
| 252 | ||
| 253 | apply (simp add: plusplus_empty) | |
| 254 | apply (cases "q \<in> w") | |
| 255 | apply simp | |
| 256 | apply (rule ub1') | |
| 257 | apply assumption | |
| 258 | apply clarify | |
| 259 | apply (rule pres_typeD) | |
| 260 | apply assumption | |
| 261 | prefer 3 apply assumption | |
| 262 | apply (blast intro: listE_nth_in dest: boundedD) | |
| 263 | apply (blast intro: pres_typeD dest: boundedD) | |
| 264 | apply (blast intro: listE_nth_in dest: boundedD) | |
| 265 | apply assumption | |
| 10496 | 266 | |
| 12516 | 267 | apply simp | 
| 268 | apply (erule allE, erule impE, assumption, erule impE, assumption) | |
| 269 | apply (rule order_trans) | |
| 270 | apply simp | |
| 271 | defer | |
| 13074 | 272 | apply (rule pp_ub2)(* | 
| 273 | apply assumption*) | |
| 12516 | 274 | apply simp | 
| 275 | apply clarify | |
| 276 | apply simp | |
| 277 | apply (rule pres_typeD) | |
| 278 | apply assumption | |
| 279 | prefer 3 apply assumption | |
| 280 | apply (blast intro: listE_nth_in dest: boundedD) | |
| 281 | apply (blast intro: pres_typeD dest: boundedD) | |
| 282 | apply (blast intro: listE_nth_in dest: boundedD) | |
| 283 | apply blast | |
| 284 | done | |
| 13074 | 285 | |
| 286 | ||
| 287 | lemma (in semilat) merges_bounded_lemma: | |
| 288 | "\<lbrakk> mono r step n A; bounded step n; | |
| 289 | \<forall>(p',s') \<in> set (step p (ss!p)). s' \<in> A; ss \<in> list n A; ts \<in> list n A; p < n; | |
| 290 | ss <=[r] ts; \<forall>p. p < n \<longrightarrow> stable r step ts p \<rbrakk> | |
| 13066 | 291 | \<Longrightarrow> merges f (step p (ss!p)) ss <=[r] ts" | 
| 10496 | 292 | apply (unfold stable_def) | 
| 12516 | 293 | apply (rule merges_pres_le_ub) | 
| 294 | apply simp | |
| 295 | apply simp | |
| 296 | prefer 2 apply assumption | |
| 297 | ||
| 298 | apply clarsimp | |
| 299 | apply (drule boundedD, assumption+) | |
| 300 | apply (erule allE, erule impE, assumption) | |
| 301 | apply (drule bspec, assumption) | |
| 302 | apply simp | |
| 303 | ||
| 304 | apply (drule monoD [of _ _ _ _ p "ss!p" "ts!p"]) | |
| 305 | apply assumption | |
| 306 | apply simp | |
| 307 | apply (simp add: le_listD) | |
| 308 | ||
| 13066 | 309 | apply (drule lesub_step_typeD, assumption) | 
| 12516 | 310 | apply clarify | 
| 311 | apply (drule bspec, assumption) | |
| 312 | apply simp | |
| 313 | apply (blast intro: order_trans) | |
| 10496 | 314 | done | 
| 315 | ||
| 13074 | 316 | lemma termination_lemma: includes semilat | 
| 317 | shows "\<lbrakk> ss \<in> list n A; \<forall>(q,t)\<in>set qs. q<n \<and> t\<in>A; p\<in>w \<rbrakk> \<Longrightarrow> | |
| 12516 | 318 | ss <[r] merges f qs ss \<or> | 
| 319 |   merges f qs ss = ss \<and> {q. \<exists>t. (q,t)\<in>set qs \<and> t +_f ss!q \<noteq> ss!q} Un (w-{p}) < w"
 | |
| 13074 | 320 | apply(insert semilat) | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 321 | apply (unfold lesssub_def) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 322 | apply (simp (no_asm_simp) add: merges_incr) | 
| 10496 | 323 | apply (rule impI) | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 324 | apply (rule merges_same_conv [THEN iffD1, elim_format]) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 325 | apply assumption+ | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 326 | defer | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 327 | apply (rule sym, assumption) | 
| 12516 | 328 | defer apply simp | 
| 329 | apply (subgoal_tac "\<forall>q t. \<not>((q, t) \<in> set qs \<and> t +_f ss ! q \<noteq> ss ! q)") | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 330 | apply (blast intro!: psubsetI elim: equalityE) | 
| 12516 | 331 | apply clarsimp | 
| 332 | apply (drule bspec, assumption) | |
| 333 | apply (drule bspec, assumption) | |
| 334 | apply clarsimp | |
| 335 | done | |
| 10496 | 336 | |
| 13074 | 337 | lemma iter_properties[rule_format]: includes semilat | 
| 338 | shows "\<lbrakk> acc r ; pres_type step n A; mono r step n A; | |
| 12516 | 339 | bounded step n; \<forall>p\<in>w0. p < n; ss0 \<in> list n A; | 
| 340 | \<forall>p<n. p \<notin> w0 \<longrightarrow> stable r step ss0 p \<rbrakk> \<Longrightarrow> | |
| 341 | iter f step ss0 w0 = (ss',w') | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 342 | \<longrightarrow> | 
| 12516 | 343 | ss' \<in> list n A \<and> stables r step ss' \<and> ss0 <=[r] ss' \<and> | 
| 344 | (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> ss' <=[r] ts)" | |
| 13074 | 345 | apply(insert semilat) | 
| 11298 | 346 | apply (unfold iter_def stables_def) | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 347 | apply (rule_tac P = "%(ss,w). | 
| 12516 | 348 | ss \<in> list n A \<and> (\<forall>p<n. p \<notin> w \<longrightarrow> stable r step ss p) \<and> ss0 <=[r] ss \<and> | 
| 349 | (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> ss <=[r] ts) \<and> | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 350 | (\<forall>p\<in>w. p < n)" and | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 351 |  r = "{(ss',ss) . ss <[r] ss'} <*lex*> finite_psubset"
 | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 352 | in while_rule) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 353 | |
| 12516 | 354 | -- "Invariant holds initially:" | 
| 13074 | 355 | apply (simp add:stables_def) | 
| 12516 | 356 | |
| 357 | -- "Invariant is preserved:" | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 358 | apply(simp add: stables_def split_paired_all) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 359 | apply(rename_tac ss w) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 360 | apply(subgoal_tac "(SOME p. p \<in> w) \<in> w") | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 361 | prefer 2; apply (fast intro: someI) | 
| 12516 | 362 | apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A") | 
| 363 | prefer 2 | |
| 364 | apply clarify | |
| 365 | apply (rule conjI) | |
| 366 | apply(clarsimp, blast dest!: boundedD) | |
| 367 | apply (erule pres_typeD) | |
| 368 | prefer 3 | |
| 369 | apply assumption | |
| 370 | apply (erule listE_nth_in) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14191diff
changeset | 371 | apply simp | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14191diff
changeset | 372 | apply simp | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 373 | apply (subst decomp_propa) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14191diff
changeset | 374 | apply fast | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 375 | apply simp | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 376 | apply (rule conjI) | 
| 13074 | 377 | apply (rule merges_preserves_type) | 
| 378 | apply blast | |
| 12516 | 379 | apply clarify | 
| 380 | apply (rule conjI) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14191diff
changeset | 381 | apply(clarsimp, fast dest!: boundedD) | 
| 12516 | 382 | apply (erule pres_typeD) | 
| 383 | prefer 3 | |
| 384 | apply assumption | |
| 385 | apply (erule listE_nth_in) | |
| 386 | apply blast | |
| 387 | apply blast | |
| 13074 | 388 | apply (rule conjI) | 
| 389 | apply clarify | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 390 | apply (blast intro!: stable_pres_lemma) | 
| 13074 | 391 | apply (rule conjI) | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 392 | apply (blast intro!: merges_incr intro: le_list_trans) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 393 | apply (rule conjI) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 394 | apply clarsimp | 
| 12516 | 395 | apply (blast intro!: merges_bounded_lemma) | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 396 | apply (blast dest!: boundedD) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 397 | |
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 398 | |
| 12516 | 399 | -- "Postcondition holds upon termination:" | 
| 13074 | 400 | apply(clarsimp simp add: stables_def split_paired_all) | 
| 12516 | 401 | |
| 402 | -- "Well-foundedness of the termination relation:" | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 403 | apply (rule wf_lex_prod) | 
| 13074 | 404 | apply (insert orderI [THEN acc_le_listI]) | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 405 | apply (simp only: acc_def lesssub_def) | 
| 12516 | 406 | apply (rule wf_finite_psubset) | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 407 | |
| 12516 | 408 | -- "Loop decreases along termination relation:" | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 409 | apply(simp add: stables_def split_paired_all) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 410 | apply(rename_tac ss w) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 411 | apply(subgoal_tac "(SOME p. p \<in> w) \<in> w") | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 412 | prefer 2; apply (fast intro: someI) | 
| 12516 | 413 | apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A") | 
| 414 | prefer 2 | |
| 415 | apply clarify | |
| 416 | apply (rule conjI) | |
| 417 | apply(clarsimp, blast dest!: boundedD) | |
| 418 | apply (erule pres_typeD) | |
| 419 | prefer 3 | |
| 420 | apply assumption | |
| 421 | apply (erule listE_nth_in) | |
| 422 | apply blast | |
| 423 | apply blast | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 424 | apply (subst decomp_propa) | 
| 12516 | 425 | apply blast | 
| 13074 | 426 | apply clarify | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 427 | apply (simp del: listE_length | 
| 12516 | 428 | add: lex_prod_def finite_psubset_def | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 429 | bounded_nat_set_is_finite) | 
| 12516 | 430 | apply (rule termination_lemma) | 
| 431 | apply assumption+ | |
| 432 | defer | |
| 433 | apply assumption | |
| 434 | apply clarsimp | |
| 13074 | 435 | done | 
| 12516 | 436 | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 437 | |
| 13074 | 438 | lemma kildall_properties: includes semilat | 
| 439 | shows "\<lbrakk> acc r; pres_type step n A; mono r step n A; | |
| 12516 | 440 | bounded step n; ss0 \<in> list n A \<rbrakk> \<Longrightarrow> | 
| 441 | kildall r f step ss0 \<in> list n A \<and> | |
| 442 | stables r step (kildall r f step ss0) \<and> | |
| 443 | ss0 <=[r] kildall r f step ss0 \<and> | |
| 444 | (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> | |
| 445 | kildall r f step ss0 <=[r] ts)" | |
| 11229 | 446 | apply (unfold kildall_def) | 
| 12516 | 447 | apply(case_tac "iter f step ss0 (unstables r step ss0)") | 
| 11298 | 448 | apply(simp) | 
| 11229 | 449 | apply (rule iter_properties) | 
| 13074 | 450 | by (simp_all add: unstables_def stable_def) | 
| 451 | ||
| 10496 | 452 | |
| 13074 | 453 | lemma is_bcv_kildall: includes semilat | 
| 454 | shows "\<lbrakk> acc r; top r T; pres_type step n A; bounded step n; mono r step n A \<rbrakk> | |
| 13006 | 455 | \<Longrightarrow> is_bcv r T step n A (kildall r f step)" | 
| 11299 | 456 | apply(unfold is_bcv_def wt_step_def) | 
| 13074 | 457 | apply(insert semilat kildall_properties[of A]) | 
| 11229 | 458 | apply(simp add:stables_def) | 
| 459 | apply clarify | |
| 12516 | 460 | apply(subgoal_tac "kildall r f step ss \<in> list n A") | 
| 11229 | 461 | prefer 2 apply (simp(no_asm_simp)) | 
| 462 | apply (rule iffI) | |
| 12516 | 463 | apply (rule_tac x = "kildall r f step ss" in bexI) | 
| 11229 | 464 | apply (rule conjI) | 
| 13074 | 465 | apply (blast) | 
| 11229 | 466 | apply (simp (no_asm_simp)) | 
| 13074 | 467 | apply(assumption) | 
| 11229 | 468 | apply clarify | 
| 12516 | 469 | apply(subgoal_tac "kildall r f step ss!p <=_r ts!p") | 
| 11229 | 470 | apply simp | 
| 471 | apply (blast intro!: le_listD less_lengthI) | |
| 472 | done | |
| 10496 | 473 | |
| 14191 | 474 | end |