| author | wenzelm | 
| Thu, 21 Aug 2008 19:19:31 +0200 | |
| changeset 27941 | b4656b671cce | 
| parent 24699 | c6674504103f | 
| child 28524 | 644b62cf678f | 
| permissions | -rw-r--r-- | 
| 13673 | 1 | (* Title: HOL/MicroJava/Comp/CorrComp.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Martin Strecker | |
| 4 | *) | |
| 5 | ||
| 6 | (* Compiler correctness statement and proof *) | |
| 7 | ||
| 15860 | 8 | theory CorrComp | 
| 15866 | 9 | imports "../J/JTypeSafe" "LemmasComp" | 
| 15860 | 10 | begin | 
| 13673 | 11 | |
| 14045 | 12 | declare wf_prog_ws_prog [simp add] | 
| 13673 | 13 | |
| 14 | (* If no exception is present after evaluation/execution, | |
| 15 | none can have been present before *) | |
| 16 | lemma eval_evals_exec_xcpt: | |
| 22271 | 17 | "(G \<turnstile> xs -ex\<succ>val-> xs' \<longrightarrow> gx xs' = None \<longrightarrow> gx xs = None) \<and> | 
| 18 | (G \<turnstile> xs -exs[\<succ>]vals-> xs' \<longrightarrow> gx xs' = None \<longrightarrow> gx xs = None) \<and> | |
| 19 | (G \<turnstile> xs -st-> xs' \<longrightarrow> gx xs' = None \<longrightarrow> gx xs = None)" | |
| 13673 | 20 | by (induct rule: eval_evals_exec.induct, auto) | 
| 21 | ||
| 22 | ||
| 23 | (* instance of eval_evals_exec_xcpt for eval *) | |
| 22271 | 24 | lemma eval_xcpt: "G \<turnstile> xs -ex\<succ>val-> xs' \<Longrightarrow> gx xs' = None \<Longrightarrow> gx xs = None" | 
| 13673 | 25 | (is "?H1 \<Longrightarrow> ?H2 \<Longrightarrow> ?T") | 
| 26 | proof- | |
| 27 | assume h1: ?H1 | |
| 28 | assume h2: ?H2 | |
| 29 | from h1 h2 eval_evals_exec_xcpt show "?T" by simp | |
| 30 | qed | |
| 31 | ||
| 32 | (* instance of eval_evals_exec_xcpt for evals *) | |
| 22271 | 33 | lemma evals_xcpt: "G \<turnstile> xs -exs[\<succ>]vals-> xs' \<Longrightarrow> gx xs' = None \<Longrightarrow> gx xs = None" | 
| 13673 | 34 | (is "?H1 \<Longrightarrow> ?H2 \<Longrightarrow> ?T") | 
| 35 | proof- | |
| 36 | assume h1: ?H1 | |
| 37 | assume h2: ?H2 | |
| 38 | from h1 h2 eval_evals_exec_xcpt show "?T" by simp | |
| 39 | qed | |
| 40 | ||
| 41 | (* instance of eval_evals_exec_xcpt for exec *) | |
| 22271 | 42 | lemma exec_xcpt: "G \<turnstile> xs -st-> xs' \<Longrightarrow> gx xs' = None \<Longrightarrow> gx xs = None" | 
| 13673 | 43 | (is "?H1 \<Longrightarrow> ?H2 \<Longrightarrow> ?T") | 
| 44 | proof- | |
| 45 | assume h1: ?H1 | |
| 46 | assume h2: ?H2 | |
| 47 | from h1 h2 eval_evals_exec_xcpt show "?T" by simp | |
| 48 | qed | |
| 49 | ||
| 50 | (**********************************************************************) | |
| 51 | ||
| 52 | theorem exec_all_trans: "\<lbrakk>(exec_all G s0 s1); (exec_all G s1 s2)\<rbrakk> \<Longrightarrow> (exec_all G s0 s2)" | |
| 53 | apply (auto simp: exec_all_def elim: Transitive_Closure.rtrancl_trans) | |
| 54 | done | |
| 55 | ||
| 56 | theorem exec_all_refl: "exec_all G s s" | |
| 15097 
b807858b97bd
Modifications for trancl_tac (new solver in simplifier).
 ballarin parents: 
