| author | wenzelm | 
| Mon, 25 May 2009 12:46:14 +0200 | |
| changeset 31240 | 2c20bcd70fbe | 
| parent 27681 | 8cedebf55539 | 
| child 32443 | 16464c3f86bd | 
| permissions | -rw-r--r-- | 
| 9757 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 1 | (* Title: HOL/MicroJava/BV/LBVSpec.thy | 
| 8245 | 2 | ID: $Id$ | 
| 3 | Author: Gerwin Klein | |
| 4 | Copyright 1999 Technische Universitaet Muenchen | |
| 9054 | 5 | *) | 
| 8245 | 6 | |
| 12911 | 7 | header {* \isaheader{The Lightweight Bytecode Verifier} *}
 | 
| 9054 | 8 | |
| 27681 | 9 | theory LBVSpec | 
| 10 | imports SemilatAlg Opt | |
| 11 | begin | |
| 13062 | 12 | |
| 8245 | 13 | types | 
| 13078 | 14 | 's certificate = "'s list" | 
| 8245 | 15 | |
| 12516 | 16 | consts | 
| 13078 | 17 | merge :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> nat \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's \<Rightarrow> 's" | 
| 12516 | 18 | primrec | 
| 13078 | 19 | "merge cert f r T pc [] x = x" | 
| 20 | "merge cert f r T pc (s#ss) x = merge cert f r T pc ss (let (pc',s') = s in | |
| 21 | if pc'=pc+1 then s' +_f x | |
| 22 | else if s' <=_r (cert!pc') then x | |
| 23 | else T)" | |
| 9757 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 24 | |
| 13062 | 25 | constdefs | 
| 13078 | 26 | wtl_inst :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> | 
| 27 | 's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's" | |
| 28 | "wtl_inst cert f r T step pc s \<equiv> merge cert f r T pc (step pc s) (cert!(pc+1))" | |
| 8245 | 29 | |
| 13078 | 30 | wtl_cert :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> | 
| 31 | 's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's" | |
| 32 | "wtl_cert cert f r T B step pc s \<equiv> | |
| 33 | if cert!pc = B then | |
| 34 | wtl_inst cert f r T step pc s | |
| 35 | else | |
| 36 | if s <=_r (cert!pc) then wtl_inst cert f r T step pc (cert!pc) else T" | |
| 9757 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 37 | |
| 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 38 | consts | 
| 13078 | 39 | wtl_inst_list :: "'a list \<Rightarrow> 's certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> | 
| 40 | 's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's" | |
| 8245 | 41 | primrec | 
| 13101 | 42 | "wtl_inst_list [] cert f r T B step pc s = s" | 
| 43 | "wtl_inst_list (i#is) cert f r T B step pc s = | |
| 13078 | 44 | (let s' = wtl_cert cert f r T B step pc s in | 
| 13101 | 45 | if s' = T \<or> s = T then T else wtl_inst_list is cert f r T B step (pc+1) s')" | 
| 8245 | 46 | |
| 47 | constdefs | |
| 13078 | 48 | cert_ok :: "'s certificate \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> 's set \<Rightarrow> bool" | 
| 49 | "cert_ok cert n T B A \<equiv> (\<forall>i < n. cert!i \<in> A \<and> cert!i \<noteq> T) \<and> (cert!n = B)" | |
| 50 | ||
| 51 | constdefs | |
| 52 | bottom :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool" | |
| 53 | "bottom r B \<equiv> \<forall>x. B <=_r x" | |
| 54 | ||
| 55 | ||
| 27681 | 56 | locale lbv = Semilat + | 
| 13078 | 57 |   fixes T :: "'a" ("\<top>") 
 | 
| 58 |   fixes B :: "'a" ("\<bottom>") 
 | |
