| author | haftmann | 
| Tue, 05 Jan 2010 11:38:51 +0100 | |
| changeset 34271 | 70af06abb13d | 
| parent 33954 | 1bc3b688548c | 
| child 35416 | d8d7d1b785af | 
| permissions | -rwxr-xr-x | 
| 12516 | 1 | (* Title: HOL/MicroJava/BV/Effect.thy | 
| 2 | Author: Gerwin Klein | |
| 3 | Copyright 2000 Technische Universitaet Muenchen | |
| 4 | *) | |
| 5 | ||
| 12911 | 6 | header {* \isaheader{Effect of Instructions on the State Type} *}
 | 
| 12516 | 7 | |
| 15481 | 8 | theory Effect | 
| 9 | imports JVMType "../JVM/JVMExceptions" | |
| 10 | begin | |
| 11 | ||
| 12516 | 12 | types | 
| 13 | succ_type = "(p_count \<times> state_type option) list" | |
| 14 | ||
| 15 | text {* Program counter of successor instructions: *}
 | |
| 16 | consts | |
| 13006 | 17 | succs :: "instr \<Rightarrow> p_count \<Rightarrow> p_count list" | 
| 12516 | 18 | primrec | 
| 19 | "succs (Load idx) pc = [pc+1]" | |
| 20 | "succs (Store idx) pc = [pc+1]" | |
| 21 | "succs (LitPush v) pc = [pc+1]" | |
| 22 | "succs (Getfield F C) pc = [pc+1]" | |
| 23 | "succs (Putfield F C) pc = [pc+1]" | |
| 24 | "succs (New C) pc = [pc+1]" | |
| 25 | "succs (Checkcast C) pc = [pc+1]" | |
| 26 | "succs Pop pc = [pc+1]" | |
| 27 | "succs Dup pc = [pc+1]" | |
| 28 | "succs Dup_x1 pc = [pc+1]" | |
| 29 | "succs Dup_x2 pc = [pc+1]" | |
| 30 | "succs Swap pc = [pc+1]" | |
| 31 | "succs IAdd pc = [pc+1]" | |
| 32 | "succs (Ifcmpeq b) pc = [pc+1, nat (int pc + b)]" | |
| 33 | "succs (Goto b) pc = [nat (int pc + b)]" | |
| 34 | "succs Return pc = [pc]" | |
| 35 | "succs (Invoke C mn fpTs) pc = [pc+1]" | |
| 36 | "succs Throw pc = [pc]" | |
| 37 | ||
| 38 | text "Effect of instruction on the state type:" | |
| 39 | consts | |
| 13006 | 40 | eff' :: "instr \<times> jvm_prog \<times> state_type \<Rightarrow> state_type" | 
| 12516 | 41 | |
| 42 | recdef eff' "{}"
 | |