15076diff
changeset | 57 | by (simp only: exec_all_def) | 
| 13673 | 58 | |
| 59 | ||
| 60 | theorem exec_instr_in_exec_all: | |
| 61 | "\<lbrakk> exec_instr i G hp stk lvars C S pc frs = (None, hp', frs'); | |
| 62 | gis (gmb G C S) ! pc = i\<rbrakk> \<Longrightarrow> | |
| 63 | G \<turnstile> (None, hp, (stk, lvars, C, S, pc) # frs) -jvm\<rightarrow> (None, hp', frs')" | |
| 64 | apply (simp only: exec_all_def) | |
| 65 | apply (rule rtrancl_refl [THEN rtrancl_into_rtrancl]) | |
| 66 | apply (simp add: gis_def gmb_def) | |
| 67 | apply (case_tac frs', simp+) | |
| 68 | done | |
| 69 | ||
| 70 | theorem exec_all_one_step: " | |
| 71 | \<lbrakk> gis (gmb G C S) = pre @ (i # post); pc0 = length pre; | |
| 72 | (exec_instr i G hp0 stk0 lvars0 C S pc0 frs) = | |
| 73 | (None, hp1, (stk1,lvars1,C,S, Suc pc0)#frs) \<rbrakk> | |
| 74 | \<Longrightarrow> | |
| 75 | G \<turnstile> (None, hp0, (stk0,lvars0,C,S, pc0)#frs) -jvm\<rightarrow> | |
| 76 | (None, hp1, (stk1,lvars1,C,S, Suc pc0)#frs)" | |
| 77 | apply (unfold exec_all_def) | |
| 78 | apply (rule r_into_rtrancl) | |
| 79 | apply (simp add: gis_def gmb_def split_beta) | |
| 80 | done | |
| 81 | ||
| 82 | ||
| 83 | (***********************************************************************) | |
| 84 | ||
| 85 | constdefs | |
| 86 | progression :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> | |
| 87 | aheap \<Rightarrow> opstack \<Rightarrow> locvars \<Rightarrow> | |
| 88 | bytecode \<Rightarrow> | |
| 89 | aheap \<Rightarrow> opstack \<Rightarrow> locvars \<Rightarrow> | |
| 90 | bool" | |
| 91 |   ("{_,_,_} \<turnstile> {_, _, _} >- _ \<rightarrow> {_, _, _}" [61,61,61,61,61,61,90,61,61,61]60)
 | |
| 92 |   "{G,C,S} \<turnstile> {hp0, os0, lvars0} >- instrs \<rightarrow> {hp1, os1, lvars1} ==
 | |
| 93 | \<forall> pre post frs. | |
| 94 | (gis (gmb G C S) = pre @ instrs @ post) \<longrightarrow> | |
| 95 | G \<turnstile> (None,hp0,(os0,lvars0,C,S,length pre)#frs) -jvm\<rightarrow> | |
| 96 | (None,hp1,(os1,lvars1,C,S,(length pre) + (length instrs))#frs)" | |
| 97 | ||
| 98 | ||
| 99 | ||
| 100 | lemma progression_call: | |
| 101 | "\<lbrakk> \<forall> pc frs. | |
| 102 | exec_instr instr G hp0 os0 lvars0 C S pc frs = | |
| 103 | (None, hp', (os', lvars', C', S', 0) # (fr pc) # frs) \<and> | |
| 104 | gis (gmb G C' S') = instrs' @ [Return] \<and> | |
| 105 |   {G, C', S'} \<turnstile> {hp', os', lvars'} >- instrs' \<rightarrow> {hp'', os'', lvars''}  \<and>
 | |
| 106 | exec_instr Return G hp'' os'' lvars'' C' S' (length instrs') | |
| 107 | ((fr pc) # frs) = | |
| 108 | (None, hp2, (os2, lvars2, C, S, Suc pc) # frs) \<rbrakk> \<Longrightarrow> | |
| 109 |   {G, C, S} \<turnstile> {hp0, os0, lvars0} >-[instr]\<rightarrow> {hp2,os2,lvars2}"
 | |
| 110 | apply (simp only: progression_def) | |
| 111 | apply (intro strip) | |
| 112 | apply (drule_tac x="length pre" in spec) | |
| 113 | apply (drule_tac x="frs" in spec) | |
| 114 | apply clarify | |
| 115 | apply (rule exec_all_trans) | |
| 116 | apply (rule exec_instr_in_exec_all) apply simp | |
| 117 | apply simp | |
| 118 | apply (rule exec_all_trans) | |
| 119 | apply (simp only: append_Nil) | |
| 120 | apply (drule_tac x="[]" in spec) | |
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24074diff
changeset | 121 | apply (simp only: list.simps list.size) | 
| 13673 | 122 | apply blast | 
| 123 | apply (rule exec_instr_in_exec_all) | |
| 124 | apply auto | |
| 125 | done | |
| 126 | ||
| 127 | ||
| 128 | lemma progression_transitive: | |
| 129 | "\<lbrakk> instrs_comb = instrs0 @ instrs1; | |
| 130 |   {G, C, S} \<turnstile> {hp0, os0, lvars0} >- instrs0 \<rightarrow> {hp1, os1, lvars1};
 | |
| 131 |   {G, C, S} \<turnstile> {hp1, os1, lvars1} >- instrs1 \<rightarrow> {hp2, os2, lvars2} \<rbrakk>
 | |
| 132 | \<Longrightarrow> | |
| 133 |   {G, C, S} \<turnstile> {hp0, os0, lvars0} >- instrs_comb \<rightarrow> {hp2, os2, lvars2}"
 | |
| 134 | apply (simp only: progression_def) | |
| 135 | apply (intro strip) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 136 | apply (rule_tac ?s1.0 = "Norm (hp1, (os1, lvars1, C, S, | 
| 13673 | 137 | length pre + length instrs0) # frs)" | 
| 138 | in exec_all_trans) | |
| 139 | apply (simp only: append_assoc) | |
| 140 | apply (erule thin_rl, erule thin_rl) | |
| 141 | apply (drule_tac x="pre @ instrs0" in spec) | |
| 142 | apply (simp add: add_assoc) | |
| 143 | done | |
| 144 | ||
| 145 | lemma progression_refl: | |
| 146 |   "{G, C, S} \<turnstile> {hp0, os0, lvars0} >- [] \<rightarrow> {hp0, os0, lvars0}"
 | |
| 147 | apply (simp add: progression_def) | |
| 148 | apply (intro strip) | |
| 149 | apply (rule exec_all_refl) | |
| 150 | done | |
| 151 | ||
| 152 | lemma progression_one_step: " | |
| 153 | \<forall> pc frs. | |
| 154 | (exec_instr i G hp0 os0 lvars0 C S pc frs) = | |
| 155 | (None, hp1, (os1,lvars1,C,S, Suc pc)#frs) | |
| 156 |   \<Longrightarrow> {G, C, S} \<turnstile> {hp0, os0, lvars0} >- [i] \<rightarrow> {hp1, os1, lvars1}"
 | |
| 157 | apply (unfold progression_def) | |
| 158 | apply (intro strip) | |
| 159 | apply simp | |
| 160 | apply (rule exec_all_one_step) | |
| 161 | apply auto | |
| 162 | done | |
| 163 | ||
| 164 | (*****) | |
| 165 | constdefs | |
| 166 | jump_fwd :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> | |
| 167 | aheap \<Rightarrow> locvars \<Rightarrow> opstack \<Rightarrow> opstack \<Rightarrow> | |
| 168 | instr \<Rightarrow> bytecode \<Rightarrow> bool" | |
| 169 | "jump_fwd G C S hp lvars os0 os1 instr instrs == | |
| 170 | \<forall> pre post frs. | |
| 171 | (gis (gmb G C S) = pre @ instr # instrs @ post) \<longrightarrow> | |
| 172 | exec_all G (None,hp,(os0,lvars,C,S, length pre)#frs) | |
| 173 | (None,hp,(os1,lvars,C,S,(length pre) + (length instrs) + 1)#frs)" | |
| 174 | ||
| 175 | ||
| 176 | lemma jump_fwd_one_step: | |
| 177 | "\<forall> pc frs. | |
| 178 | exec_instr instr G hp os0 lvars C S pc frs = | |
| 179 | (None, hp, (os1, lvars, C, S, pc + (length instrs) + 1)#frs) | |
| 180 | \<Longrightarrow> jump_fwd G C S hp lvars os0 os1 instr instrs" | |
| 181 | apply (unfold jump_fwd_def) | |
| 182 | apply (intro strip) | |
| 183 | apply (rule exec_instr_in_exec_all) | |
| 184 | apply auto | |
| 185 | done | |
| 186 | ||
| 187 | ||
| 188 | lemma jump_fwd_progression_aux: | |
| 189 | "\<lbrakk> instrs_comb = instr # instrs0 @ instrs1; | |
| 190 | jump_fwd G C S hp lvars os0 os1 instr instrs0; | |
| 191 |      {G, C, S} \<turnstile> {hp, os1, lvars} >- instrs1 \<rightarrow> {hp2, os2, lvars2} \<rbrakk> 
 | |
| 192 |   \<Longrightarrow> {G, C, S} \<turnstile> {hp, os0, lvars} >- instrs_comb \<rightarrow> {hp2, os2, lvars2}"
 | |
| 193 | apply (simp only: progression_def jump_fwd_def) | |
| 194 | apply (intro strip) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 195 | apply (rule_tac ?s1.0 = "Norm(hp, (os1, lvars, C, S, length pre + length instrs0 + 1) # frs)" in exec_all_trans) | 
| 13673 | 196 | apply (simp only: append_assoc) | 
| 197 | apply (subgoal_tac "pre @ (instr # instrs0 @ instrs1) @ post = pre @ instr # instrs0 @ (instrs1 @ post)") | |
| 198 | apply blast | |
| 199 | apply simp | |
| 200 | apply (erule thin_rl, erule thin_rl) | |
| 201 | apply (drule_tac x="pre @ instr # instrs0" in spec) | |
| 202 | apply (simp add: add_assoc) | |
| 203 | done | |
| 204 | ||
| 205 | ||
| 206 | lemma jump_fwd_progression: | |
| 207 | "\<lbrakk> instrs_comb = instr # instrs0 @ instrs1; | |
| 208 | \<forall> pc frs. | |
| 209 | exec_instr instr G hp os0 lvars C S pc frs = | |
| 210 | (None, hp, (os1, lvars, C, S, pc + (length instrs0) + 1)#frs); | |
| 211 |   {G, C, S} \<turnstile> {hp, os1, lvars} >- instrs1 \<rightarrow> {hp2, os2, lvars2} \<rbrakk> 
 | |
| 212 |   \<Longrightarrow> {G, C, S}  \<turnstile> {hp, os0, lvars} >- instrs_comb \<rightarrow> {hp2, os2, lvars2}"
 | |
| 213 | apply (rule jump_fwd_progression_aux) | |
| 214 | apply assumption | |
| 215 | apply (rule jump_fwd_one_step) apply assumption+ | |
| 216 | done | |
| 217 | ||
| 218 | ||
| 219 | (* note: instrs and instr reversed wrt. jump_fwd *) | |
| 220 | constdefs | |
| 221 | jump_bwd :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> | |
| 222 | aheap \<Rightarrow> locvars \<Rightarrow> opstack \<Rightarrow> opstack \<Rightarrow> | |
| 223 | bytecode \<Rightarrow> instr \<Rightarrow> bool" | |
| 224 | "jump_bwd G C S hp lvars os0 os1 instrs instr == | |
| 225 | \<forall> pre post frs. | |
| 226 | (gis (gmb G C S) = pre @ instrs @ instr # post) \<longrightarrow> | |
| 227 | exec_all G (None,hp,(os0,lvars,C,S, (length pre) + (length instrs))#frs) | |
| 228 | (None,hp,(os1,lvars,C,S, (length pre))#frs)" | |
| 229 | ||
| 230 | ||
| 231 | lemma jump_bwd_one_step: | |
| 232 | "\<forall> pc frs. | |
| 233 | exec_instr instr G hp os0 lvars C S (pc + (length instrs)) frs = | |
| 234 | (None, hp, (os1, lvars, C, S, pc)#frs) | |
| 235 | \<Longrightarrow> | |
| 236 | jump_bwd G C S hp lvars os0 os1 instrs instr" | |
| 237 | apply (unfold jump_bwd_def) | |
| 238 | apply (intro strip) | |
| 239 | apply (rule exec_instr_in_exec_all) | |
| 240 | apply auto | |
| 241 | done | |
| 242 | ||
| 243 | lemma jump_bwd_progression: | |
| 244 | "\<lbrakk> instrs_comb = instrs @ [instr]; | |
| 245 |   {G, C, S} \<turnstile> {hp0, os0, lvars0} >- instrs \<rightarrow> {hp1, os1, lvars1};
 | |
| 246 | jump_bwd G C S hp1 lvars1 os1 os2 instrs instr; | |
| 247 |   {G, C, S} \<turnstile> {hp1, os2, lvars1} >- instrs_comb \<rightarrow> {hp3, os3, lvars3} \<rbrakk> 
 | |
| 248 |   \<Longrightarrow> {G, C, S}  \<turnstile> {hp0, os0, lvars0} >- instrs_comb \<rightarrow> {hp3, os3, lvars3}"
 | |
| 249 | apply (simp only: progression_def jump_bwd_def) | |
| 250 | apply (intro strip) | |
| 251 | apply (rule exec_all_trans, force) | |
| 252 | apply (rule exec_all_trans, force) | |
| 253 | apply (rule exec_all_trans, force) | |
| 254 | apply simp | |
| 255 | apply (rule exec_all_refl) | |
| 256 | done | |
| 257 | ||
| 258 | ||
| 259 | (**********************************************************************) | |
| 260 | ||
| 261 | (* class C with signature S is defined in program G *) | |
| 262 | constdefs class_sig_defined :: "'c prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> bool" | |
| 263 | "class_sig_defined G C S == | |
| 264 | is_class G C \<and> (\<exists> D rT mb. (method (G, C) S = Some (D, rT, mb)))" | |
| 265 | ||
| 266 | ||
| 267 | (* The environment of a java method body | |
| 268 | (characterized by class and signature) *) | |
| 269 | constdefs env_of_jmb :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> java_mb env" | |
| 270 | "env_of_jmb G C S == | |
| 271 | (let (mn,pTs) = S; | |
| 272 | (D,rT,(pns,lvars,blk,res)) = the(method (G, C) S) in | |
| 273 | (G,map_of lvars(pns[\<mapsto>]pTs)(This\<mapsto>Class C)))" | |
| 274 | ||
| 275 | lemma env_of_jmb_fst [simp]: "fst (env_of_jmb G C S) = G" | |
| 276 | by (simp add: env_of_jmb_def split_beta) | |
| 277 | ||
| 278 | ||
| 279 | (**********************************************************************) | |
| 280 | ||
| 281 | ||
| 282 | lemma method_preserves [rule_format (no_asm)]: | |
| 283 | "\<lbrakk> wf_prog wf_mb G; is_class G C; | |
| 284 | \<forall> S rT mb. \<forall> cn \<in> fst ` set G. wf_mdecl wf_mb G cn (S,rT,mb) \<longrightarrow> (P cn S (rT,mb))\<rbrakk> | |
| 285 | \<Longrightarrow> \<forall> D. | |
| 286 | method (G, C) S = Some (D, rT, mb) \<longrightarrow> (P D S (rT,mb))" | |
| 287 | ||
| 14045 | 288 | apply (frule wf_prog_ws_prog [THEN wf_subcls1]) | 
| 13673 | 289 | apply (rule subcls1_induct, assumption, assumption) | 
| 290 | ||
| 291 | apply (intro strip) | |
| 292 | apply ((drule spec)+, drule_tac x="Object" in bspec) | |
| 14045 | 293 | apply (simp add: wf_prog_def ws_prog_def wf_syscls_def) | 
| 13673 | 294 | apply (subgoal_tac "D=Object") apply simp | 
| 295 | apply (drule mp) | |
| 296 | apply (frule_tac C=Object in method_wf_mdecl) | |
| 14045 | 297 | apply simp | 
| 298 | apply assumption apply simp apply assumption apply simp | |
| 13673 | 299 | |
| 15481 | 300 | apply (simplesubst method_rec) apply simp | 
| 13673 | 301 | apply force | 
| 14045 | 302 | apply simp | 
| 14025 | 303 | apply (simp only: map_add_def) | 
| 13673 | 304 | apply (split option.split) | 
| 305 | apply (rule conjI) | |
| 306 | apply force | |
| 307 | apply (intro strip) | |
| 308 | apply (frule_tac | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 309 | ?P1.0 = "wf_mdecl wf_mb G Ca" and | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 310 | ?P2.0 = "%(S, (Da, rT, mb)). P Da S (rT, mb)" in map_of_map_prop) | 
| 13673 | 311 | apply (force simp: wf_cdecl_def) | 
| 312 | ||
| 313 | apply clarify | |
| 314 | ||
| 315 | apply (simp only: class_def) | |
| 316 | apply (drule map_of_SomeD)+ | |
| 317 | apply (frule_tac A="set G" and f=fst in imageI, simp) | |
| 318 | apply blast | |
| 319 | apply simp | |
| 320 | done | |
| 321 | ||
| 322 | ||
| 323 | lemma method_preserves_length: | |
| 324 | "\<lbrakk> wf_java_prog G; is_class G C; | |
| 325 | method (G, C) (mn,pTs) = Some (D, rT, pns, lvars, blk, res)\<rbrakk> | |
| 326 | \<Longrightarrow> length pns = length pTs" | |
| 327 | apply (frule_tac | |
| 328 | P="%D (mn,pTs) (rT, pns, lvars, blk, res). length pns = length pTs" | |
| 329 | in method_preserves) | |
| 330 | apply (auto simp: wf_mdecl_def wf_java_mdecl_def) | |
| 331 | done | |
| 332 | ||
| 333 | (**********************************************************************) | |
| 334 | ||
| 335 | constdefs wtpd_expr :: "java_mb env \<Rightarrow> expr \<Rightarrow> bool" | |
| 336 | "wtpd_expr E e == (\<exists> T. E\<turnstile>e :: T)" | |
| 337 | wtpd_exprs :: "java_mb env \<Rightarrow> (expr list) \<Rightarrow> bool" | |
| 338 | "wtpd_exprs E e == (\<exists> T. E\<turnstile>e [::] T)" | |
| 339 | wtpd_stmt :: "java_mb env \<Rightarrow> stmt \<Rightarrow> bool" | |
| 340 | "wtpd_stmt E c == (E\<turnstile>c \<surd>)" | |
| 341 | ||
| 342 | lemma wtpd_expr_newc: "wtpd_expr E (NewC C) \<Longrightarrow> is_class (prg E) C" | |
| 343 | by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) | |
| 344 | ||
| 345 | lemma wtpd_expr_cast: "wtpd_expr E (Cast cn e) \<Longrightarrow> (wtpd_expr E e)" | |
| 346 | by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) | |
| 347 | ||
| 348 | lemma wtpd_expr_lacc: "\<lbrakk> wtpd_expr (env_of_jmb G C S) (LAcc vn); | |
| 349 | class_sig_defined G C S \<rbrakk> | |
| 350 | \<Longrightarrow> vn \<in> set (gjmb_plns (gmb G C S)) \<or> vn = This" | |
| 351 | apply (simp only: wtpd_expr_def env_of_jmb_def class_sig_defined_def galldefs) | |
| 352 | apply clarify | |
| 353 | apply (case_tac S) | |
| 354 | apply simp | |
| 355 | apply (erule ty_expr.cases) | |
| 356 | apply (auto dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma) | |
| 357 | apply (drule map_upds_SomeD) | |
| 358 | apply (erule disjE) | |
| 359 | apply assumption | |
| 360 | apply (drule map_of_SomeD) apply (auto dest: fst_in_set_lemma) | |
| 361 | done | |
| 362 | ||
| 363 | lemma wtpd_expr_lass: "wtpd_expr E (vn::=e) | |
| 364 | \<Longrightarrow> (vn \<noteq> This) & (wtpd_expr E (LAcc vn)) & (wtpd_expr E e)" | |
| 365 | by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) | |
| 366 | ||
| 367 | lemma wtpd_expr_facc: "wtpd_expr E ({fd}a..fn) 
 | |
| 368 | \<Longrightarrow> (wtpd_expr E a)" | |
| 369 | by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) | |
| 370 | ||
| 371 | lemma wtpd_expr_fass: "wtpd_expr E ({fd}a..fn:=v) 
 | |
| 372 |   \<Longrightarrow> (wtpd_expr E ({fd}a..fn)) & (wtpd_expr E v)"
 | |
| 373 | by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) | |
| 374 | ||
| 375 | ||
| 376 | lemma wtpd_expr_binop: "wtpd_expr E (BinOp bop e1 e2) | |
| 377 | \<Longrightarrow> (wtpd_expr E e1) & (wtpd_expr E e2)" | |
| 378 | by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) | |
| 379 | ||
| 380 | lemma wtpd_exprs_cons: "wtpd_exprs E (e # es) | |
| 381 | \<Longrightarrow> (wtpd_expr E e) & (wtpd_exprs E es)" | |
| 382 | by (simp only: wtpd_exprs_def wtpd_expr_def, erule exE, erule ty_exprs.cases, auto) | |
| 383 | ||
| 384 | lemma wtpd_stmt_expr: "wtpd_stmt E (Expr e) \<Longrightarrow> (wtpd_expr E e)" | |
| 385 | by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) | |
| 386 | ||
| 387 | lemma wtpd_stmt_comp: "wtpd_stmt E (s1;; s2) \<Longrightarrow> | |
| 388 | (wtpd_stmt E s1) & (wtpd_stmt E s2)" | |
| 389 | by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) | |
| 390 | ||
| 391 | lemma wtpd_stmt_cond: "wtpd_stmt E (If(e) s1 Else s2) \<Longrightarrow> | |
| 392 | (wtpd_expr E e) & (wtpd_stmt E s1) & (wtpd_stmt E s2) | |
| 393 | & (E\<turnstile>e::PrimT Boolean)" | |
| 394 | by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) | |
| 395 | ||
| 396 | lemma wtpd_stmt_loop: "wtpd_stmt E (While(e) s) \<Longrightarrow> | |
| 397 | (wtpd_expr E e) & (wtpd_stmt E s) & (E\<turnstile>e::PrimT Boolean)" | |
| 398 | by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) | |
| 399 | ||
| 400 | lemma wtpd_expr_call: "wtpd_expr E ({C}a..mn({pTs'}ps))
 | |
| 401 | \<Longrightarrow> (wtpd_expr E a) & (wtpd_exprs E ps) | |
| 402 | & (length ps = length pTs') & (E\<turnstile>a::Class C) | |
| 403 | & (\<exists> pTs md rT. | |
| 404 |        E\<turnstile>ps[::]pTs & max_spec (prg E) C (mn, pTs) = {((md,rT),pTs')})"
 | |
| 405 | apply (simp only: wtpd_expr_def wtpd_exprs_def) | |
| 406 | apply (erule exE) | |
| 23757 | 407 | apply (ind_cases "E \<turnstile> {C}a..mn( {pTs'}ps) :: T" for T)
 | 
| 13673 | 408 | apply (auto simp: max_spec_preserves_length) | 
| 409 | done | |
| 410 | ||
| 411 | lemma wtpd_blk: | |
| 412 | "\<lbrakk> method (G, D) (md, pTs) = Some (D, rT, (pns, lvars, blk, res)); | |
| 413 | wf_prog wf_java_mdecl G; is_class G D \<rbrakk> | |
| 414 | \<Longrightarrow> wtpd_stmt (env_of_jmb G D (md, pTs)) blk" | |
| 415 | apply (simp add: wtpd_stmt_def env_of_jmb_def) | |
| 416 | apply (frule_tac P="%D (md, pTs) (rT, (pns, lvars, blk, res)). (G, map_of lvars(pns[\<mapsto>]pTs)(This\<mapsto>Class D)) \<turnstile> blk \<surd> " in method_preserves) | |
| 417 | apply (auto simp: wf_mdecl_def wf_java_mdecl_def) | |
| 418 | done | |
| 419 | ||
| 420 | lemma wtpd_res: | |
| 421 | "\<lbrakk> method (G, D) (md, pTs) = Some (D, rT, (pns, lvars, blk, res)); | |
| 422 | wf_prog wf_java_mdecl G; is_class G D \<rbrakk> | |
| 423 | \<Longrightarrow> wtpd_expr (env_of_jmb G D (md, pTs)) res" | |
| 424 | apply (simp add: wtpd_expr_def env_of_jmb_def) | |
| 425 | apply (frule_tac P="%D (md, pTs) (rT, (pns, lvars, blk, res)). \<exists>T. (G, map_of lvars(pns[\<mapsto>]pTs)(This\<mapsto>Class D)) \<turnstile> res :: T " in method_preserves) | |
| 426 | apply (auto simp: wf_mdecl_def wf_java_mdecl_def) | |
| 427 | done | |
| 428 | ||
| 429 | ||
| 430 | (**********************************************************************) | |
| 431 | ||
| 432 | ||
| 433 | (* Is there a more elegant proof? *) | |
| 434 | lemma evals_preserves_length: | |
| 435 | "G\<turnstile> xs -es[\<succ>]vs-> (None, s) \<Longrightarrow> length es = length vs" | |
| 436 | apply (subgoal_tac | |
| 437 | "\<forall> xs'. (G \<turnstile> xk -xj\<succ>xi-> xh \<longrightarrow> True) & | |
| 438 | (G\<turnstile> xs -es[\<succ>]vs-> xs' \<longrightarrow> (\<exists> s. (xs' = (None, s))) \<longrightarrow> | |
| 439 | length es = length vs) & | |
| 22271 | 440 | (G \<turnstile> xc -xb-> xa \<longrightarrow> True)") | 
| 13673 | 441 | apply blast | 
| 442 | apply (rule allI) | |
| 443 | apply (rule Eval.eval_evals_exec.induct) | |
| 444 | apply auto | |
| 445 | done | |
| 446 | ||
| 447 | (***********************************************************************) | |
| 448 | ||
| 449 | (* required for translation of BinOp *) | |
| 450 | ||
| 451 | ||
| 452 | lemma progression_Eq : "{G, C, S} \<turnstile>
 | |
| 453 |   {hp, (v2 # v1 # os), lvars} 
 | |
| 454 | >- [Ifcmpeq 3, LitPush (Bool False), Goto 2, LitPush (Bool True)] \<rightarrow> | |
| 455 |   {hp, (Bool (v1 = v2) # os), lvars}"
 | |
| 456 | apply (case_tac "v1 = v2") | |
| 457 | ||
| 458 | (* case v1 = v2 *) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 459 | apply (rule_tac ?instrs1.0 = "[LitPush (Bool True)]" in jump_fwd_progression) | 
| 13837 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13673diff
changeset | 460 | apply (auto simp: nat_add_distrib) | 
| 13673 | 461 | apply (rule progression_one_step) apply simp | 
| 462 | ||
| 463 | (* case v1 \<noteq> v2 *) | |
| 464 | apply (rule progression_one_step [THEN HOL.refl [THEN progression_transitive], simplified]) | |
| 465 | apply auto | |
| 466 | apply (rule progression_one_step [THEN HOL.refl [THEN progression_transitive], simplified]) | |
| 467 | apply auto | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 468 | apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression) | 
| 13837 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13673diff
changeset | 469 | apply (auto simp: nat_add_distrib intro: progression_refl) | 
| 13673 | 470 | done | 
| 471 | ||
| 472 | ||
| 473 | (**********************************************************************) | |
| 474 | ||
| 475 | ||
| 476 | (* to avoid automatic pair splits *) | |
| 477 | ||
| 478 | declare split_paired_All [simp del] split_paired_Ex [simp del] | |
| 24074 | 479 | declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *}
 | 
| 13673 | 480 | |
| 481 | lemma distinct_method: "\<lbrakk> wf_java_prog G; is_class G C; | |
| 482 | method (G, C) S = Some (D, rT, pns, lvars, blk, res) \<rbrakk> \<Longrightarrow> | |
| 483 | distinct (gjmb_plns (gmb G C S))" | |
| 484 | apply (frule method_wf_mdecl [THEN conjunct2], assumption, assumption) | |
| 485 | apply (case_tac S) | |
| 486 | apply (simp add: wf_mdecl_def wf_java_mdecl_def galldefs distinct_append) | |
| 487 | apply (simp add: unique_def map_of_in_set) | |
| 488 | apply blast | |
| 489 | done | |
| 490 | ||
| 491 | lemma distinct_method_if_class_sig_defined : | |
| 492 | "\<lbrakk> wf_java_prog G; class_sig_defined G C S \<rbrakk> \<Longrightarrow> | |
| 493 | distinct (gjmb_plns (gmb G C S))" | |
| 494 | by (auto intro: distinct_method simp: class_sig_defined_def) | |
| 495 | ||
| 496 | ||
| 497 | lemma method_yields_wf_java_mdecl: "\<lbrakk> wf_java_prog G; is_class G C; | |
| 498 | method (G, C) S = Some (D, rT, pns, lvars, blk, res) \<rbrakk> \<Longrightarrow> | |
| 499 | wf_java_mdecl G D (S,rT,(pns,lvars,blk,res))" | |
| 500 | apply (frule method_wf_mdecl) | |
| 501 | apply (auto simp: wf_mdecl_def) | |
| 502 | done | |
| 503 | ||
| 504 | (**********************************************************************) | |
| 505 | ||
| 506 | ||
| 507 | lemma progression_lvar_init_aux [rule_format (no_asm)]: " | |
| 508 | \<forall> zs prfx lvals lvars0. | |
| 509 | lvars0 = (zs @ lvars) \<longrightarrow> | |
| 510 | (disjoint_varnames pns lvars0 \<longrightarrow> | |
| 511 | (length lvars = length lvals) \<longrightarrow> | |
| 512 | (Suc(length pns + length zs) = length prfx) \<longrightarrow> | |
| 513 |    ({cG, D, S} \<turnstile> 
 | |
| 514 |     {h, os, (prfx @ lvals)}
 | |
| 515 | >- (concat (map (compInit (pns, lvars0, blk, res)) lvars)) \<rightarrow> | |
| 516 |     {h, os, (prfx @ (map (\<lambda>p. (default_val (snd p))) lvars))}))"
 | |
| 517 | apply simp | |
| 518 | apply (induct lvars) | |
| 519 | apply (clarsimp, rule progression_refl) | |
| 520 | apply (intro strip) | |
| 521 | apply (case_tac lvals) apply simp | |
| 522 | apply (simp (no_asm_simp) ) | |
| 523 | ||
| 15236 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
15097diff
changeset | 524 | apply (rule_tac ?lvars1.0 = "(prfx @ [default_val (snd a)]) @ list" in progression_transitive, rule HOL.refl) | 
| 13673 | 525 | apply (case_tac a) apply (simp (no_asm_simp) add: compInit_def) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 526 | apply (rule_tac ?instrs0.0 = "[load_default_val b]" in progression_transitive, simp) | 
| 13673 | 527 | apply (rule progression_one_step) | 
| 528 | apply (simp (no_asm_simp) add: load_default_val_def) | |
| 529 | apply (rule conjI, simp)+ apply (rule HOL.refl) | |
| 530 | ||
| 531 | apply (rule progression_one_step) | |
| 532 | apply (simp (no_asm_simp)) | |
| 533 | apply (rule conjI, simp)+ | |
| 534 | apply (simp add: index_of_var2) | |
| 535 | apply (drule_tac x="zs @ [a]" in spec) (* instantiate zs *) | |
| 536 | apply (drule mp, simp) | |
| 537 | apply (drule_tac x="(prfx @ [default_val (snd a)])" in spec) (* instantiate prfx *) | |
| 538 | apply auto | |
| 539 | done | |
| 540 | ||
| 541 | lemma progression_lvar_init [rule_format (no_asm)]: | |
| 542 | "\<lbrakk> wf_java_prog G; is_class G C; | |
| 543 | method (G, C) S = Some (D, rT, pns, lvars, blk, res) \<rbrakk> \<Longrightarrow> | |
| 544 | length pns = length pvs \<longrightarrow> | |
| 545 | (\<forall> lvals. | |
| 546 | length lvars = length lvals \<longrightarrow> | |
| 547 |    {cG, D, S} \<turnstile>
 | |
| 548 |    {h, os, (a' # pvs @ lvals)}
 | |
| 549 | >- (compInitLvars (pns, lvars, blk, res) lvars) \<rightarrow> | |
| 550 |    {h, os, (locvars_xstate G C S (Norm (h, init_vars lvars(pns[\<mapsto>]pvs)(This\<mapsto>a'))))})"
 | |
| 551 | apply (simp only: compInitLvars_def) | |
| 552 | apply (frule method_yields_wf_java_mdecl, assumption, assumption) | |
| 553 | ||
| 554 | apply (simp only: wf_java_mdecl_def) | |
| 555 | apply (subgoal_tac "(\<forall>y\<in>set pns. y \<notin> set (map fst lvars))") | |
| 556 | apply (simp add: init_vars_def locvars_xstate_def locvars_locals_def galldefs unique_def split_def map_of_map_as_map_upd) | |
| 557 | apply (intro strip) | |
| 558 | apply (simp (no_asm_simp) only: append_Cons [THEN sym]) | |
| 559 | apply (rule progression_lvar_init_aux) | |
| 560 | apply (auto simp: unique_def map_of_in_set disjoint_varnames_def) | |
| 561 | done | |
| 562 | ||
| 563 | ||
| 564 | ||
| 565 | ||
| 566 | (**********************************************************************) | |
| 567 | ||
| 14045 | 568 | lemma state_ok_eval: "\<lbrakk>xs::\<preceq>E; wf_java_prog (prg E); wtpd_expr E e; | 
| 569 | (prg E)\<turnstile>xs -e\<succ>v -> xs'\<rbrakk> \<Longrightarrow> xs'::\<preceq>E" | |
| 570 | apply (simp only: wtpd_expr_def) | |
| 13673 | 571 | apply (erule exE) | 
| 14045 | 572 | apply (case_tac xs', case_tac xs) | 
| 573 | apply (auto intro: eval_type_sound [THEN conjunct1]) | |
| 13673 | 574 | done | 
| 575 | ||
| 14045 | 576 | lemma state_ok_evals: "\<lbrakk>xs::\<preceq>E; wf_java_prog (prg E); wtpd_exprs E es; | 
| 22271 | 577 | prg E \<turnstile> xs -es[\<succ>]vs-> xs'\<rbrakk> \<Longrightarrow> xs'::\<preceq>E" | 
| 14045 | 578 | apply (simp only: wtpd_exprs_def) | 
| 579 | apply (erule exE) | |
| 580 | apply (case_tac xs) apply (case_tac xs') | |
| 581 | apply (auto intro: evals_type_sound [THEN conjunct1]) | |
| 13673 | 582 | done | 
| 583 | ||
| 14045 | 584 | lemma state_ok_exec: "\<lbrakk>xs::\<preceq>E; wf_java_prog (prg E); wtpd_stmt E st; | 
| 22271 | 585 | prg E \<turnstile> xs -st-> xs'\<rbrakk> \<Longrightarrow> xs'::\<preceq>E" | 
| 14045 | 586 | apply (simp only: wtpd_stmt_def) | 
| 587 | apply (case_tac xs', case_tac xs) | |
| 14143 | 588 | apply (auto dest: exec_type_sound) | 
| 13673 | 589 | done | 
| 590 | ||
| 591 | ||
| 592 | lemma state_ok_init: | |
| 14045 | 593 | "\<lbrakk> wf_java_prog G; (x, h, l)::\<preceq>(env_of_jmb G C S); | 
| 13673 | 594 | is_class G dynT; | 
| 595 | method (G, dynT) (mn, pTs) = Some (md, rT, pns, lvars, blk, res); | |
| 596 | list_all2 (conf G h) pvs pTs; G,h \<turnstile> a' ::\<preceq> Class md\<rbrakk> | |
| 597 | \<Longrightarrow> | |
| 14045 | 598 | (np a' x, h, init_vars lvars(pns[\<mapsto>]pvs)(This\<mapsto>a'))::\<preceq>(env_of_jmb G md (mn, pTs))" | 
| 599 | apply (frule wf_prog_ws_prog) | |
| 13673 | 600 | apply (frule method_in_md [THEN conjunct2], assumption+) | 
| 601 | apply (frule method_yields_wf_java_mdecl, assumption, assumption) | |
| 14045 | 602 | apply (simp add: env_of_jmb_def gs_def conforms_def split_beta) | 
| 13673 | 603 | apply (simp add: wf_java_mdecl_def) | 
| 604 | apply (rule conjI) | |
| 605 | apply (rule lconf_ext) | |
| 606 | apply (rule lconf_ext_list) | |
| 607 | apply (rule lconf_init_vars) | |
| 608 | apply (auto dest: Ball_set_table) | |
| 609 | apply (simp add: np_def xconf_raise_if) | |
| 610 | done | |
| 611 | ||
| 612 | ||
| 613 | lemma ty_exprs_list_all2 [rule_format (no_asm)]: | |
| 614 | "(\<forall> Ts. (E \<turnstile> es [::] Ts) = list_all2 (\<lambda>e T. E \<turnstile> e :: T) es Ts)" | |
| 615 | apply (rule list.induct) | |
| 616 | apply simp | |
| 617 | apply (rule allI) | |
| 618 | apply (rule iffI) | |
| 23757 | 619 | apply (ind_cases "E \<turnstile> [] [::] Ts" for Ts, assumption) | 
| 13673 | 620 | apply simp apply (rule WellType.Nil) | 
| 621 | apply (simp add: list_all2_Cons1) | |
| 622 | apply (rule allI) | |
| 623 | apply (rule iffI) | |
| 624 | apply (rename_tac a exs Ts) | |
| 23757 | 625 | apply (ind_cases "E \<turnstile> a # exs [::] Ts" for a exs Ts) apply blast | 
| 13673 | 626 | apply (auto intro: WellType.Cons) | 
| 627 | done | |
| 628 | ||
| 629 | ||
| 630 | lemma conf_bool: "G,h \<turnstile> v::\<preceq>PrimT Boolean \<Longrightarrow> \<exists> b. v = Bool b" | |
| 631 | apply (simp add: conf_def) | |
| 632 | apply (erule exE) | |
| 633 | apply (case_tac v) | |
| 634 | apply (auto elim: widen.cases) | |
| 635 | done | |
| 636 | ||
| 637 | ||
| 14045 | 638 | lemma class_expr_is_class: "\<lbrakk>E \<turnstile> e :: Class C; ws_prog (prg E)\<rbrakk> | 
| 13673 | 639 | \<Longrightarrow> is_class (prg E) C" | 
| 640 | by (case_tac E, auto dest: ty_expr_is_type) | |
| 641 | ||
| 642 | ||
| 643 | lemma max_spec_widen: "max_spec G C (mn, pTs) = {((md,rT),pTs')} \<Longrightarrow> 
 | |
| 644 | list_all2 (\<lambda> T T'. G \<turnstile> T \<preceq> T') pTs pTs'" | |
| 645 | by (blast dest: singleton_in_set max_spec2appl_meths appl_methsD) | |
| 646 | ||
| 647 | ||
| 14045 | 648 | lemma eval_conf: "\<lbrakk>G \<turnstile> s -e\<succ>v-> s'; wf_java_prog G; s::\<preceq>E; | 
| 13673 | 649 | E\<turnstile>e::T; gx s' = None; prg E = G \<rbrakk> | 
| 650 | \<Longrightarrow> G,gh s'\<turnstile>v::\<preceq>T" | |
| 651 | apply (simp add: gh_def) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 652 | apply (rule_tac x3="fst s" and s3="snd s"and x'3="fst s'" | 
| 14143 | 653 | in eval_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]], simplified]) | 
| 654 | apply assumption+ | |
| 13673 | 655 | apply (simp (no_asm_use) add: surjective_pairing [THEN sym]) | 
| 14045 | 656 | apply (simp only: surjective_pairing [THEN sym]) | 
| 657 | apply (auto simp add: gs_def gx_def) | |
| 13673 | 658 | done | 
| 659 | ||
| 660 | lemma evals_preserves_conf: | |
| 661 | "\<lbrakk> G\<turnstile> s -es[\<succ>]vs-> s'; G,gh s \<turnstile> t ::\<preceq> T; E \<turnstile>es[::]Ts; | |
| 14045 | 662 | wf_java_prog G; s::\<preceq>E; | 
| 13673 | 663 | prg E = G \<rbrakk> \<Longrightarrow> G,gh s' \<turnstile> t ::\<preceq> T" | 
| 664 | apply (subgoal_tac "gh s\<le>| gh s'") | |
| 665 | apply (frule conf_hext, assumption, assumption) | |
| 666 | apply (frule eval_evals_exec_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]]]) | |
| 667 | apply (subgoal_tac "G \<turnstile> (gx s, (gh s, gl s)) -es[\<succ>]vs-> (gx s', (gh s', gl s'))") | |
| 668 | apply assumption | |
| 669 | apply (simp add: gx_def gh_def gl_def surjective_pairing [THEN sym]) | |
| 670 | apply (case_tac E) | |
| 14045 | 671 | apply (simp add: gx_def gs_def gh_def gl_def surjective_pairing [THEN sym]) | 
| 13673 | 672 | done | 
| 673 | ||
| 674 | lemma eval_of_class: "\<lbrakk> G \<turnstile> s -e\<succ>a'-> s'; E \<turnstile> e :: Class C; | |
| 14045 | 675 | wf_java_prog G; s::\<preceq>E; gx s'=None; a' \<noteq> Null; G=prg E\<rbrakk> | 
| 13673 | 676 | \<Longrightarrow> (\<exists> lc. a' = Addr lc)" | 
| 677 | apply (case_tac s, case_tac s', simp) | |
| 14045 | 678 | apply (frule eval_type_sound, (simp add: gs_def)+) | 
| 13673 | 679 | apply (case_tac a') | 
| 680 | apply (auto simp: conf_def) | |
| 681 | done | |
| 682 | ||
| 683 | ||
| 684 | lemma dynT_subcls: | |
| 685 | "\<lbrakk> a' \<noteq> Null; G,h\<turnstile>a'::\<preceq> Class C; dynT = fst (the (h (the_Addr a'))); | |
| 14045 | 686 | is_class G dynT; ws_prog G \<rbrakk> \<Longrightarrow> G\<turnstile>dynT \<preceq>C C" | 
| 13673 | 687 | apply (case_tac "C = Object") | 
| 688 | apply (simp, rule subcls_C_Object, assumption+) | |
| 14045 | 689 | apply simp | 
| 13673 | 690 | apply (frule non_np_objD, auto) | 
| 691 | done | |
| 692 | ||
| 693 | ||
| 694 | lemma method_defined: "\<lbrakk> | |
| 695 | m = the (method (G, dynT) (mn, pTs)); | |
| 696 | dynT = fst (the (h a)); is_class G dynT; wf_java_prog G; | |
| 697 | a' \<noteq> Null; G,h\<turnstile>a'::\<preceq> Class C; a = the_Addr a'; | |
| 698 |   \<exists>pTsa md rT. max_spec G C (mn, pTsa) = {((md, rT), pTs)} \<rbrakk>
 | |
| 699 | \<Longrightarrow> (method (G, dynT) (mn, pTs)) = Some m" | |
| 700 | apply (erule exE)+ | |
| 701 | apply (drule singleton_in_set, drule max_spec2appl_meths) | |
| 702 | apply (simp add: appl_methds_def) | |
| 703 | apply ((erule exE)+, (erule conjE)+, (erule exE)+) | |
| 704 | apply (drule widen_methd) | |
| 705 | apply assumption | |
| 14045 | 706 | apply (rule dynT_subcls) apply assumption+ apply simp apply simp | 
| 13673 | 707 | apply (erule exE)+ apply simp | 
| 708 | done | |
| 709 | ||
| 710 | ||
| 711 | ||
| 712 | (**********************************************************************) | |
| 713 | ||
| 714 | ||
| 715 | (* 1. any difference between locvars_xstate \<dots> and L ??? *) | |
| 716 | (* 2. possibly skip env_of_jmb ??? *) | |
| 717 | theorem compiler_correctness: | |
| 718 | "wf_java_prog G \<Longrightarrow> | |
| 22271 | 719 | (G \<turnstile> xs -ex\<succ>val-> xs' \<longrightarrow> | 
| 13673 | 720 | gx xs = None \<longrightarrow> gx xs' = None \<longrightarrow> | 
| 721 | (\<forall> os CL S. | |
| 722 | (class_sig_defined G CL S) \<longrightarrow> | |
| 723 | (wtpd_expr (env_of_jmb G CL S) ex) \<longrightarrow> | |
| 14045 | 724 | (xs ::\<preceq> (env_of_jmb G CL S)) \<longrightarrow> | 
| 13673 | 725 |   ( {TranslComp.comp G, CL, S} \<turnstile>
 | 
| 726 |     {gh xs, os, (locvars_xstate G CL S xs)}
 | |
| 727 | >- (compExpr (gmb G CL S) ex) \<rightarrow> | |
| 728 |     {gh xs', val#os, locvars_xstate G CL S xs'}))) \<and> 
 | |
| 729 | ||
| 22271 | 730 | (G \<turnstile> xs -exs[\<succ>]vals-> xs' \<longrightarrow> | 
| 13673 | 731 | gx xs = None \<longrightarrow> gx xs' = None \<longrightarrow> | 
| 732 | (\<forall> os CL S. | |
| 733 | (class_sig_defined G CL S) \<longrightarrow> | |
| 734 | (wtpd_exprs (env_of_jmb G CL S) exs) \<longrightarrow> | |
| 14045 | 735 | (xs::\<preceq>(env_of_jmb G CL S)) \<longrightarrow> | 
| 13673 | 736 |   ( {TranslComp.comp G, CL, S} \<turnstile>
 | 
| 737 |     {gh xs, os, (locvars_xstate G CL S xs)}
 | |
| 738 | >- (compExprs (gmb G CL S) exs) \<rightarrow> | |
| 739 |     {gh xs', (rev vals)@os, (locvars_xstate G CL S xs')}))) \<and> 
 | |
| 740 | ||
| 22271 | 741 | (G \<turnstile> xs -st-> xs' \<longrightarrow> | 
| 13673 | 742 | gx xs = None \<longrightarrow> gx xs' = None \<longrightarrow> | 
| 743 | (\<forall> os CL S. | |
| 744 | (class_sig_defined G CL S) \<longrightarrow> | |
| 745 | (wtpd_stmt (env_of_jmb G CL S) st) \<longrightarrow> | |
| 14045 | 746 | (xs::\<preceq>(env_of_jmb G CL S)) \<longrightarrow> | 
| 13673 | 747 |   ( {TranslComp.comp G, CL, S} \<turnstile>
 | 
| 748 |     {gh xs, os, (locvars_xstate G CL S xs)}
 | |
| 749 | >- (compStmt (gmb G CL S) st) \<rightarrow> | |
| 750 |     {gh xs', os, (locvars_xstate G CL S xs')})))"
 | |
| 751 | apply (rule Eval.eval_evals_exec.induct) | |
| 752 | ||
| 753 | (* case XcptE *) | |
| 754 | apply simp | |
| 755 | ||
| 756 | (* case NewC *) | |
| 14045 | 757 | apply clarify | 
| 758 | apply (frule wf_prog_ws_prog [THEN wf_subcls1]) (* establish wf ((subcls1 G)^-1) *) | |
| 13673 | 759 | apply (simp add: c_hupd_hp_invariant) | 
| 760 | apply (rule progression_one_step) | |
| 761 | apply (rotate_tac 1, drule sym) (* reverse equation (a, None) = new_Addr (fst s) *) | |
| 762 | apply (simp add: locvars_xstate_def locvars_locals_def comp_fields) | |
| 763 | ||
| 764 | ||
| 765 | (* case Cast *) | |
| 766 | apply (intro allI impI) | |
| 767 | apply simp | |
| 768 | apply (frule raise_if_NoneD) | |
| 769 | apply (frule wtpd_expr_cast) | |
| 770 | apply simp | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 771 | apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, simp) | 
| 13673 | 772 | apply blast | 
| 773 | apply (rule progression_one_step) | |
| 774 | apply (simp add: raise_system_xcpt_def gh_def comp_cast_ok) | |
| 775 | ||
| 776 | ||
| 777 | (* case Lit *) | |
| 778 | apply simp | |
| 779 | apply (intro strip) | |
| 780 | apply (rule progression_one_step) | |
| 781 | apply simp | |
| 782 | ||
| 783 | ||
| 784 | (* case BinOp *) | |
| 785 | apply (intro allI impI) | |
| 786 | apply (frule_tac xs=s1 in eval_xcpt, assumption) (* establish (gx s1 = None) *) | |
| 787 | apply (frule wtpd_expr_binop) | |
| 14045 | 788 | (* establish (s1::\<preceq> \<dots>) *) | 
| 13673 | 789 | apply (frule_tac e=e1 in state_ok_eval) apply (simp (no_asm_simp)) apply simp apply (simp (no_asm_use) only: env_of_jmb_fst) | 
| 790 | ||
| 791 | ||
| 792 | apply (simp (no_asm_use) only: compExpr_compExprs.simps) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 793 | apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e1" in progression_transitive, simp) apply blast | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 794 | apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e2" in progression_transitive, simp) apply blast | 
| 13673 | 795 | apply (case_tac bop) | 
| 796 | (*subcase bop = Eq *) apply simp apply (rule progression_Eq) | |
| 797 | (*subcase bop = Add *) apply simp apply (rule progression_one_step) apply simp | |
| 798 | ||
| 799 | ||
| 800 | (* case LAcc *) | |
| 801 | apply simp | |
| 802 | apply (intro strip) | |
| 803 | apply (rule progression_one_step) | |
| 804 | apply (simp add: locvars_xstate_def locvars_locals_def) | |
| 805 | apply (frule wtpd_expr_lacc) | |
| 806 | apply assumption | |
| 807 | apply (simp add: gl_def) | |
| 808 | apply (erule select_at_index) | |
| 809 | ||
| 810 | ||
| 811 | (* case LAss *) | |
| 812 | apply (intro allI impI) | |
| 813 | apply (frule wtpd_expr_lass, erule conjE, erule conjE) | |
| 814 | apply (simp add: compExpr_compExprs.simps) | |
| 815 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 816 | apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) | 
| 13673 | 817 | apply blast | 
| 818 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 819 | apply (rule_tac ?instrs0.0 = "[Dup]" in progression_transitive, simp) | 
| 13673 | 820 | apply (rule progression_one_step) | 
| 821 | apply (simp add: gh_def) | |
| 822 | apply (rule conjI, simp)+ apply simp | |
| 823 | apply (rule progression_one_step) | |
| 824 | apply (simp add: gh_def) | |
| 825 | (* the following falls out of the general scheme *) | |
| 826 | apply (frule wtpd_expr_lacc) apply assumption | |
| 827 | apply (rule update_at_index) | |
| 828 | apply (rule distinct_method_if_class_sig_defined) apply assumption | |
| 829 | apply assumption apply simp apply assumption | |
| 830 | ||
| 831 | ||
| 832 | (* case FAcc *) | |
| 833 | apply (intro allI impI) | |
| 834 | (* establish x1 = None \<and> a' \<noteq> Null *) | |
| 835 | apply (simp (no_asm_use) only: gx_conv, frule np_NoneD) | |
| 836 | apply (frule wtpd_expr_facc) | |
| 837 | ||
| 838 | apply (simp (no_asm_use) only: compExpr_compExprs.simps) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 839 | apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) | 
| 13673 | 840 | apply blast | 
| 841 | apply (rule progression_one_step) | |
| 842 | apply (simp add: gh_def) | |
| 843 | apply (case_tac "(the (fst s1 (the_Addr a')))") | |
| 844 | apply (simp add: raise_system_xcpt_def) | |
| 845 | ||
| 846 | ||
| 847 | (* case FAss *) | |
| 848 | apply (intro allI impI) | |
| 849 | apply (frule wtpd_expr_fass) apply (erule conjE) apply (frule wtpd_expr_facc) | |
| 850 | apply (simp only: c_hupd_xcpt_invariant) (* establish x2 = None *) | |
| 851 | (* establish x1 = None and a' \<noteq> Null *) | |
| 852 | apply (frule_tac xs="(np a' x1, s1)" in eval_xcpt) | |
| 853 | apply (simp only: gx_conv, simp only: gx_conv, frule np_NoneD, erule conjE) | |
| 854 | ||
| 855 | ||
| 14045 | 856 | (* establish ((Norm s1)::\<preceq> \<dots>) *) | 
| 857 | apply (frule_tac e=e1 in state_ok_eval) apply (simp (no_asm_simp) only: env_of_jmb_fst) | |
| 858 | apply assumption apply (simp (no_asm_use) only: env_of_jmb_fst) | |
| 13673 | 859 | |
| 860 | apply (simp only: compExpr_compExprs.simps) | |
| 861 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 862 | apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e1)" in progression_transitive, rule HOL.refl) | 
| 13673 | 863 | apply fast (* blast does not seem to work - why? *) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 864 | apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e2)" in progression_transitive, rule HOL.refl) | 
| 13673 | 865 | apply fast | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 866 | apply (rule_tac ?instrs0.0 = "[Dup_x1]" and ?instrs1.0 = "[Putfield fn T]" in progression_transitive, simp) | 
| 13673 | 867 | |
| 868 | (* Dup_x1 *) | |
| 869 | apply (rule progression_one_step) | |
| 870 | apply (simp add: gh_def) | |
| 871 | apply (rule conjI, simp)+ apply simp | |
| 872 | ||
| 873 | ||
| 874 | (* Putfield \<longrightarrow> still looks nasty*) | |
| 875 | apply (rule progression_one_step) | |
| 876 | apply simp | |
| 877 | apply (case_tac "(the (fst s2 (the_Addr a')))") | |
| 878 | apply (simp add: c_hupd_hp_invariant) | |
| 879 | apply (case_tac s2) | |
| 880 | apply (simp add: c_hupd_conv raise_system_xcpt_def) | |
| 881 | apply (rule locvars_xstate_par_dep, rule HOL.refl) | |
| 882 | ||
| 883 | defer (* method call *) | |
| 884 | ||
| 885 | (* case XcptEs *) | |
| 886 | apply simp | |
| 887 | ||
| 888 | (* case Nil *) | |
| 889 | apply (simp add: compExpr_compExprs.simps) | |
| 890 | apply (intro strip) | |
| 891 | apply (rule progression_refl) | |
| 892 | ||
| 893 | (* case Cons *) | |
| 894 | apply (intro allI impI) | |
| 895 | apply (frule_tac xs=s1 in evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *) | |
| 896 | apply (frule wtpd_exprs_cons) | |
| 14045 | 897 | (* establish ((Norm s0)::\<preceq> \<dots>) *) | 
| 13673 | 898 | apply (frule_tac e=e in state_ok_eval) apply (simp (no_asm_simp) only: env_of_jmb_fst) apply simp apply (simp (no_asm_use) only: env_of_jmb_fst) | 
| 899 | ||
| 900 | apply simp | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 901 | apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) | 
| 13673 | 902 | apply fast | 
| 903 | apply fast | |
| 904 | ||
| 905 | ||
| 906 | (* case Statement: exception *) | |
| 907 | apply simp | |
| 908 | ||
| 909 | (* case Skip *) | |
| 910 | apply (intro allI impI) | |
| 911 | apply simp | |
| 912 | apply (rule progression_refl) | |
| 913 | ||
| 914 | (* case Expr *) | |
| 915 | apply (intro allI impI) | |
| 916 | apply (frule wtpd_stmt_expr) | |
| 917 | apply simp | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 918 | apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) | 
| 13673 | 919 | apply fast | 
| 920 | apply (rule progression_one_step) | |
| 921 | apply simp | |
| 922 | ||
| 923 | (* case Comp *) | |
| 924 | apply (intro allI impI) | |
| 925 | apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *) | |
| 926 | apply (frule wtpd_stmt_comp) | |
| 927 | ||
| 14045 | 928 | (* establish (s1::\<preceq> \<dots>) *) | 
| 13673 | 929 | apply (frule_tac st=c1 in state_ok_exec) apply (simp (no_asm_simp) only: env_of_jmb_fst) apply simp apply (simp (no_asm_use) only: env_of_jmb_fst) | 
| 930 | ||
| 931 | apply simp | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 932 | apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, rule HOL.refl) | 
| 13673 | 933 | apply fast | 
| 934 | apply fast | |
| 935 | ||
| 936 | ||
| 937 | (* case Cond *) | |
| 938 | apply (intro allI impI) | |
| 939 | apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *) | |
| 940 | apply (frule wtpd_stmt_cond, (erule conjE)+) | |
| 14045 | 941 | (* establish (s1::\<preceq> \<dots>) *) | 
| 13673 | 942 | apply (frule_tac e=e in state_ok_eval) | 
| 943 | apply (simp (no_asm_simp) only: env_of_jmb_fst) | |
| 944 | apply assumption | |
| 945 | apply (simp (no_asm_use) only: env_of_jmb_fst) | |
| 946 | (* establish G,gh s1\<turnstile>v::\<preceq>PrimT Boolean *) | |
| 947 | apply (frule eval_conf, assumption+, rule env_of_jmb_fst) | |
| 948 | apply (frule conf_bool) (* establish \<exists>b. v = Bool b *) | |
| 949 | apply (erule exE) | |
| 950 | ||
| 951 | apply simp | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 952 | apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp)) | 
| 13673 | 953 | apply (rule progression_one_step, simp) | 
| 954 | apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl) | |
| 955 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 956 | apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl) | 
| 13673 | 957 | apply fast | 
| 958 | ||
| 959 | apply (case_tac b) | |
| 960 | (* case b= True *) | |
| 961 | apply simp | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 962 | apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c1)))]" in progression_transitive, simp) | 
| 13673 | 963 | apply (rule progression_one_step) apply simp | 
| 964 | apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 965 | apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, simp) | 
| 13673 | 966 | apply fast | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 967 | apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression) | 
| 13673 | 968 | apply (simp, rule conjI, (rule HOL.refl)+) | 
| 20432 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20272diff
changeset | 969 | apply simp apply (rule conjI, simp) apply (simp add: nat_add_distrib) | 
| 13673 | 970 | apply (rule progression_refl) | 
| 971 | ||
| 972 | (* case b= False *) | |
| 973 | apply simp | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 974 | apply (rule_tac ?instrs1.0 = "compStmt (gmb G CL S) c2" in jump_fwd_progression) | 
| 13673 | 975 | apply (simp, rule conjI, (rule HOL.refl)+) | 
| 20432 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20272diff
changeset | 976 | apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib) | 
| 13673 | 977 | apply fast | 
| 978 | ||
| 979 | (* case exit Loop *) | |
| 980 | apply (intro allI impI) | |
| 981 | apply (frule wtpd_stmt_loop, (erule conjE)+) | |
| 982 | ||
| 983 | (* establish G,gh s1\<turnstile>v::\<preceq>PrimT Boolean *) | |
| 984 | apply (frule eval_conf, assumption+, rule env_of_jmb_fst) | |
| 985 | apply (frule conf_bool) (* establish \<exists>b. v = Bool b *) | |
| 986 | apply (erule exE) | |
| 987 | apply (case_tac b) | |
| 988 | ||
| 989 | (* case b= True \<longrightarrow> contradiction *) | |
| 990 | apply simp | |
| 991 | ||
| 992 | (* case b= False *) | |
| 993 | apply simp | |
| 994 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 995 | apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp)) | 
| 13673 | 996 | apply (rule progression_one_step) | 
| 997 | apply simp | |
| 998 | apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl) | |
| 999 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 1000 | apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl) | 
| 13673 | 1001 | apply fast | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 1002 | apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression) | 
| 13673 | 1003 | apply (simp, rule conjI, rule HOL.refl, rule HOL.refl) | 
| 20432 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20272diff
changeset | 1004 | apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib) | 
| 13673 | 1005 | apply (rule progression_refl) | 
| 1006 | ||
| 1007 | ||
| 1008 | (* case continue Loop *) | |
| 1009 | apply (intro allI impI) | |
| 1010 | apply (frule_tac xs=s2 in exec_xcpt, assumption) (* establish (gx s2 = None) *) | |
| 1011 | apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *) | |
| 1012 | apply (frule wtpd_stmt_loop, (erule conjE)+) | |
| 1013 | ||
| 14045 | 1014 | (* establish (s1::\<preceq> \<dots>) *) | 
| 13673 | 1015 | apply (frule_tac e=e in state_ok_eval) apply (simp (no_asm_simp) only: env_of_jmb_fst) apply simp apply (simp (no_asm_use) only: env_of_jmb_fst) | 
| 14045 | 1016 | (* establish (s2::\<preceq> \<dots>) *) | 
| 13673 | 1017 | apply (frule_tac xs=s1 and st=c in state_ok_exec) | 
| 1018 | apply (simp (no_asm_simp) only: env_of_jmb_fst) apply assumption apply (simp (no_asm_use) only: env_of_jmb_fst) | |
| 1019 | ||
| 1020 | (* establish G,gh s1\<turnstile>v::\<preceq>PrimT Boolean *) | |
| 1021 | apply (frule eval_conf, assumption+, rule env_of_jmb_fst) | |
| 1022 | apply (frule conf_bool) (* establish \<exists>b. v = Bool b *) | |
| 1023 | apply (erule exE) | |
| 1024 | ||
| 1025 | apply simp | |
| 1026 | apply (rule jump_bwd_progression) | |
| 1027 | apply simp | |
| 1028 | apply (rule conjI, (rule HOL.refl)+) | |
| 1029 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 1030 | apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp)) | 
| 13673 | 1031 | apply (rule progression_one_step) | 
| 1032 | apply simp | |
| 1033 | apply (rule conjI, simp)+ apply simp | |
| 1034 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 1035 | apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl) | 
| 13673 | 1036 | apply fast | 
| 1037 | ||
| 1038 | apply (case_tac b) | |
| 1039 | (* case b= True *) | |
| 1040 | apply simp | |
| 1041 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 1042 | apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c)))]" in progression_transitive, simp) | 
| 13673 | 1043 | apply (rule progression_one_step) apply simp | 
| 1044 | apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl) | |
| 1045 | apply fast | |
| 1046 | ||
| 1047 | (* case b= False \<longrightarrow> contradiction*) | |
| 1048 | apply simp | |
| 1049 | ||
| 1050 | apply (rule jump_bwd_one_step) | |
| 1051 | apply simp | |
| 1052 | apply blast | |
| 1053 | ||
| 1054 | (*****) | |
| 14045 | 1055 | (* case method call *) | 
| 13673 | 1056 | |
| 1057 | apply (intro allI impI) | |
| 1058 | ||
| 1059 | apply (frule_tac xs=s3 in eval_xcpt, simp only: gx_conv) (* establish gx s3 = None *) | |
| 1060 | apply (frule exec_xcpt, assumption, simp (no_asm_use) only: gx_conv, frule np_NoneD) (* establish x = None \<and> a' \<noteq> Null *) | |
| 1061 | apply (frule evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *) | |
| 1062 | ||
| 1063 | apply (frule wtpd_expr_call, (erule conjE)+) | |
| 1064 | ||
| 1065 | ||
| 1066 | (* assumptions about state_ok and is_class *) | |
| 1067 | ||
| 14045 | 1068 | (* establish s1::\<preceq> (env_of_jmb G CL S) *) | 
| 13673 | 1069 | apply (frule_tac xs="Norm s0" and e=e in state_ok_eval) | 
| 1070 | apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp (no_asm_use) only: env_of_jmb_fst) | |
| 1071 | ||
| 14045 | 1072 | (* establish (x, h, l)::\<preceq>(env_of_jmb G CL S) *) | 
| 13673 | 1073 | apply (frule_tac xs=s1 and xs'="(x, h, l)" in state_ok_evals) | 
| 1074 | apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp only: env_of_jmb_fst) | |
| 1075 | ||
| 1076 | (* establish \<exists> lc. a' = Addr lc *) | |
| 1077 | apply (frule (5) eval_of_class, rule env_of_jmb_fst [THEN sym]) | |
| 1078 | apply (subgoal_tac "G,h \<turnstile> a' ::\<preceq> Class C") | |
| 1079 | apply (subgoal_tac "is_class G dynT") | |
| 1080 | ||
| 1081 | (* establish method (G, dynT) (mn, pTs) = Some(md, rT, pns, lvars, blk, res) *) | |
| 1082 | apply (drule method_defined, assumption+) | |
| 1083 | apply (simp only: env_of_jmb_fst) | |
| 1084 | apply ((erule exE)+, erule conjE, (rule exI)+, assumption) | |
| 1085 | ||
| 1086 | apply (subgoal_tac "is_class G md") | |
| 1087 | apply (subgoal_tac "G\<turnstile>Class dynT \<preceq> Class md") | |
| 1088 | apply (subgoal_tac " method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res)") | |
| 1089 | apply (subgoal_tac "list_all2 (conf G h) pvs pTs") | |
| 1090 | ||
| 14045 | 1091 | (* establish (np a' x, h, init_vars lvars(pns[\<mapsto>]pvs)(This\<mapsto>a'))::\<preceq>(env_of_jmb G md (mn, pTs)) *) | 
| 13673 | 1092 | apply (subgoal_tac "G,h \<turnstile> a' ::\<preceq> Class dynT") | 
| 1093 | apply (frule (2) conf_widen) | |
| 1094 | apply (frule state_ok_init, assumption+) | |
| 1095 | ||
| 1096 | apply (subgoal_tac "class_sig_defined G md (mn, pTs)") | |
| 1097 | apply (frule wtpd_blk, assumption, assumption) | |
| 1098 | apply (frule wtpd_res, assumption, assumption) | |
| 14045 | 1099 | apply (subgoal_tac "s3::\<preceq>(env_of_jmb G md (mn, pTs))") | 
| 13673 | 1100 | |
| 14045 | 1101 | apply (subgoal_tac "method (TranslComp.comp G, md) (mn, pTs) = | 
| 1102 | Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))") | |
| 1103 | prefer 2 apply (simp add: wf_prog_ws_prog [THEN comp_method]) | |
| 1104 | apply (subgoal_tac "method (TranslComp.comp G, dynT) (mn, pTs) = | |
| 1105 | Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))") | |
| 1106 | prefer 2 apply (simp add: wf_prog_ws_prog [THEN comp_method]) | |
| 13673 | 1107 | apply (simp only: fst_conv snd_conv) | 
| 1108 | ||
| 1109 | (* establish length pns = length pTs *) | |
| 1110 | apply (frule method_preserves_length, assumption, assumption) | |
| 1111 | (* establish length pvs = length ps *) | |
| 1112 | apply (frule evals_preserves_length [THEN sym]) | |
| 1113 | ||
| 1114 | (** start evaluating subexpressions **) | |
| 1115 | apply (simp (no_asm_use) only: compExpr_compExprs.simps) | |
| 1116 | ||
| 1117 | (* evaluate e *) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 1118 | apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) | 
| 13673 | 1119 | apply fast | 
| 1120 | ||
| 1121 | (* evaluate parameters *) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 1122 | apply (rule_tac ?instrs0.0 = "compExprs (gmb G CL S) ps" in progression_transitive, rule HOL.refl) | 
| 13673 | 1123 | apply fast | 
| 1124 | ||
| 1125 | (* invokation *) | |
| 1126 | apply (rule progression_call) | |
| 1127 | apply (intro allI impI conjI) | |
| 1128 | (* execute Invoke statement *) | |
| 1129 | apply (simp (no_asm_use) only: exec_instr.simps) | |
| 1130 | apply (erule thin_rl, erule thin_rl, erule thin_rl) | |
| 1131 | apply (simp add: compMethod_def raise_system_xcpt_def) | |
| 1132 | apply (rule conjI, simp)+ apply (rule HOL.refl) | |
| 1133 | ||
| 1134 | (* get instructions of invoked method *) | |
| 1135 | apply (simp (no_asm_simp) add: gis_def gmb_def compMethod_def) | |
| 1136 | ||
| 1137 | (* var. initialization *) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 1138 | apply (rule_tac ?instrs0.0 = "(compInitLvars (pns, lvars, blk, res) lvars)" in progression_transitive, rule HOL.refl) | 
| 13673 | 1139 | apply (rule_tac C=md in progression_lvar_init, assumption, assumption, assumption) | 
| 1140 | apply (simp (no_asm_simp)) (* length pns = length pvs *) | |
| 1141 | apply (simp (no_asm_simp)) (* length lvars = length (replicate (length lvars) arbitrary) *) | |
| 1142 | ||
| 1143 | ||
| 1144 | (* body statement *) | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14143diff
changeset | 1145 | apply (rule_tac ?instrs0.0 = "compStmt (pns, lvars, blk, res) blk" in progression_transitive, rule HOL.refl) | 
| 13673 | 1146 | apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)") | 
| 1147 | apply (simp (no_asm_simp)) | |
| 15866 | 1148 | apply (simp only: gh_conv) | 
| 1149 | apply (drule mp [OF _ TrueI])+ | |
| 1150 | apply (erule allE, erule allE, erule allE, erule impE, assumption)+ | |
| 1151 | apply ((erule impE, assumption)+, assumption) | |
| 15860 | 1152 | |
| 13673 | 1153 | apply (simp (no_asm_use)) | 
| 1154 | apply (simp (no_asm_simp) add: gmb_def) | |
| 1155 | ||
| 1156 | (* return expression *) | |
| 1157 | apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)") | |
| 1158 | apply (simp (no_asm_simp)) | |
| 1159 | apply (simp only: gh_conv) | |
| 1160 | apply ((drule mp, rule TrueI)+, (drule spec)+, (drule mp, assumption)+, assumption) | |
| 1161 | apply (simp (no_asm_use)) | |
| 1162 | apply (simp (no_asm_simp) add: gmb_def) | |
| 1163 | ||
| 1164 | (* execute return statement *) | |
| 1165 | apply (simp (no_asm_use) add: gh_def locvars_xstate_def gl_def del: drop_append) | |
| 1166 | apply (subgoal_tac "rev pvs @ a' # os = (rev (a' # pvs)) @ os") | |
| 1167 | apply (simp only: drop_append) | |
| 1168 | apply (simp (no_asm_simp)) | |
| 1169 | apply (simp (no_asm_simp)) | |
| 1170 | ||
| 14045 | 1171 | (* show s3::\<preceq>\<dots> *) | 
| 13673 | 1172 | apply (rule_tac xs = "(np a' x, h, init_vars lvars(pns[\<mapsto>]pvs)(This\<mapsto>a'))" and st=blk in state_ok_exec) | 
| 1173 | apply assumption apply (simp (no_asm_simp) only: env_of_jmb_fst) | |
| 1174 | apply assumption apply (simp (no_asm_use) only: env_of_jmb_fst) | |
| 1175 | ||
| 1176 | (* show class_sig_defined G md (mn, pTs) *) | |
| 1177 | apply (simp (no_asm_simp) add: class_sig_defined_def) | |
| 1178 | ||
| 1179 | (* show G,h \<turnstile> a' ::\<preceq> Class dynT *) | |
| 1180 | apply (frule non_npD) apply assumption | |
| 1181 | apply (erule exE)+ apply simp | |
| 1182 | apply (rule conf_obj_AddrI) apply simp | |
| 1183 | apply (rule conjI, (rule HOL.refl)+) | |
| 1184 | apply (rule widen_Class_Class [THEN iffD1], rule widen.refl) | |
| 1185 | ||
| 1186 | ||
| 1187 | (* show list_all2 (conf G h) pvs pTs *) | |
| 1188 | apply (erule exE)+ apply (erule conjE)+ | |
| 1189 | apply (rule_tac Ts="pTsa" in conf_list_gext_widen) apply assumption | |
| 22271 | 1190 | apply (subgoal_tac "G \<turnstile> (gx s1, gs s1) -ps[\<succ>]pvs-> (x, h, l)") | 
| 13673 | 1191 | apply (frule_tac E="env_of_jmb G CL S" in evals_type_sound) | 
| 14143 | 1192 | apply assumption+ | 
| 13673 | 1193 | apply (simp only: env_of_jmb_fst) | 
| 1194 | apply (simp add: conforms_def xconf_def gs_def) | |
| 1195 | apply simp | |
| 1196 | apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [THEN sym]) | |
| 1197 | apply (simp (no_asm_use) only: ty_exprs_list_all2) apply simp | |
| 14143 | 1198 | apply simp | 
| 13673 | 1199 | apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [THEN sym]) | 
| 1200 | (* list_all2 (\<lambda>T T'. G \<turnstile> T \<preceq> T') pTsa pTs *) | |
| 1201 | apply (rule max_spec_widen, simp only: env_of_jmb_fst) | |
| 1202 | ||
| 1203 | ||
| 1204 | (* show method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res) *) | |
| 14045 | 1205 | apply (frule wf_prog_ws_prog [THEN method_in_md [THEN conjunct2]], assumption+) | 
| 13673 | 1206 | |
| 1207 | (* show G\<turnstile>Class dynT \<preceq> Class md *) | |
| 1208 | apply (simp (no_asm_use) only: widen_Class_Class) | |
| 1209 | apply (rule method_wf_mdecl [THEN conjunct1], assumption+) | |
| 1210 | ||
| 1211 | (* is_class G md *) | |
| 14045 | 1212 | apply (rule wf_prog_ws_prog [THEN method_in_md [THEN conjunct1]], assumption+) | 
| 13673 | 1213 | |
| 1214 | (* show is_class G dynT *) | |
| 1215 | apply (frule non_npD) apply assumption | |
| 1216 | apply (erule exE)+ apply (erule conjE)+ | |
| 1217 | apply simp | |
| 1218 | apply (rule subcls_is_class2) apply assumption | |
| 1219 | apply (frule class_expr_is_class) apply (simp only: env_of_jmb_fst) | |
| 14045 | 1220 | apply (rule wf_prog_ws_prog, assumption) | 
| 13673 | 1221 | apply (simp only: env_of_jmb_fst) | 
| 1222 | ||
| 1223 | (* show G,h \<turnstile> a' ::\<preceq> Class C *) | |
| 1224 | apply (simp only: wtpd_exprs_def, erule exE) | |
| 1225 | apply (frule evals_preserves_conf) | |
| 1226 | apply (rule eval_conf, assumption+) | |
| 1227 | apply (rule env_of_jmb_fst, assumption+) | |
| 1228 | apply (rule env_of_jmb_fst) | |
| 1229 | apply (simp only: gh_conv) | |
| 1230 | done | |
| 1231 | ||
| 1232 | ||
| 1233 | theorem compiler_correctness_eval: " | |
| 1234 | \<lbrakk> G \<turnstile> (None,hp,loc) -ex \<succ> val-> (None,hp',loc'); | |
| 1235 | wf_java_prog G; | |
| 1236 | class_sig_defined G C S; | |
| 1237 | wtpd_expr (env_of_jmb G C S) ex; | |
| 1238 | (None,hp,loc) ::\<preceq> (env_of_jmb G C S) \<rbrakk> \<Longrightarrow> | |
| 1239 |   {(TranslComp.comp G), C, S} \<turnstile> 
 | |
| 1240 |     {hp, os, (locvars_locals G C S loc)}
 | |
| 1241 | >- (compExpr (gmb G C S) ex) \<rightarrow> | |
| 1242 |     {hp', val#os, (locvars_locals G C S loc')}"
 | |
| 1243 | apply (frule compiler_correctness [THEN conjunct1]) | |
| 14045 | 1244 | apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def) | 
| 13673 | 1245 | done | 
| 1246 | ||
| 1247 | theorem compiler_correctness_exec: " | |
| 22271 | 1248 | \<lbrakk> G \<turnstile> Norm (hp, loc) -st-> Norm (hp', loc'); | 
| 13673 | 1249 | wf_java_prog G; | 
| 1250 | class_sig_defined G C S; | |
| 1251 | wtpd_stmt (env_of_jmb G C S) st; | |
| 1252 | (None,hp,loc) ::\<preceq> (env_of_jmb G C S) \<rbrakk> \<Longrightarrow> | |
| 1253 |   {(TranslComp.comp G), C, S} \<turnstile> 
 | |
| 1254 |     {hp, os, (locvars_locals G C S loc)}
 | |
| 1255 | >- (compStmt (gmb G C S) st) \<rightarrow> | |
| 1256 |     {hp', os, (locvars_locals G C S loc')}"
 | |
| 1257 | apply (frule compiler_correctness [THEN conjunct2 [THEN conjunct2]]) | |
| 14045 | 1258 | apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def) | 
| 13673 | 1259 | done | 
| 1260 | ||
| 1261 | (**********************************************************************) | |
| 1262 | ||
| 1263 | ||
| 1264 | (* reinstall pair splits *) | |
| 1265 | declare split_paired_All [simp] split_paired_Ex [simp] | |
| 24074 | 1266 | declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *}
 | 
| 13673 | 1267 | |
| 14045 | 1268 | declare wf_prog_ws_prog [simp del] | 
| 1269 | ||
| 13673 | 1270 | end |