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