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