| author | paulson | 
| Wed, 07 Mar 2001 18:35:27 +0100 | |
| changeset 11199 | 97cde35cec10 | 
| parent 11184 | 10a307328d2c | 
| child 11229 | f417841385b7 | 
| permissions | -rw-r--r-- | 
| 10496 | 1 | (* Title: HOL/BCV/Kildall.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 2000 TUM | |
| 5 | ||
| 6 | Kildall's algorithm | |
| 7 | *) | |
| 8 | ||
| 9 | header "Kildall's Algorithm" | |
| 10 | ||
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 11 | theory Kildall = DFA_Framework + While_Combinator: | 
| 10496 | 12 | |
| 13 | constdefs | |
| 14 | pres_type :: "(nat => 's => 's) => nat => 's set => bool" | |
| 15 | "pres_type step n A == !s:A. !p<n. step p s : A" | |
| 16 | ||
| 17 | mono :: "'s ord => (nat => 's => 's) => nat => 's set => bool" | |
| 18 | "mono r step n A == | |
| 19 | !s p t. s : A & p < n & s <=_r t --> step p s <=_r step p t" | |
| 20 | ||
| 21 | consts | |
| 11184 | 22 | iter :: "'s binop \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> 's) \<Rightarrow> (nat \<Rightarrow> nat list) \<Rightarrow> | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 23 | 's list \<Rightarrow> nat set \<Rightarrow> 's list" | 
| 10496 | 24 | propa :: "'s binop => nat list => 's => 's list => nat set => 's list * nat set" | 
| 25 | ||
| 26 | primrec | |
| 27 | "propa f [] t ss w = (ss,w)" | |
| 28 | "propa f (q#qs) t ss w = (let u = t +_f ss!q; | |
| 29 | w' = (if u = ss!q then w else insert q w) | |
| 30 | in propa f qs t (ss[q := u]) w')" | |
| 31 | ||
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 32 | defs iter_def: | 
| 11184 | 33 | "iter f step succs ss w == | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 34 |  fst(while (%(ss,w). w \<noteq> {})
 | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 35 | (%(ss,w). let p = SOME p. p : w | 
| 11184 | 36 |                      in propa f (succs p) (step p (ss!p)) ss (w-{p}))
 | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 37 | (ss,w))" | 
| 10496 | 38 | |
| 39 | constdefs | |
| 40 | unstables :: | |
| 41 | "'s binop => (nat => 's => 's) => (nat => nat list) => 's list => nat set" | |
| 42 | "unstables f step succs ss == | |
| 43 |  {p. p < size ss & (? q:set(succs p). step p (ss!p) +_f ss!q ~= ss!q)}"
 | |