| 59 | fixes step :: "'a step_type" | |
| 60 | assumes top: "top r \<top>" | |
| 61 | assumes T_A: "\<top> \<in> A" | |
| 62 | assumes bot: "bottom r \<bottom>" | |
| 63 | assumes B_A: "\<bottom> \<in> A" | |
| 64 | ||
| 65 | fixes merge :: "'a certificate \<Rightarrow> nat \<Rightarrow> (nat \<times> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 66 | defines mrg_def: "merge cert \<equiv> LBVSpec.merge cert f r \<top>" | |
| 13062 | 67 | |
| 13078 | 68 | fixes wti :: "'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 69 | defines wti_def: "wti cert \<equiv> wtl_inst cert f r \<top> step" | |
| 70 | ||
| 71 | fixes wtc :: "'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 72 | defines wtc_def: "wtc cert \<equiv> wtl_cert cert f r \<top> \<bottom> step" | |
| 73 | ||
| 74 | fixes wtl :: "'b list \<Rightarrow> 'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 75 | defines wtl_def: "wtl ins cert \<equiv> wtl_inst_list ins cert f r \<top> \<bottom> step" | |
| 76 | ||
| 77 | ||
| 78 | lemma (in lbv) wti: | |
| 79 | "wti c pc s \<equiv> merge c pc (step pc s) (c!(pc+1))" | |
| 80 | by (simp add: wti_def mrg_def wtl_inst_def) | |
| 81 | ||
| 82 | lemma (in lbv) wtc: | |
| 83 | "wtc c pc s \<equiv> if c!pc = \<bottom> then wti c pc s else if s <=_r c!pc then wti c pc (c!pc) else \<top>" | |
| 84 | by (unfold wtc_def wti_def wtl_cert_def) | |
| 85 | ||
| 86 | ||
| 87 | lemma cert_okD1 [intro?]: | |
| 88 | "cert_ok c n T B A \<Longrightarrow> pc < n \<Longrightarrow> c!pc \<in> A" | |
| 13062 | 89 | by (unfold cert_ok_def) fast | 
| 90 | ||
| 13078 | 91 | lemma cert_okD2 [intro?]: | 
| 92 | "cert_ok c n T B A \<Longrightarrow> c!n = B" | |
| 13071 | 93 | by (simp add: cert_ok_def) | 
| 94 | ||
| 13078 | 95 | lemma cert_okD3 [intro?]: | 
| 96 | "cert_ok c n T B A \<Longrightarrow> B \<in> A \<Longrightarrow> pc < n \<Longrightarrow> c!Suc pc \<in> A" | |
| 13071 | 97 | by (drule Suc_leI) (auto simp add: le_eq_less_or_eq dest: cert_okD1 cert_okD2) | 
| 98 | ||
| 13078 | 99 | lemma cert_okD4 [intro?]: | 
| 100 | "cert_ok c n T B A \<Longrightarrow> pc < n \<Longrightarrow> c!pc \<noteq> T" | |
| 101 | by (simp add: cert_ok_def) | |
| 13062 | 102 | |
| 103 | declare Let_def [simp] | |
| 104 | ||
| 105 | section "more semilattice lemmas" | |
| 106 | ||
| 13078 | 107 | |
| 108 | lemma (in lbv) sup_top [simp, elim]: | |
| 109 | assumes x: "x \<in> A" | |
| 110 | shows "x +_f \<top> = \<top>" | |
| 13062 | 111 | proof - | 
| 13078 | 112 | from top have "x +_f \<top> <=_r \<top>" .. | 
| 23464 | 113 | moreover from x T_A have "\<top> <=_r x +_f \<top>" .. | 
| 13078 | 114 | ultimately show ?thesis .. | 
| 13062 | 115 | qed | 
| 116 | ||
| 13078 | 117 | lemma (in lbv) plusplussup_top [simp, elim]: | 
| 118 | "set xs \<subseteq> A \<Longrightarrow> xs ++_f \<top> = \<top>" | |
| 13062 | 119 | by (induct xs) auto | 
| 13078 | 120 | |
| 121 | ||
| 13062 | 122 | |
| 27681 | 123 | lemma (in Semilat) pp_ub1': | 
| 13078 | 124 | assumes S: "snd`set S \<subseteq> A" | 
| 125 | assumes y: "y \<in> A" and ab: "(a, b) \<in> set S" | |
| 23281 | 126 | shows "b <=_r map snd [(p', t') \<leftarrow> S . p' = a] ++_f y" | 
| 13078 | 127 | proof - | 
| 128 | from S have "\<forall>(x,y) \<in> set S. y \<in> A" by auto | |
| 129 | with semilat y ab show ?thesis by - (rule ub1') | |
| 130 | qed | |
| 13062 | 131 | |
| 13078 | 132 | lemma (in lbv) bottom_le [simp, intro]: | 
| 133 | "\<bottom> <=_r x" | |
| 134 | by (insert bot) (simp add: bottom_def) | |
| 13062 | 135 | |
| 13078 | 136 | lemma (in lbv) le_bottom [simp]: | 
| 137 | "x <=_r \<bottom> = (x = \<bottom>)" | |
| 138 | by (blast intro: antisym_r) | |
| 139 | ||
| 13062 | 140 | |
| 141 | ||
| 142 | section "merge" | |
| 143 | ||
| 13078 | 144 | lemma (in lbv) merge_Nil [simp]: | 
| 145 | "merge c pc [] x = x" by (simp add: mrg_def) | |
| 146 | ||
| 147 | lemma (in lbv) merge_Cons [simp]: | |
| 148 | "merge c pc (l#ls) x = merge c pc ls (if fst l=pc+1 then snd l +_f x | |
| 149 | else if snd l <=_r (c!fst l) then x | |
| 150 | else \<top>)" | |
| 151 | by (simp add: mrg_def split_beta) | |
| 152 | ||
| 153 | lemma (in lbv) merge_Err [simp]: | |
| 154 | "snd`set ss \<subseteq> A \<Longrightarrow> merge c pc ss \<top> = \<top>" | |
| 13062 | 155 | by (induct ss) auto | 
| 8245 | 156 | |
| 13078 | 157 | lemma (in lbv) merge_not_top: | 
| 158 | "\<And>x. snd`set ss \<subseteq> A \<Longrightarrow> merge c pc ss x \<noteq> \<top> \<Longrightarrow> | |
| 159 | \<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' <=_r (c!pc'))" | |
| 160 | (is "\<And>x. ?set ss \<Longrightarrow> ?merge ss x \<Longrightarrow> ?P ss") | |
| 13062 | 161 | proof (induct ss) | 
| 162 | show "?P []" by simp | |
| 163 | next | |
| 13078 | 164 | fix x ls l | 
| 165 | assume "?set (l#ls)" then obtain set: "snd`set ls \<subseteq> A" by simp | |
| 166 | assume merge: "?merge (l#ls) x" | |
| 13062 | 167 | moreover | 
| 23464 | 168 | obtain pc' s' where l: "l = (pc',s')" by (cases l) | 
| 13062 | 169 | ultimately | 
| 23464 | 170 | obtain x' where merge': "?merge ls x'" by simp | 
| 171 | assume "\<And>x. ?set ls \<Longrightarrow> ?merge ls x \<Longrightarrow> ?P ls" hence "?P ls" using set merge' . | |
| 13062 | 172 | moreover | 
| 13078 | 173 | from merge set | 
| 23464 | 174 | have "pc' \<noteq> pc+1 \<longrightarrow> s' <=_r (c!pc')" by (simp add: l split: split_if_asm) | 
| 13062 | 175 | ultimately | 
| 23464 | 176 | show "?P (l#ls)" by (simp add: l) | 
| 13062 | 177 | qed | 
| 178 | ||
| 9012 | 179 | |
| 13078 | 180 | lemma (in lbv) merge_def: | 
| 13062 | 181 | shows | 
| 13078 | 182 | "\<And>x. x \<in> A \<Longrightarrow> snd`set ss \<subseteq> A \<Longrightarrow> | 
| 183 | merge c pc ss x = | |
| 184 | (if \<forall>(pc',s') \<in> set ss. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc' then | |
| 23281 | 185 | map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x | 
| 13078 | 186 | else \<top>)" | 
| 13062 | 187 | (is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?merge ss x = ?if ss x" is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?P ss x") | 
| 188 | proof (induct ss) | |
| 189 | fix x show "?P [] x" by simp | |
| 190 | next | |
| 13078 | 191 | fix x assume x: "x \<in> A" | 
| 192 | fix l::"nat \<times> 'a" and ls | |
| 193 | assume "snd`set (l#ls) \<subseteq> A" | |
| 194 | then obtain l: "snd l \<in> A" and ls: "snd`set ls \<subseteq> A" by auto | |
| 195 | assume "\<And>x. x \<in> A \<Longrightarrow> snd`set ls \<subseteq> A \<Longrightarrow> ?P ls x" | |
| 23464 | 196 | hence IH: "\<And>x. x \<in> A \<Longrightarrow> ?P ls x" using ls by iprover | 
| 13062 | 197 | obtain pc' s' where [simp]: "l = (pc',s')" by (cases l) | 
| 198 | hence "?merge (l#ls) x = ?merge ls | |
| 13078 | 199 | (if pc'=pc+1 then s' +_f x else if s' <=_r c!pc' then x else \<top>)" | 
| 13062 | 200 | (is "?merge (l#ls) x = ?merge ls ?if'") | 
| 201 | by simp | |
| 202 | also have "\<dots> = ?if ls ?if'" | |
| 203 | proof - | |
| 13078 | 204 | from l have "s' \<in> A" by simp | 
| 205 | with x have "s' +_f x \<in> A" by simp | |
| 23464 | 206 | with x T_A have "?if' \<in> A" by auto | 
| 13062 | 207 | hence "?P ls ?if'" by (rule IH) thus ?thesis by simp | 
| 208 | qed | |
| 209 | also have "\<dots> = ?if (l#ls) x" | |
| 13078 | 210 | proof (cases "\<forall>(pc', s')\<in>set (l#ls). pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'") | 
| 13062 | 211 | case True | 
| 13078 | 212 | hence "\<forall>(pc', s')\<in>set ls. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" by auto | 
| 13062 | 213 | moreover | 
| 214 | from True have | |
| 23281 | 215 | "map snd [(p',t')\<leftarrow>ls . p'=pc+1] ++_f ?if' = | 
| 216 | (map snd [(p',t')\<leftarrow>l#ls . p'=pc+1] ++_f x)" | |
| 13062 | 217 | by simp | 
| 218 | ultimately | |
| 219 | show ?thesis using True by simp | |
| 220 | next | |
| 13078 | 221 | case False | 
| 222 | moreover | |
| 23281 | 223 | from ls have "set (map snd [(p', t')\<leftarrow>ls . p' = Suc pc]) \<subseteq> A" by auto | 
| 13078 | 224 | ultimately show ?thesis by auto | 
| 13062 | 225 | qed | 
| 226 | finally show "?P (l#ls) x" . | |
| 227 | qed | |
| 228 | ||
| 13078 | 229 | lemma (in lbv) merge_not_top_s: | 
| 230 | assumes x: "x \<in> A" and ss: "snd`set ss \<subseteq> A" | |
| 231 | assumes m: "merge c pc ss x \<noteq> \<top>" | |
| 23281 | 232 | shows "merge c pc ss x = (map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x)" | 
| 13062 | 233 | proof - | 
| 13078 | 234 | from ss m have "\<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' <=_r c!pc')" | 
| 235 | by (rule merge_not_top) | |
| 236 | with x ss m show ?thesis by - (drule merge_def, auto split: split_if_asm) | |
| 13062 | 237 | qed | 
| 238 | ||
| 239 | section "wtl-inst-list" | |
| 10042 | 240 | |
| 10812 
ead84e90bfeb
merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
 kleing parents: 
10638diff
changeset | 241 | lemmas [iff] = not_Err_eq | 
| 
ead84e90bfeb
merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
 kleing parents: 
10638diff
changeset | 242 | |
| 13078 | 243 | lemma (in lbv) wtl_Nil [simp]: "wtl [] c pc s = s" | 
| 244 | by (simp add: wtl_def) | |
| 245 | ||
| 246 | lemma (in lbv) wtl_Cons [simp]: | |
| 247 | "wtl (i#is) c pc s = | |
| 248 | (let s' = wtc c pc s in if s' = \<top> \<or> s = \<top> then \<top> else wtl is c (pc+1) s')" | |
| 249 | by (simp add: wtl_def wtc_def) | |
| 250 | ||
| 251 | lemma (in lbv) wtl_Cons_not_top: | |
| 252 | "wtl (i#is) c pc s \<noteq> \<top> = | |
| 253 | (wtc c pc s \<noteq> \<top> \<and> s \<noteq> T \<and> wtl is c (pc+1) (wtc c pc s) \<noteq> \<top>)" | |
| 10812 
ead84e90bfeb
merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
 kleing parents: 
10638diff
changeset | 254 | by (auto simp del: split_paired_Ex) | 
| 9549 
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
 kleing parents: 
9376diff
changeset | 255 | |
| 13078 | 256 | lemma (in lbv) wtl_top [simp]: "wtl ls c pc \<top> = \<top>" | 
| 257 | by (cases ls) auto | |
| 258 | ||
| 259 | lemma (in lbv) wtl_not_top: | |
| 260 | "wtl ls c pc s \<noteq> \<top> \<Longrightarrow> s \<noteq> \<top>" | |
| 261 | by (cases "s=\<top>") auto | |
| 9012 | 262 | |
| 13078 | 263 | lemma (in lbv) wtl_append [simp]: | 
| 264 | "\<And>pc s. wtl (a@b) c pc s = wtl b c (pc+length a) (wtl a c pc s)" | |
| 265 | by (induct a) auto | |
| 266 | ||
| 267 | lemma (in lbv) wtl_take: | |
| 268 | "wtl is c pc s \<noteq> \<top> \<Longrightarrow> wtl (take pc' is) c pc s \<noteq> \<top>" | |
| 269 | (is "?wtl is \<noteq> _ \<Longrightarrow> _") | |
| 9183 | 270 | proof - | 
| 13078 | 271 | assume "?wtl is \<noteq> \<top>" | 
| 272 | hence "?wtl (take pc' is @ drop pc' is) \<noteq> \<top>" by simp | |
| 273 | thus ?thesis by (auto dest!: wtl_not_top simp del: append_take_drop_id) | |
| 9757 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 274 | qed | 
| 9012 | 275 | |
| 9757 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 276 | lemma take_Suc: | 
| 13078 | 277 | "\<forall>n. n < length l \<longrightarrow> take (Suc n) l = (take n l)@[l!n]" (is "?P l") | 
| 9757 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 278 | proof (induct l) | 
| 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 279 | show "?P []" by simp | 
| 12516 | 280 | next | 
| 281 | fix x xs assume IH: "?P xs" | |
| 9757 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 282 | show "?P (x#xs)" | 
| 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 283 | proof (intro strip) | 
| 12516 | 284 | fix n assume "n < length (x#xs)" | 
| 285 | with IH show "take (Suc n) (x # xs) = take n (x # xs) @ [(x # xs) ! n]" | |
| 286 | by (cases n, auto) | |
| 9183 | 287 | qed | 
| 288 | qed | |
| 9012 | 289 | |
| 13078 | 290 | lemma (in lbv) wtl_Suc: | 
| 291 | assumes suc: "pc+1 < length is" | |
| 292 | assumes wtl: "wtl (take pc is) c 0 s \<noteq> \<top>" | |
| 293 | shows "wtl (take (pc+1) is) c 0 s = wtc c pc (wtl (take pc is) c 0 s)" | |
| 9183 | 294 | proof - | 
| 13062 | 295 | from suc have "take (pc+1) is=(take pc is)@[is!pc]" by (simp add: take_Suc) | 
| 13078 | 296 | with suc wtl show ?thesis by (simp add: min_def) | 
| 9183 | 297 | qed | 
| 9012 | 298 | |
| 13078 | 299 | lemma (in lbv) wtl_all: | 
| 300 | assumes all: "wtl is c 0 s \<noteq> \<top>" (is "?wtl is \<noteq> _") | |
| 301 | assumes pc: "pc < length is" | |
| 302 | shows "wtc c pc (wtl (take pc is) c 0 s) \<noteq> \<top>" | |
| 9183 | 303 | proof - | 
| 13062 | 304 | from pc have "0 < length (drop pc is)" by simp | 
| 12516 | 305 | then obtain i r where Cons: "drop pc is = i#r" | 
| 15109 | 306 | by (auto simp add: neq_Nil_conv simp del: length_drop drop_eq_Nil) | 
| 9757 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9664diff
changeset | 307 | hence "i#r = drop pc is" .. | 
| 13078 | 308 | with all have take: "?wtl (take pc is@i#r) \<noteq> \<top>" by simp | 
| 12516 | 309 | from pc have "is!pc = drop pc is ! 0" by simp | 
| 310 | with Cons have "is!pc = i" by simp | |
| 13078 | 311 | with take pc show ?thesis by (auto simp add: min_def split: split_if_asm) | 
| 9183 | 312 | qed | 
| 9012 | 313 | |
| 13062 | 314 | section "preserves-type" | 
| 315 | ||
| 13078 | 316 | lemma (in lbv) merge_pres: | 
| 317 | assumes s0: "snd`set ss \<subseteq> A" and x: "x \<in> A" | |
| 318 | shows "merge c pc ss x \<in> A" | |
| 13062 | 319 | proof - | 
| 23281 | 320 | from s0 have "set (map snd [(p', t')\<leftarrow>ss . p'=pc+1]) \<subseteq> A" by auto | 
| 321 | with x have "(map snd [(p', t')\<leftarrow>ss . p'=pc+1] ++_f x) \<in> A" | |
| 23464 | 322 | by (auto intro!: plusplus_closed semilat) | 
| 13078 | 323 | with s0 x show ?thesis by (simp add: merge_def T_A) | 
| 13062 | 324 | qed | 
| 325 | ||
| 326 | ||
| 13078 | 327 | lemma pres_typeD2: | 
| 328 | "pres_type step n A \<Longrightarrow> s \<in> A \<Longrightarrow> p < n \<Longrightarrow> snd`set (step p s) \<subseteq> A" | |
| 329 | by auto (drule pres_typeD) | |
| 13062 | 330 | |
| 13078 | 331 | |
| 332 | lemma (in lbv) wti_pres [intro?]: | |
| 333 | assumes pres: "pres_type step n A" | |
| 334 | assumes cert: "c!(pc+1) \<in> A" | |
| 335 | assumes s_pc: "s \<in> A" "pc < n" | |
| 336 | shows "wti c pc s \<in> A" | |
| 13062 | 337 | proof - | 
| 13078 | 338 | from pres s_pc have "snd`set (step pc s) \<subseteq> A" by (rule pres_typeD2) | 
| 339 | with cert show ?thesis by (simp add: wti merge_pres) | |
| 13062 | 340 | qed | 
| 341 | ||
| 342 | ||
| 13078 | 343 | lemma (in lbv) wtc_pres: | 
| 23464 | 344 | assumes pres: "pres_type step n A" | 
| 345 | assumes cert: "c!pc \<in> A" and cert': "c!(pc+1) \<in> A" | |
| 346 | assumes s: "s \<in> A" and pc: "pc < n" | |
| 13078 | 347 | shows "wtc c pc s \<in> A" | 
| 13062 | 348 | proof - | 
| 23464 | 349 | have "wti c pc s \<in> A" using pres cert' s pc .. | 
| 350 | moreover have "wti c pc (c!pc) \<in> A" using pres cert' cert pc .. | |
| 13078 | 351 | ultimately show ?thesis using T_A by (simp add: wtc) | 
| 13062 | 352 | qed | 
| 353 | ||
| 13078 | 354 | |
| 355 | lemma (in lbv) wtl_pres: | |
| 356 | assumes pres: "pres_type step (length is) A" | |
| 357 | assumes cert: "cert_ok c (length is) \<top> \<bottom> A" | |
| 358 | assumes s: "s \<in> A" | |
| 359 | assumes all: "wtl is c 0 s \<noteq> \<top>" | |
| 360 | shows "pc < length is \<Longrightarrow> wtl (take pc is) c 0 s \<in> A" | |
| 361 | (is "?len pc \<Longrightarrow> ?wtl pc \<in> A") | |
| 13062 | 362 | proof (induct pc) | 
| 13078 | 363 | from s show "?wtl 0 \<in> A" by simp | 
| 13062 | 364 | next | 
| 23464 | 365 | fix n assume IH: "Suc n < length is" | 
| 366 | then have n: "n < length is" by simp | |
| 367 | from IH have n1: "n+1 < length is" by simp | |
| 368 | assume prem: "n < length is \<Longrightarrow> ?wtl n \<in> A" | |
| 369 | have "wtc c n (?wtl n) \<in> A" | |
| 370 | using pres _ _ _ n | |
| 371 | proof (rule wtc_pres) | |
| 372 | from prem n show "?wtl n \<in> A" . | |
| 373 | from cert n show "c!n \<in> A" by (rule cert_okD1) | |
| 374 | from cert n1 show "c!(n+1) \<in> A" by (rule cert_okD1) | |
| 375 | qed | |
| 13078 | 376 | also | 
| 377 | from all n have "?wtl n \<noteq> \<top>" by - (rule wtl_take) | |
| 378 | with n1 have "wtc c n (?wtl n) = ?wtl (n+1)" by (rule wtl_Suc [symmetric]) | |
| 379 | finally show "?wtl (Suc n) \<in> A" by simp | |
| 13062 | 380 | qed | 
| 381 | ||
| 9183 | 382 | end |