| 43 | "eff' (Load idx, G, (ST, LT)) = (ok_val (LT ! idx) # ST, LT)" | |
| 44 | "eff' (Store idx, G, (ts#ST, LT)) = (ST, LT[idx:= OK ts])" | |
| 45 | "eff' (LitPush v, G, (ST, LT)) = (the (typeof (\<lambda>v. None) v) # ST, LT)" | |
| 46 | "eff' (Getfield F C, G, (oT#ST, LT)) = (snd (the (field (G,C) F)) # ST, LT)" | |
| 47 | "eff' (Putfield F C, G, (vT#oT#ST, LT)) = (ST,LT)" | |
| 48 | "eff' (New C, G, (ST,LT)) = (Class C # ST, LT)" | |
| 49 | "eff' (Checkcast C, G, (RefT rt#ST,LT)) = (Class C # ST,LT)" | |
| 50 | "eff' (Pop, G, (ts#ST,LT)) = (ST,LT)" | |
| 51 | "eff' (Dup, G, (ts#ST,LT)) = (ts#ts#ST,LT)" | |
| 52 | "eff' (Dup_x1, G, (ts1#ts2#ST,LT)) = (ts1#ts2#ts1#ST,LT)" | |
| 53 | "eff' (Dup_x2, G, (ts1#ts2#ts3#ST,LT)) = (ts1#ts2#ts3#ts1#ST,LT)" | |
| 54 | "eff' (Swap, G, (ts1#ts2#ST,LT)) = (ts2#ts1#ST,LT)" | |
| 55 | "eff' (IAdd, G, (PrimT Integer#PrimT Integer#ST,LT)) | |
| 56 | = (PrimT Integer#ST,LT)" | |
| 57 | "eff' (Ifcmpeq b, G, (ts1#ts2#ST,LT)) = (ST,LT)" | |
| 58 | "eff' (Goto b, G, s) = s" | |
| 59 | -- "Return has no successor instruction in the same method" | |
| 60 | "eff' (Return, G, s) = s" | |
| 61 | -- "Throw always terminates abruptly" | |
| 62 | "eff' (Throw, G, s) = s" | |
| 63 | "eff' (Invoke C mn fpTs, G, (ST,LT)) = (let ST' = drop (length fpTs) ST | |
| 64 | in (fst (snd (the (method (G,C) (mn,fpTs))))#(tl ST'),LT))" | |
| 65 | ||
| 66 | ||
| 67 | consts | |
| 68 | match_any :: "jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list" | |
| 69 | primrec | |
| 70 | "match_any G pc [] = []" | |
| 71 | "match_any G pc (e#es) = (let (start_pc, end_pc, handler_pc, catch_type) = e; | |
| 72 | es' = match_any G pc es | |
| 73 | in | |
| 74 | if start_pc <= pc \<and> pc < end_pc then catch_type#es' else es')" | |
| 75 | ||
| 12951 | 76 | consts | 
| 13717 | 77 | match :: "jvm_prog \<Rightarrow> xcpt \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list" | 
| 12951 | 78 | primrec | 
| 79 | "match G X pc [] = []" | |
| 80 | "match G X pc (e#es) = | |
| 13717 | 81 | (if match_exception_entry G (Xcpt X) pc e then [Xcpt X] else match G X pc es)" | 
| 12951 | 82 | |
| 83 | lemma match_some_entry: | |
| 13717 | 84 | "match G X pc et = (if \<exists>e \<in> set et. match_exception_entry G (Xcpt X) pc e then [Xcpt X] else [])" | 
| 12951 | 85 | by (induct et) auto | 
| 12516 | 86 | |
| 87 | consts | |
| 13006 | 88 | xcpt_names :: "instr \<times> jvm_prog \<times> p_count \<times> exception_table \<Rightarrow> cname list" | 
| 12516 | 89 | recdef xcpt_names "{}"
 | 
| 13717 | 90 | "xcpt_names (Getfield F C, G, pc, et) = match G NullPointer pc et" | 
| 91 | "xcpt_names (Putfield F C, G, pc, et) = match G NullPointer pc et" | |
| 92 | "xcpt_names (New C, G, pc, et) = match G OutOfMemory pc et" | |
| 93 | "xcpt_names (Checkcast C, G, pc, et) = match G ClassCast pc et" | |
| 12516 | 94 | "xcpt_names (Throw, G, pc, et) = match_any G pc et" | 
| 95 | "xcpt_names (Invoke C m p, G, pc, et) = match_any G pc et" | |
| 96 | "xcpt_names (i, G, pc, et) = []" | |
| 97 | ||
| 98 | ||
| 99 | constdefs | |
| 100 | xcpt_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> state_type option \<Rightarrow> exception_table \<Rightarrow> succ_type" | |
| 101 | "xcpt_eff i G pc s et == | |
| 102 | map (\<lambda>C. (the (match_exception_table G C pc et), case s of None \<Rightarrow> None | Some s' \<Rightarrow> Some ([Class C], snd s'))) | |
| 103 | (xcpt_names (i,G,pc,et))" | |
| 104 | ||
| 105 | norm_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> state_type option \<Rightarrow> state_type option" | |
| 30235 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 nipkow parents: 
25362diff
changeset | 106 | "norm_eff i G == Option.map (\<lambda>s. eff' (i,G,s))" | 
| 12516 | 107 | |
| 13006 | 108 | eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> succ_type" | 
| 12516 | 109 | "eff i G pc et s == (map (\<lambda>pc'. (pc',norm_eff i G s)) (succs i pc)) @ (xcpt_eff i G pc s et)" | 
| 110 | ||
| 12772 | 111 | constdefs | 
| 112 | isPrimT :: "ty \<Rightarrow> bool" | |
| 113 | "isPrimT T == case T of PrimT T' \<Rightarrow> True | RefT T' \<Rightarrow> False" | |
| 114 | ||
| 115 | isRefT :: "ty \<Rightarrow> bool" | |
| 116 | "isRefT T == case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> True" | |
| 117 | ||
| 118 | lemma isPrimT [simp]: | |
| 119 | "isPrimT T = (\<exists>T'. T = PrimT T')" by (simp add: isPrimT_def split: ty.splits) | |
| 120 | ||
| 121 | lemma isRefT [simp]: | |
| 122 | "isRefT T = (\<exists>T'. T = RefT T')" by (simp add: isRefT_def split: ty.splits) | |
| 123 | ||
| 124 | ||
| 125 | lemma "list_all2 P a b \<Longrightarrow> \<forall>(x,y) \<in> set (zip a b). P x y" | |
| 126 | by (simp add: list_all2_def) | |
| 127 | ||
| 12516 | 128 | |
| 129 | text "Conditions under which eff is applicable:" | |
| 130 | consts | |
| 13006 | 131 | app' :: "instr \<times> jvm_prog \<times> p_count \<times> nat \<times> ty \<times> state_type \<Rightarrow> bool" | 
| 12516 | 132 | |
| 133 | recdef app' "{}"
 | |
| 12974 | 134 | "app' (Load idx, G, pc, maxs, rT, s) = | 
| 135 | (idx < length (snd s) \<and> (snd s) ! idx \<noteq> Err \<and> length (fst s) < maxs)" | |
| 136 | "app' (Store idx, G, pc, maxs, rT, (ts#ST, LT)) = | |
| 137 | (idx < length LT)" | |
| 138 | "app' (LitPush v, G, pc, maxs, rT, s) = | |
| 139 | (length (fst s) < maxs \<and> typeof (\<lambda>t. None) v \<noteq> None)" | |
| 140 | "app' (Getfield F C, G, pc, maxs, rT, (oT#ST, LT)) = | |
| 141 | (is_class G C \<and> field (G,C) F \<noteq> None \<and> fst (the (field (G,C) F)) = C \<and> | |
| 142 | G \<turnstile> oT \<preceq> (Class C))" | |
| 143 | "app' (Putfield F C, G, pc, maxs, rT, (vT#oT#ST, LT)) = | |
| 144 | (is_class G C \<and> field (G,C) F \<noteq> None \<and> fst (the (field (G,C) F)) = C \<and> | |
| 145 | G \<turnstile> oT \<preceq> (Class C) \<and> G \<turnstile> vT \<preceq> (snd (the (field (G,C) F))))" | |
| 146 | "app' (New C, G, pc, maxs, rT, s) = | |
| 147 | (is_class G C \<and> length (fst s) < maxs)" | |
| 148 | "app' (Checkcast C, G, pc, maxs, rT, (RefT rt#ST,LT)) = | |
| 149 | (is_class G C)" | |
| 150 | "app' (Pop, G, pc, maxs, rT, (ts#ST,LT)) = | |
| 151 | True" | |
| 152 | "app' (Dup, G, pc, maxs, rT, (ts#ST,LT)) = | |
| 153 | (1+length ST < maxs)" | |
| 154 | "app' (Dup_x1, G, pc, maxs, rT, (ts1#ts2#ST,LT)) = | |
| 155 | (2+length ST < maxs)" | |
| 156 | "app' (Dup_x2, G, pc, maxs, rT, (ts1#ts2#ts3#ST,LT)) = | |
| 157 | (3+length ST < maxs)" | |
| 158 | "app' (Swap, G, pc, maxs, rT, (ts1#ts2#ST,LT)) = | |
| 159 | True" | |
| 160 | "app' (IAdd, G, pc, maxs, rT, (PrimT Integer#PrimT Integer#ST,LT)) = | |
| 161 | True" | |
| 162 | "app' (Ifcmpeq b, G, pc, maxs, rT, (ts#ts'#ST,LT)) = | |
| 163 | (0 \<le> int pc + b \<and> (isPrimT ts \<and> ts' = ts \<or> isRefT ts \<and> isRefT ts'))" | |
| 164 | "app' (Goto b, G, pc, maxs, rT, s) = | |
| 165 | (0 \<le> int pc + b)" | |
| 166 | "app' (Return, G, pc, maxs, rT, (T#ST,LT)) = | |
| 167 | (G \<turnstile> T \<preceq> rT)" | |
| 168 | "app' (Throw, G, pc, maxs, rT, (T#ST,LT)) = | |
| 169 | isRefT T" | |
| 170 | "app' (Invoke C mn fpTs, G, pc, maxs, rT, s) = | |
| 171 | (length fpTs < length (fst s) \<and> | |
| 172 | (let apTs = rev (take (length fpTs) (fst s)); | |
| 173 | X = hd (drop (length fpTs) (fst s)) | |
| 174 | in | |
| 175 | G \<turnstile> X \<preceq> Class C \<and> is_class G C \<and> method (G,C) (mn,fpTs) \<noteq> None \<and> | |
| 176 | list_all2 (\<lambda>x y. G \<turnstile> x \<preceq> y) apTs fpTs))" | |
| 12772 | 177 | |
| 12974 | 178 | "app' (i,G, pc,maxs,rT,s) = False" | 
| 12516 | 179 | |
| 180 | constdefs | |
| 181 | xcpt_app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> bool" | |
| 182 | "xcpt_app i G pc et \<equiv> \<forall>C\<in>set(xcpt_names (i,G,pc,et)). is_class G C" | |
| 183 | ||
| 13006 | 184 | app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> bool" | 
| 185 | "app i G maxs rT pc et s == case s of None \<Rightarrow> True | Some t \<Rightarrow> app' (i,G,pc,maxs,rT,t) \<and> xcpt_app i G pc et" | |
| 12516 | 186 | |
| 187 | ||
| 13066 | 188 | lemma match_any_match_table: | 
| 189 | "C \<in> set (match_any G pc et) \<Longrightarrow> match_exception_table G C pc et \<noteq> None" | |
| 190 | apply (induct et) | |
| 191 | apply simp | |
| 192 | apply simp | |
| 193 | apply clarify | |
| 194 | apply (simp split: split_if_asm) | |
| 195 | apply (auto simp add: match_exception_entry_def) | |
| 196 | done | |
| 197 | ||
| 198 | lemma match_X_match_table: | |
| 199 | "C \<in> set (match G X pc et) \<Longrightarrow> match_exception_table G C pc et \<noteq> None" | |
| 200 | apply (induct et) | |
| 201 | apply simp | |
| 202 | apply (simp split: split_if_asm) | |
| 203 | done | |
| 204 | ||
| 205 | lemma xcpt_names_in_et: | |
| 206 | "C \<in> set (xcpt_names (i,G,pc,et)) \<Longrightarrow> | |
| 207 | \<exists>e \<in> set et. the (match_exception_table G C pc et) = fst (snd (snd e))" | |
| 208 | apply (cases i) | |
| 18576 | 209 | apply (auto dest!: match_any_match_table match_X_match_table | 
| 13066 | 210 | dest: match_exception_table_in_et) | 
| 211 | done | |
| 212 | ||
| 213 | ||
| 13006 | 214 | lemma 1: "2 < length a \<Longrightarrow> (\<exists>l l' l'' ls. a = l#l'#l''#ls)" | 
| 12516 | 215 | proof (cases a) | 
| 216 | fix x xs assume "a = x#xs" "2 < length a" | |
| 217 | thus ?thesis by - (cases xs, simp, cases "tl xs", auto) | |
| 218 | qed auto | |
| 219 | ||
| 13006 | 220 | lemma 2: "\<not>(2 < length a) \<Longrightarrow> a = [] \<or> (\<exists> l. a = [l]) \<or> (\<exists> l l'. a = [l,l'])" | 
| 12516 | 221 | proof -; | 
| 222 | assume "\<not>(2 < length a)" | |
| 223 | hence "length a < (Suc (Suc (Suc 0)))" by simp | |
| 224 | hence * : "length a = 0 \<or> length a = Suc 0 \<or> length a = Suc (Suc 0)" | |
| 225 | by (auto simp add: less_Suc_eq) | |
| 226 | ||
| 227 |   { 
 | |
| 228 | fix x | |
| 229 | assume "length x = Suc 0" | |
| 230 | hence "\<exists> l. x = [l]" by - (cases x, auto) | |
| 231 | } note 0 = this | |
| 232 | ||
| 13006 | 233 | have "length a = Suc (Suc 0) \<Longrightarrow> \<exists>l l'. a = [l,l']" by (cases a, auto dest: 0) | 
| 12516 | 234 | with * show ?thesis by (auto dest: 0) | 
| 235 | qed | |
| 236 | ||
| 237 | lemmas [simp] = app_def xcpt_app_def | |
| 238 | ||
| 239 | text {* 
 | |
| 240 | \medskip | |
| 241 | simp rules for @{term app}
 | |
| 242 | *} | |
| 243 | lemma appNone[simp]: "app i G maxs rT pc et None = True" by simp | |
| 244 | ||
| 245 | ||
| 246 | lemma appLoad[simp]: | |
| 247 | "(app (Load idx) G maxs rT pc et (Some s)) = (\<exists>ST LT. s = (ST,LT) \<and> idx < length LT \<and> LT!idx \<noteq> Err \<and> length ST < maxs)" | |
| 248 | by (cases s, simp) | |
| 249 | ||
| 250 | lemma appStore[simp]: | |
| 251 | "(app (Store idx) G maxs rT pc et (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT) \<and> idx < length LT)" | |
| 252 | by (cases s, cases "2 < length (fst s)", auto dest: 1 2) | |
| 253 | ||
| 254 | lemma appLitPush[simp]: | |
| 255 | "(app (LitPush v) G maxs rT pc et (Some s)) = (\<exists>ST LT. s = (ST,LT) \<and> length ST < maxs \<and> typeof (\<lambda>v. None) v \<noteq> None)" | |
| 256 | by (cases s, simp) | |
| 257 | ||
| 258 | lemma appGetField[simp]: | |
| 259 | "(app (Getfield F C) G maxs rT pc et (Some s)) = | |
| 260 | (\<exists> oT vT ST LT. s = (oT#ST, LT) \<and> is_class G C \<and> | |
| 13717 | 261 | field (G,C) F = Some (C,vT) \<and> G \<turnstile> oT \<preceq> (Class C) \<and> (\<forall>x \<in> set (match G NullPointer pc et). is_class G x))" | 
| 12516 | 262 | by (cases s, cases "2 <length (fst s)", auto dest!: 1 2) | 
| 263 | ||
| 264 | lemma appPutField[simp]: | |
| 265 | "(app (Putfield F C) G maxs rT pc et (Some s)) = | |
| 266 | (\<exists> vT vT' oT ST LT. s = (vT#oT#ST, LT) \<and> is_class G C \<and> | |
| 12951 | 267 | field (G,C) F = Some (C, vT') \<and> G \<turnstile> oT \<preceq> (Class C) \<and> G \<turnstile> vT \<preceq> vT' \<and> | 
| 13717 | 268 | (\<forall>x \<in> set (match G NullPointer pc et). is_class G x))" | 
| 12516 | 269 | by (cases s, cases "2 <length (fst s)", auto dest!: 1 2) | 
| 270 | ||
| 271 | lemma appNew[simp]: | |
| 272 | "(app (New C) G maxs rT pc et (Some s)) = | |
| 12951 | 273 | (\<exists>ST LT. s=(ST,LT) \<and> is_class G C \<and> length ST < maxs \<and> | 
| 13717 | 274 | (\<forall>x \<in> set (match G OutOfMemory pc et). is_class G x))" | 
| 12516 | 275 | by (cases s, simp) | 
| 276 | ||
| 277 | lemma appCheckcast[simp]: | |
| 278 | "(app (Checkcast C) G maxs rT pc et (Some s)) = | |
| 12951 | 279 | (\<exists>rT ST LT. s = (RefT rT#ST,LT) \<and> is_class G C \<and> | 
| 13717 | 280 | (\<forall>x \<in> set (match G ClassCast pc et). is_class G x))" | 
| 12516 | 281 | by (cases s, cases "fst s", simp add: app_def) (cases "hd (fst s)", auto) | 
| 282 | ||
| 283 | lemma appPop[simp]: | |
| 284 | "(app Pop G maxs rT pc et (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT))" | |
| 285 | by (cases s, cases "2 <length (fst s)", auto dest: 1 2) | |
| 286 | ||
| 287 | ||
| 288 | lemma appDup[simp]: | |
| 289 | "(app Dup G maxs rT pc et (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT) \<and> 1+length ST < maxs)" | |
| 290 | by (cases s, cases "2 <length (fst s)", auto dest: 1 2) | |
| 291 | ||
| 292 | ||
| 293 | lemma appDup_x1[simp]: | |
| 294 | "(app Dup_x1 G maxs rT pc et (Some s)) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT) \<and> 2+length ST < maxs)" | |
| 295 | by (cases s, cases "2 <length (fst s)", auto dest: 1 2) | |
| 296 | ||
| 297 | ||
| 298 | lemma appDup_x2[simp]: | |
| 299 | "(app Dup_x2 G maxs rT pc et (Some s)) = (\<exists>ts1 ts2 ts3 ST LT. s = (ts1#ts2#ts3#ST,LT) \<and> 3+length ST < maxs)" | |
| 300 | by (cases s, cases "2 <length (fst s)", auto dest: 1 2) | |
| 301 | ||
| 302 | ||
| 303 | lemma appSwap[simp]: | |
| 304 | "app Swap G maxs rT pc et (Some s) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT))" | |
| 25362 | 305 | by (cases s, cases "2 <length (fst s)") (auto dest: 1 2) | 
| 12516 | 306 | |
| 307 | ||
| 308 | lemma appIAdd[simp]: | |
| 309 | "app IAdd G maxs rT pc et (Some s) = (\<exists> ST LT. s = (PrimT Integer#PrimT Integer#ST,LT))" | |
| 310 | (is "?app s = ?P s") | |
| 25362 | 311 | proof (cases s) | 
| 312 | case (Pair a b) | |
| 12516 | 313 | have "?app (a,b) = ?P (a,b)" | 
| 25362 | 314 | proof (cases a) | 
| 12516 | 315 | fix t ts assume a: "a = t#ts" | 
| 316 | show ?thesis | |
| 317 | proof (cases t) | |
| 318 | fix p assume p: "t = PrimT p" | |
| 319 | show ?thesis | |
| 320 | proof (cases p) | |
| 321 | assume ip: "p = Integer" | |
| 322 | show ?thesis | |
| 323 | proof (cases ts) | |
| 324 | fix t' ts' assume t': "ts = t' # ts'" | |
| 325 | show ?thesis | |
| 326 | proof (cases t') | |
| 327 | fix p' assume "t' = PrimT p'" | |
| 328 | with t' ip p a | |
| 25362 | 329 | show ?thesis by (cases p') auto | 
| 12516 | 330 | qed (auto simp add: a p ip t') | 
| 331 | qed (auto simp add: a p ip) | |
| 332 | qed (auto simp add: a p) | |
| 333 | qed (auto simp add: a) | |
| 334 | qed auto | |
| 335 | with Pair show ?thesis by simp | |
| 336 | qed | |
| 337 | ||
| 338 | ||
| 339 | lemma appIfcmpeq[simp]: | |
| 12974 | 340 | "app (Ifcmpeq b) G maxs rT pc et (Some s) = | 
| 341 | (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT) \<and> 0 \<le> int pc + b \<and> | |
| 342 | ((\<exists> p. ts1 = PrimT p \<and> ts2 = PrimT p) \<or> (\<exists>r r'. ts1 = RefT r \<and> ts2 = RefT r')))" | |
| 12772 | 343 | by (cases s, cases "2 <length (fst s)", auto dest!: 1 2) | 
| 12516 | 344 | |
| 345 | lemma appReturn[simp]: | |
| 346 | "app Return G maxs rT pc et (Some s) = (\<exists>T ST LT. s = (T#ST,LT) \<and> (G \<turnstile> T \<preceq> rT))" | |
| 347 | by (cases s, cases "2 <length (fst s)", auto dest: 1 2) | |
| 348 | ||
| 349 | lemma appGoto[simp]: | |
| 12974 | 350 | "app (Goto b) G maxs rT pc et (Some s) = (0 \<le> int pc + b)" | 
| 12516 | 351 | by simp | 
| 352 | ||
| 353 | lemma appThrow[simp]: | |
| 354 | "app Throw G maxs rT pc et (Some s) = | |
| 355 | (\<exists>T ST LT r. s=(T#ST,LT) \<and> T = RefT r \<and> (\<forall>C \<in> set (match_any G pc et). is_class G C))" | |
| 356 | by (cases s, cases "2 < length (fst s)", auto dest: 1 2) | |
| 357 | ||
| 358 | lemma appInvoke[simp]: | |
| 359 | "app (Invoke C mn fpTs) G maxs rT pc et (Some s) = (\<exists>apTs X ST LT mD' rT' b'. | |
| 360 | s = ((rev apTs) @ (X # ST), LT) \<and> length apTs = length fpTs \<and> is_class G C \<and> | |
| 361 | G \<turnstile> X \<preceq> Class C \<and> (\<forall>(aT,fT)\<in>set(zip apTs fpTs). G \<turnstile> aT \<preceq> fT) \<and> | |
| 362 | method (G,C) (mn,fpTs) = Some (mD', rT', b') \<and> | |
| 363 | (\<forall>C \<in> set (match_any G pc et). is_class G C))" (is "?app s = ?P s") | |
| 25362 | 364 | proof (cases s) | 
| 12772 | 365 | note list_all2_def [simp] | 
| 25362 | 366 | case (Pair a b) | 
| 13006 | 367 | have "?app (a,b) \<Longrightarrow> ?P (a,b)" | 
| 12516 | 368 | proof - | 
| 369 | assume app: "?app (a,b)" | |
| 370 | hence "a = (rev (rev (take (length fpTs) a))) @ (drop (length fpTs) a) \<and> | |
| 371 | length fpTs < length a" (is "?a \<and> ?l") | |
| 372 | by (auto simp add: app_def) | |
| 373 | hence "?a \<and> 0 < length (drop (length fpTs) a)" (is "?a \<and> ?l") | |
| 374 | by auto | |
| 375 | hence "?a \<and> ?l \<and> length (rev (take (length fpTs) a)) = length fpTs" | |
| 32443 | 376 | by (auto) | 
| 12516 | 377 | hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> 0 < length ST" | 
| 378 | by blast | |
| 379 | hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> ST \<noteq> []" | |
| 380 | by blast | |
| 381 | hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> | |
| 382 | (\<exists>X ST'. ST = X#ST')" | |
| 383 | by (simp add: neq_Nil_conv) | |
| 384 | hence "\<exists>apTs X ST. a = rev apTs @ X # ST \<and> length apTs = length fpTs" | |
| 385 | by blast | |
| 386 | with app | |
| 387 | show ?thesis by (unfold app_def, clarsimp) blast | |
| 388 | qed | |
| 389 | with Pair | |
| 12772 | 390 | have "?app s \<Longrightarrow> ?P s" by (simp only:) | 
| 12516 | 391 | moreover | 
| 32642 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 haftmann parents: 
32443diff
changeset | 392 | have "?P s \<Longrightarrow> ?app s" by (clarsimp simp add: min_max.inf_absorb2) | 
| 12516 | 393 | ultimately | 
| 12772 | 394 | show ?thesis by (rule iffI) | 
| 12516 | 395 | qed | 
| 396 | ||
| 397 | lemma effNone: | |
| 398 | "(pc', s') \<in> set (eff i G pc et None) \<Longrightarrow> s' = None" | |
| 399 | by (auto simp add: eff_def xcpt_eff_def norm_eff_def) | |
| 400 | ||
| 12772 | 401 | |
| 402 | lemma xcpt_app_lemma [code]: | |
| 403 | "xcpt_app i G pc et = list_all (is_class G) (xcpt_names (i, G, pc, et))" | |
| 17088 | 404 | by (simp add: list_all_iff) | 
| 12772 | 405 | |
| 12516 | 406 | lemmas [simp del] = app_def xcpt_app_def | 
| 407 | ||
| 408 | end |