| 44 | ||
| 11184 | 45 | kildall :: "'s binop => (nat => 's => 's) => (nat => nat list) | 
| 10496 | 46 | => 's list => 's list" | 
| 11184 | 47 | "kildall f step succs ss == iter f step succs ss (unstables f step succs ss)" | 
| 10496 | 48 | |
| 49 | consts merges :: "'s binop => 's => nat list => 's list => 's list" | |
| 50 | primrec | |
| 51 | "merges f s [] ss = ss" | |
| 52 | "merges f s (p#ps) ss = merges f s ps (ss[p := s +_f ss!p])" | |
| 53 | ||
| 54 | ||
| 10774 | 55 | lemmas [simp] = Let_def le_iff_plus_unchanged [symmetric] | 
| 10496 | 56 | |
| 57 | ||
| 58 | lemma pres_typeD: | |
| 59 | "[| pres_type step n A; s:A; p<n |] ==> step p s : A" | |
| 60 | by (unfold pres_type_def, blast) | |
| 61 | ||
| 62 | lemma boundedD: | |
| 63 | "[| bounded succs n; p < n; q : set(succs p) |] ==> q < n" | |
| 64 | by (unfold bounded_def, blast) | |
| 65 | ||
| 66 | lemma monoD: | |
| 67 | "[| mono r step n A; p < n; s:A; s <=_r t |] ==> step p s <=_r step p t" | |
| 68 | by (unfold mono_def, blast) | |
| 69 | ||
| 70 | (** merges **) | |
| 71 | ||
| 72 | lemma length_merges [rule_format, simp]: | |
| 73 | "!ss. size(merges f t ps ss) = size ss" | |
| 74 | by (induct_tac ps, auto) | |
| 75 | ||
| 76 | lemma merges_preserves_type [rule_format, simp]: | |
| 77 | "[| semilat(A,r,f) |] ==> | |
| 78 | !xs. xs : list n A --> x : A --> (!p : set ps. p<n) | |
| 79 | --> merges f x ps xs : list n A" | |
| 80 | apply (frule semilatDclosedI) | |
| 81 | apply (unfold closed_def) | |
| 82 | apply (induct_tac ps) | |
| 83 | apply auto | |
| 84 | done | |
| 85 | ||
| 86 | lemma merges_incr [rule_format]: | |
| 87 | "[| semilat(A,r,f) |] ==> | |
| 88 | !xs. xs : list n A --> x : A --> (!p:set ps. p<size xs) --> xs <=[r] merges f x ps xs" | |
| 89 | apply (induct_tac ps) | |
| 90 | apply simp | |
| 91 | apply simp | |
| 92 | apply clarify | |
| 93 | apply (rule order_trans) | |
| 94 | apply simp | |
| 95 | apply (erule list_update_incr) | |
| 96 | apply assumption | |
| 97 | apply simp | |
| 98 | apply simp | |
| 99 | apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in]) | |
| 100 | done | |
| 101 | ||
| 102 | lemma merges_same_conv [rule_format]: | |
| 103 | "[| semilat(A,r,f) |] ==> | |
| 104 | (!xs. xs : list n A --> x:A --> (!p:set ps. p<size xs) --> | |
| 105 | (merges f x ps xs = xs) = (!p:set ps. x <=_r xs!p))" | |
| 106 | apply (induct_tac ps) | |
| 107 | apply simp | |
| 108 | apply clarsimp | |
| 109 | apply (rename_tac p ps xs) | |
| 110 | apply (rule iffI) | |
| 111 | apply (rule context_conjI) | |
| 112 | apply (subgoal_tac "xs[p := x +_f xs!p] <=[r] xs") | |
| 113 | apply (force dest!: le_listD simp add: nth_list_update) | |
| 114 | apply (erule subst, rule merges_incr) | |
| 115 | apply assumption | |
| 116 | apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in]) | |
| 117 | apply assumption | |
| 118 | apply simp | |
| 119 | apply clarify | |
| 120 | apply (rotate_tac -2) | |
| 121 | apply (erule allE) | |
| 122 | apply (erule impE) | |
| 123 | apply assumption | |
| 124 | apply (erule impE) | |
| 125 | apply assumption | |
| 126 | apply (drule bspec) | |
| 127 | apply assumption | |
| 128 | apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2]) | |
| 129 | apply clarify | |
| 130 | apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2]) | |
| 131 | done | |
| 132 | ||
| 133 | ||
| 134 | lemma list_update_le_listI [rule_format]: | |
| 135 | "set xs <= A --> set ys <= A --> xs <=[r] ys --> p < size xs --> | |
| 136 | x <=_r ys!p --> semilat(A,r,f) --> x:A --> | |
| 137 | xs[p := x +_f xs!p] <=[r] ys" | |
| 138 | apply (unfold Listn.le_def lesub_def semilat_def) | |
| 139 | apply (simp add: list_all2_conv_all_nth nth_list_update) | |
| 140 | done | |
| 141 | ||
| 142 | lemma merges_pres_le_ub: | |
| 143 | "[| semilat(A,r,f); t:A; set ts <= A; set ss <= A; | |
| 144 | !p. p:set ps --> t <=_r ts!p; | |
| 145 | !p. p:set ps --> p<size ts; | |
| 146 | ss <=[r] ts |] | |
| 147 | ==> merges f t ps ss <=[r] ts" | |
| 148 | proof - | |
| 149 |   { fix A r f t ts ps
 | |
| 150 | have | |
| 151 | "!!qs. [| semilat(A,r,f); set ts <= A; t:A; | |
| 152 | !p. p:set ps --> t <=_r ts!p; | |
| 153 | !p. p:set ps --> p<size ts |] ==> | |
| 154 | set qs <= set ps --> | |
| 155 | (!ss. set ss <= A --> ss <=[r] ts --> merges f t qs ss <=[r] ts)" | |
| 156 | apply (induct_tac qs) | |
| 157 | apply simp | |
| 158 | apply (simp (no_asm_simp)) | |
| 159 | apply clarify | |
| 160 | apply (rotate_tac -2) | |
| 161 | apply simp | |
| 162 | apply (erule allE, erule impE, erule_tac [2] mp) | |
| 163 | apply (simp (no_asm_simp) add: closedD) | |
| 164 | apply (simp add: list_update_le_listI) | |
| 165 | done | |
| 166 | } note this [dest] | |
| 167 | ||
| 168 | case antecedent | |
| 169 | thus ?thesis by blast | |
| 170 | qed | |
| 171 | ||
| 172 | ||
| 173 | lemma nth_merges [rule_format]: | |
| 174 | "[| semilat (A, r, f); t:A; p < n |] ==> !ss. ss : list n A --> | |
| 175 | (!p:set ps. p<n) --> | |
| 176 | (merges f t ps ss)!p = (if p : set ps then t +_f ss!p else ss!p)" | |
| 177 | apply (induct_tac "ps") | |
| 178 | apply simp | |
| 179 | apply (simp add: nth_list_update closedD listE_nth_in) | |
| 180 | done | |
| 181 | ||
| 182 | ||
| 183 | (** propa **) | |
| 184 | ||
| 185 | lemma decomp_propa [rule_format]: | |
| 186 | "!ss w. (!q:set qs. q < size ss) --> | |
| 187 | propa f qs t ss w = | |
| 188 |    (merges f t qs ss, {q. q:set qs & t +_f ss!q ~= ss!q} Un w)"
 | |
| 189 | apply (induct_tac qs) | |
| 190 | apply simp | |
| 191 | apply (simp (no_asm)) | |
| 192 | apply clarify | |
| 193 | apply (rule conjI) | |
| 194 | apply (simp add: nth_list_update) | |
| 195 | apply blast | |
| 196 | apply (simp add: nth_list_update) | |
| 197 | apply blast | |
| 198 | done | |
| 199 | ||
| 200 | (** iter **) | |
| 201 | ||
| 202 | lemma stable_pres_lemma: | |
| 203 | "[| semilat (A,r,f); pres_type step n A; bounded succs n; | |
| 204 | ss : list n A; p : w; ! q:w. q < n; | |
| 205 | ! q. q < n --> q ~: w --> stable r step succs ss q; q < n; | |
| 206 | q : set (succs p) --> step p (ss ! p) +_f ss ! q = ss ! q; | |
| 207 | q ~: w | q = p |] | |
| 208 | ==> stable r step succs (merges f (step p (ss!p)) (succs p) ss) q" | |
| 209 | apply (unfold stable_def) | |
| 210 | apply (subgoal_tac "step p (ss!p) : A") | |
| 211 | defer | |
| 212 | apply (blast intro: listE_nth_in pres_typeD) | |
| 213 | apply simp | |
| 214 | apply clarify | |
| 215 | apply (subst nth_merges) | |
| 216 | apply assumption | |
| 217 | apply assumption | |
| 218 | prefer 2 | |
| 219 | apply assumption | |
| 220 | apply (blast dest: boundedD) | |
| 221 | apply (blast dest: boundedD) | |
| 222 | apply (subst nth_merges) | |
| 223 | apply assumption | |
| 224 | apply assumption | |
| 225 | prefer 2 | |
| 226 | apply assumption | |
| 227 | apply (blast dest: boundedD) | |
| 228 | apply (blast dest: boundedD) | |
| 229 | apply simp | |
| 230 | apply (rule conjI) | |
| 231 | apply clarify | |
| 232 | apply (blast intro!: semilat_ub1 semilat_ub2 listE_nth_in | |
| 233 | intro: order_trans dest: boundedD) | |
| 234 | apply (blast intro!: semilat_ub1 semilat_ub2 listE_nth_in | |
| 235 | intro: order_trans dest: boundedD) | |
| 236 | done | |
| 237 | ||
| 238 | lemma merges_bounded_lemma [rule_format]: | |
| 239 | "[| semilat (A,r,f); mono r step n A; bounded succs n; | |
| 240 | step p (ss!p) : A; ss : list n A; ts : list n A; p < n; | |
| 241 | ss <=[r] ts; ! p. p < n --> stable r step succs ts p |] | |
| 242 | ==> merges f (step p (ss!p)) (succs p) ss <=[r] ts" | |
| 243 | apply (unfold stable_def) | |
| 244 | apply (blast intro!: listE_set monoD listE_nth_in le_listD less_lengthI | |
| 245 | intro: merges_pres_le_ub order_trans | |
| 246 | dest: pres_typeD boundedD) | |
| 247 | done | |
| 248 | ||
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 249 | lemma termination_lemma: | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 250 | "[| semilat(A,r,f); ss : list n A; t:A; ! q:set qs. q < n; p:w |] ==> | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 251 | ss <[r] merges f t qs ss | | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 252 |   merges f t qs ss = ss & {q. q:set qs & t +_f ss!q ~= ss!q} Un (w-{p}) < w"
 | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 253 | apply (unfold lesssub_def) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 254 | apply (simp (no_asm_simp) add: merges_incr) | 
| 10496 | 255 | apply (rule impI) | 
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 256 | apply (rule merges_same_conv [THEN iffD1, elim_format]) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 257 | apply assumption+ | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 258 | defer | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 259 | apply (rule sym, assumption) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 260 | apply (simp cong add: conj_cong add: le_iff_plus_unchanged [symmetric]) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 261 | apply (blast intro!: psubsetI elim: equalityE) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 262 | apply simp | 
| 10496 | 263 | done | 
| 264 | ||
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 265 | lemma while_properties: | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 266 | "\<lbrakk> semilat(A,r,f); acc r ; pres_type step n A; mono r step n A; | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 267 | bounded succs n; \<forall>p\<in>w0. p < n; ss0 \<in> list n A; | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 268 | \<forall>p<n. p \<notin> w0 \<longrightarrow> stable r step succs ss0 p \<rbrakk> \<Longrightarrow> | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 269 |    ss' = fst(while (%(ss,w). w \<noteq> {})
 | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 270 | (%(ss,w). let p = SOME p. p \<in> w | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 271 |                    in propa f (succs p) (step p (ss!p)) ss (w-{p})) (ss0,w0))
 | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 272 | \<longrightarrow> | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 273 | ss' \<in> list n A \<and> stables r step succs ss' \<and> ss0 <=[r] ss' \<and> | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 274 | (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step succs ts \<longrightarrow> ss' <=[r] ts)" | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 275 | apply (unfold stables_def) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 276 | apply (rule_tac P = "%(ss,w). | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 277 | ss \<in> list n A \<and> (\<forall>p<n. p \<notin> w \<longrightarrow> stable r step succs ss p) \<and> ss0 <=[r] ss \<and> | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 278 | (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step succs ts \<longrightarrow> ss <=[r] ts) \<and> | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 279 | (\<forall>p\<in>w. p < n)" and | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 280 |  r = "{(ss',ss) . ss <[r] ss'} <*lex*> finite_psubset"
 | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 281 | in while_rule) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 282 | |
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 283 | (* Invariant holds initially: *) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 284 | apply (simp add:stables_def) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 285 | (* Invariant is preserved: *) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 286 | apply(simp add: stables_def split_paired_all) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 287 | apply(rename_tac ss w) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 288 | 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 | 289 | prefer 2; apply (fast intro: someI) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 290 | apply(subgoal_tac "step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w)) \<in> A") | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 291 | prefer 2; apply (blast intro: pres_typeD listE_nth_in dest: boundedD); | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 292 | apply (subgoal_tac "\<forall>q\<in>set(succs (SOME p. p \<in> w)). q < size ss") | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 293 | prefer 2; apply(clarsimp, blast dest!: boundedD) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 294 | apply (subst decomp_propa) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 295 | apply blast | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 296 | apply simp | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 297 | apply (rule conjI) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 298 | apply (blast intro: merges_preserves_type dest: boundedD); | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 299 | apply (rule conjI) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 300 | apply (blast intro!: stable_pres_lemma) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 301 | apply (rule conjI) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 302 | apply (blast intro!: merges_incr intro: le_list_trans) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 303 | apply (rule conjI) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 304 | apply clarsimp | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 305 | apply (best intro!: merges_bounded_lemma) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 306 | apply (blast dest!: boundedD) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 307 | |
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 308 | (* Postcondition holds upon termination: *) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 309 | apply(clarsimp simp add: stables_def split_paired_all) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 310 | |
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 311 | (* Well-foundedness of the termination relation: *) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 312 | apply (rule wf_lex_prod) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 313 | apply (drule (1) semilatDorderI [THEN acc_le_listI]) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 314 | apply (simp only: acc_def lesssub_def) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 315 | apply (rule wf_finite_psubset) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 316 | |
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 317 | (* Loop decreases along termination relation: *) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 318 | apply(simp add: stables_def split_paired_all) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 319 | apply(rename_tac ss w) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 320 | 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 | 321 | prefer 2; apply (fast intro: someI) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 322 | apply(subgoal_tac "step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w)) \<in> A") | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 323 | prefer 2; apply (blast intro: pres_typeD listE_nth_in dest: boundedD); | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 324 | apply (subgoal_tac "\<forall>q\<in>set(succs (SOME p. p \<in> w)). q < n") | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 325 | prefer 2; apply(clarsimp, blast dest!: boundedD) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 326 | apply (subst decomp_propa) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 327 | apply clarsimp | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 328 | apply clarify | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 329 | apply (simp del: listE_length | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 330 | add: lex_prod_def finite_psubset_def decomp_propa termination_lemma | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 331 | bounded_nat_set_is_finite) | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 332 | done | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 333 | |
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 334 | lemma iter_properties: | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 335 | "\<lbrakk> semilat(A,r,f); acc r ; pres_type step n A; mono r step n A; | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 336 | bounded succs n; \<forall>p\<in>w0. p < n; ss0 \<in> list n A; | 
| 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 337 | \<forall>p<n. p \<notin> w0 \<longrightarrow> stable r step succs ss0 p \<rbrakk> \<Longrightarrow> | 
| 11184 | 338 | iter f step succs ss0 w0 : list n A \<and> | 
| 339 | stables r step succs (iter f step succs ss0 w0) \<and> | |
| 340 | ss0 <=[r] iter f step succs ss0 w0 \<and> | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 341 | (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step succs ts \<longrightarrow> | 
| 11184 | 342 | iter f step succs ss0 w0 <=[r] ts)" | 
| 343 | apply(simp add:iter_def del:Let_def) | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 344 | by(rule while_properties[THEN mp,OF _ _ _ _ _ _ _ _ refl]) | 
| 10496 | 345 | |
| 346 | lemma is_dfa_kildall: | |
| 347 | "[| semilat(A,r,f); acc r; pres_type step n A; | |
| 348 | mono r step n A; bounded succs n|] | |
| 11184 | 349 | ==> is_dfa r (kildall f step succs) step succs n A" | 
| 10496 | 350 | apply (unfold is_dfa_def kildall_def) | 
| 351 | apply clarify | |
| 352 | apply (rule iter_properties) | |
| 11175 
56ab6a5ba351
recoded function iter with the help of the while-combinator.
 nipkow parents: 
10774diff
changeset | 353 | apply (simp_all add: unstables_def stable_def) | 
| 10496 | 354 | apply (blast intro!: le_iff_plus_unchanged [THEN iffD2] listE_nth_in | 
| 355 | dest: boundedD pres_typeD) | |
| 356 | done | |
| 357 | ||
| 358 | lemma is_bcv_kildall: | |
| 359 | "[| semilat(A,r,f); acc r; top r T; | |
| 360 | pres_type step n A; bounded succs n; | |
| 361 | mono r step n A |] | |
| 11184 | 362 | ==> is_bcv r T step succs n A (kildall f step succs)" | 
| 10496 | 363 | apply (intro is_bcv_dfa is_dfa_kildall semilatDorderI) | 
| 364 | apply assumption+ | |
| 365 | done | |
| 366 | ||
| 367 | end |