| author | haftmann | 
| Wed, 09 May 2007 07:53:06 +0200 | |
| changeset 22886 | cdff6ef76009 | 
| parent 22271 | 51a80e238b29 | 
| child 23757 | 087b0a241557 | 
| 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: 
15076 
diff
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)  | 
|
121  | 
apply (simp only: list.simps)  | 
|
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: 
14143 
diff
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: 
14143 
diff
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: 
14143 
diff
changeset
 | 
309  | 
?P1.0 = "wf_mdecl wf_mb G Ca" and  | 
| 
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
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)  | 
|
| 22271 | 407  | 
apply (ind_cases2 "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: 
14143 
diff
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: 
13673 
diff
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: 
14143 
diff
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: 
13673 
diff
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]  | 
|
479  | 
ML_setup {*
 | 
|
| 17876 | 480  | 
change_simpset (fn ss => ss delloop "split_all_tac");  | 
| 13673 | 481  | 
*}  | 
482  | 
||
483  | 
lemma distinct_method: "\<lbrakk> wf_java_prog G; is_class G C;  | 
|
484  | 
method (G, C) S = Some (D, rT, pns, lvars, blk, res) \<rbrakk> \<Longrightarrow>  | 
|
485  | 
distinct (gjmb_plns (gmb G C S))"  | 
|
486  | 
apply (frule method_wf_mdecl [THEN conjunct2], assumption, assumption)  | 
|
487  | 
apply (case_tac S)  | 
|
488  | 
apply (simp add: wf_mdecl_def wf_java_mdecl_def galldefs distinct_append)  | 
|
489  | 
apply (simp add: unique_def map_of_in_set)  | 
|
490  | 
apply blast  | 
|
491  | 
done  | 
|
492  | 
||
493  | 
lemma distinct_method_if_class_sig_defined :  | 
|
494  | 
"\<lbrakk> wf_java_prog G; class_sig_defined G C S \<rbrakk> \<Longrightarrow>  | 
|
495  | 
distinct (gjmb_plns (gmb G C S))"  | 
|
496  | 
by (auto intro: distinct_method simp: class_sig_defined_def)  | 
|
497  | 
||
498  | 
||
499  | 
lemma method_yields_wf_java_mdecl: "\<lbrakk> wf_java_prog G; is_class G C;  | 
|
500  | 
method (G, C) S = Some (D, rT, pns, lvars, blk, res) \<rbrakk> \<Longrightarrow>  | 
|
501  | 
wf_java_mdecl G D (S,rT,(pns,lvars,blk,res))"  | 
|
502  | 
apply (frule method_wf_mdecl)  | 
|
503  | 
apply (auto simp: wf_mdecl_def)  | 
|
504  | 
done  | 
|
505  | 
||
506  | 
(**********************************************************************)  | 
|
507  | 
||
508  | 
||
509  | 
lemma progression_lvar_init_aux [rule_format (no_asm)]: "  | 
|
510  | 
\<forall> zs prfx lvals lvars0.  | 
|
511  | 
lvars0 = (zs @ lvars) \<longrightarrow>  | 
|
512  | 
(disjoint_varnames pns lvars0 \<longrightarrow>  | 
|
513  | 
(length lvars = length lvals) \<longrightarrow>  | 
|
514  | 
(Suc(length pns + length zs) = length prfx) \<longrightarrow>  | 
|
515  | 
   ({cG, D, S} \<turnstile> 
 | 
|
516  | 
    {h, os, (prfx @ lvals)}
 | 
|
517  | 
>- (concat (map (compInit (pns, lvars0, blk, res)) lvars)) \<rightarrow>  | 
|
518  | 
    {h, os, (prfx @ (map (\<lambda>p. (default_val (snd p))) lvars))}))"
 | 
|
519  | 
apply simp  | 
|
520  | 
apply (induct lvars)  | 
|
521  | 
apply (clarsimp, rule progression_refl)  | 
|
522  | 
apply (intro strip)  | 
|
523  | 
apply (case_tac lvals) apply simp  | 
|
524  | 
apply (simp (no_asm_simp) )  | 
|
525  | 
||
| 
15236
 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 
nipkow 
parents: 
15097 
diff
changeset
 | 
526  | 
apply (rule_tac ?lvars1.0 = "(prfx @ [default_val (snd a)]) @ list" in progression_transitive, rule HOL.refl)  | 
| 13673 | 527  | 
apply (case_tac a) apply (simp (no_asm_simp) add: compInit_def)  | 
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
528  | 
apply (rule_tac ?instrs0.0 = "[load_default_val b]" in progression_transitive, simp)  | 
| 13673 | 529  | 
apply (rule progression_one_step)  | 
530  | 
apply (simp (no_asm_simp) add: load_default_val_def)  | 
|
531  | 
apply (rule conjI, simp)+ apply (rule HOL.refl)  | 
|
532  | 
||
533  | 
apply (rule progression_one_step)  | 
|
534  | 
apply (simp (no_asm_simp))  | 
|
535  | 
apply (rule conjI, simp)+  | 
|
536  | 
apply (simp add: index_of_var2)  | 
|
537  | 
apply (drule_tac x="zs @ [a]" in spec) (* instantiate zs *)  | 
|
538  | 
apply (drule mp, simp)  | 
|
539  | 
apply (drule_tac x="(prfx @ [default_val (snd a)])" in spec) (* instantiate prfx *)  | 
|
540  | 
apply auto  | 
|
541  | 
done  | 
|
542  | 
||
543  | 
lemma progression_lvar_init [rule_format (no_asm)]:  | 
|
544  | 
"\<lbrakk> wf_java_prog G; is_class G C;  | 
|
545  | 
method (G, C) S = Some (D, rT, pns, lvars, blk, res) \<rbrakk> \<Longrightarrow>  | 
|
546  | 
length pns = length pvs \<longrightarrow>  | 
|
547  | 
(\<forall> lvals.  | 
|
548  | 
length lvars = length lvals \<longrightarrow>  | 
|
549  | 
   {cG, D, S} \<turnstile>
 | 
|
550  | 
   {h, os, (a' # pvs @ lvals)}
 | 
|
551  | 
>- (compInitLvars (pns, lvars, blk, res) lvars) \<rightarrow>  | 
|
552  | 
   {h, os, (locvars_xstate G C S (Norm (h, init_vars lvars(pns[\<mapsto>]pvs)(This\<mapsto>a'))))})"
 | 
|
553  | 
apply (simp only: compInitLvars_def)  | 
|
554  | 
apply (frule method_yields_wf_java_mdecl, assumption, assumption)  | 
|
555  | 
||
556  | 
apply (simp only: wf_java_mdecl_def)  | 
|
557  | 
apply (subgoal_tac "(\<forall>y\<in>set pns. y \<notin> set (map fst lvars))")  | 
|
558  | 
apply (simp add: init_vars_def locvars_xstate_def locvars_locals_def galldefs unique_def split_def map_of_map_as_map_upd)  | 
|
559  | 
apply (intro strip)  | 
|
560  | 
apply (simp (no_asm_simp) only: append_Cons [THEN sym])  | 
|
561  | 
apply (rule progression_lvar_init_aux)  | 
|
562  | 
apply (auto simp: unique_def map_of_in_set disjoint_varnames_def)  | 
|
563  | 
done  | 
|
564  | 
||
565  | 
||
566  | 
||
567  | 
||
568  | 
(**********************************************************************)  | 
|
569  | 
||
| 14045 | 570  | 
lemma state_ok_eval: "\<lbrakk>xs::\<preceq>E; wf_java_prog (prg E); wtpd_expr E e;  | 
571  | 
(prg E)\<turnstile>xs -e\<succ>v -> xs'\<rbrakk> \<Longrightarrow> xs'::\<preceq>E"  | 
|
572  | 
apply (simp only: wtpd_expr_def)  | 
|
| 13673 | 573  | 
apply (erule exE)  | 
| 14045 | 574  | 
apply (case_tac xs', case_tac xs)  | 
575  | 
apply (auto intro: eval_type_sound [THEN conjunct1])  | 
|
| 13673 | 576  | 
done  | 
577  | 
||
| 14045 | 578  | 
lemma state_ok_evals: "\<lbrakk>xs::\<preceq>E; wf_java_prog (prg E); wtpd_exprs E es;  | 
| 22271 | 579  | 
prg E \<turnstile> xs -es[\<succ>]vs-> xs'\<rbrakk> \<Longrightarrow> xs'::\<preceq>E"  | 
| 14045 | 580  | 
apply (simp only: wtpd_exprs_def)  | 
581  | 
apply (erule exE)  | 
|
582  | 
apply (case_tac xs) apply (case_tac xs')  | 
|
583  | 
apply (auto intro: evals_type_sound [THEN conjunct1])  | 
|
| 13673 | 584  | 
done  | 
585  | 
||
| 14045 | 586  | 
lemma state_ok_exec: "\<lbrakk>xs::\<preceq>E; wf_java_prog (prg E); wtpd_stmt E st;  | 
| 22271 | 587  | 
prg E \<turnstile> xs -st-> xs'\<rbrakk> \<Longrightarrow> xs'::\<preceq>E"  | 
| 14045 | 588  | 
apply (simp only: wtpd_stmt_def)  | 
589  | 
apply (case_tac xs', case_tac xs)  | 
|
| 14143 | 590  | 
apply (auto dest: exec_type_sound)  | 
| 13673 | 591  | 
done  | 
592  | 
||
593  | 
||
594  | 
lemma state_ok_init:  | 
|
| 14045 | 595  | 
"\<lbrakk> wf_java_prog G; (x, h, l)::\<preceq>(env_of_jmb G C S);  | 
| 13673 | 596  | 
is_class G dynT;  | 
597  | 
method (G, dynT) (mn, pTs) = Some (md, rT, pns, lvars, blk, res);  | 
|
598  | 
list_all2 (conf G h) pvs pTs; G,h \<turnstile> a' ::\<preceq> Class md\<rbrakk>  | 
|
599  | 
\<Longrightarrow>  | 
|
| 14045 | 600  | 
(np a' x, h, init_vars lvars(pns[\<mapsto>]pvs)(This\<mapsto>a'))::\<preceq>(env_of_jmb G md (mn, pTs))"  | 
601  | 
apply (frule wf_prog_ws_prog)  | 
|
| 13673 | 602  | 
apply (frule method_in_md [THEN conjunct2], assumption+)  | 
603  | 
apply (frule method_yields_wf_java_mdecl, assumption, assumption)  | 
|
| 14045 | 604  | 
apply (simp add: env_of_jmb_def gs_def conforms_def split_beta)  | 
| 13673 | 605  | 
apply (simp add: wf_java_mdecl_def)  | 
606  | 
apply (rule conjI)  | 
|
607  | 
apply (rule lconf_ext)  | 
|
608  | 
apply (rule lconf_ext_list)  | 
|
609  | 
apply (rule lconf_init_vars)  | 
|
610  | 
apply (auto dest: Ball_set_table)  | 
|
611  | 
apply (simp add: np_def xconf_raise_if)  | 
|
612  | 
done  | 
|
613  | 
||
614  | 
||
615  | 
lemma ty_exprs_list_all2 [rule_format (no_asm)]:  | 
|
616  | 
"(\<forall> Ts. (E \<turnstile> es [::] Ts) = list_all2 (\<lambda>e T. E \<turnstile> e :: T) es Ts)"  | 
|
617  | 
apply (rule list.induct)  | 
|
618  | 
apply simp  | 
|
619  | 
apply (rule allI)  | 
|
620  | 
apply (rule iffI)  | 
|
| 22271 | 621  | 
apply (ind_cases2 "E \<turnstile> [] [::] Ts" for Ts, assumption)  | 
| 13673 | 622  | 
apply simp apply (rule WellType.Nil)  | 
623  | 
apply (simp add: list_all2_Cons1)  | 
|
624  | 
apply (rule allI)  | 
|
625  | 
apply (rule iffI)  | 
|
626  | 
apply (rename_tac a exs Ts)  | 
|
| 22271 | 627  | 
apply (ind_cases2 "E \<turnstile> a # exs [::] Ts" for a exs Ts) apply blast  | 
| 13673 | 628  | 
apply (auto intro: WellType.Cons)  | 
629  | 
done  | 
|
630  | 
||
631  | 
||
632  | 
lemma conf_bool: "G,h \<turnstile> v::\<preceq>PrimT Boolean \<Longrightarrow> \<exists> b. v = Bool b"  | 
|
633  | 
apply (simp add: conf_def)  | 
|
634  | 
apply (erule exE)  | 
|
635  | 
apply (case_tac v)  | 
|
636  | 
apply (auto elim: widen.cases)  | 
|
637  | 
done  | 
|
638  | 
||
639  | 
||
| 14045 | 640  | 
lemma class_expr_is_class: "\<lbrakk>E \<turnstile> e :: Class C; ws_prog (prg E)\<rbrakk>  | 
| 13673 | 641  | 
\<Longrightarrow> is_class (prg E) C"  | 
642  | 
by (case_tac E, auto dest: ty_expr_is_type)  | 
|
643  | 
||
644  | 
||
645  | 
lemma max_spec_widen: "max_spec G C (mn, pTs) = {((md,rT),pTs')} \<Longrightarrow> 
 | 
|
646  | 
list_all2 (\<lambda> T T'. G \<turnstile> T \<preceq> T') pTs pTs'"  | 
|
647  | 
by (blast dest: singleton_in_set max_spec2appl_meths appl_methsD)  | 
|
648  | 
||
649  | 
||
| 14045 | 650  | 
lemma eval_conf: "\<lbrakk>G \<turnstile> s -e\<succ>v-> s'; wf_java_prog G; s::\<preceq>E;  | 
| 13673 | 651  | 
E\<turnstile>e::T; gx s' = None; prg E = G \<rbrakk>  | 
652  | 
\<Longrightarrow> G,gh s'\<turnstile>v::\<preceq>T"  | 
|
653  | 
apply (simp add: gh_def)  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
654  | 
apply (rule_tac x3="fst s" and s3="snd s"and x'3="fst s'"  | 
| 14143 | 655  | 
in eval_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]], simplified])  | 
656  | 
apply assumption+  | 
|
| 13673 | 657  | 
apply (simp (no_asm_use) add: surjective_pairing [THEN sym])  | 
| 14045 | 658  | 
apply (simp only: surjective_pairing [THEN sym])  | 
659  | 
apply (auto simp add: gs_def gx_def)  | 
|
| 13673 | 660  | 
done  | 
661  | 
||
662  | 
lemma evals_preserves_conf:  | 
|
663  | 
"\<lbrakk> G\<turnstile> s -es[\<succ>]vs-> s'; G,gh s \<turnstile> t ::\<preceq> T; E \<turnstile>es[::]Ts;  | 
|
| 14045 | 664  | 
wf_java_prog G; s::\<preceq>E;  | 
| 13673 | 665  | 
prg E = G \<rbrakk> \<Longrightarrow> G,gh s' \<turnstile> t ::\<preceq> T"  | 
666  | 
apply (subgoal_tac "gh s\<le>| gh s'")  | 
|
667  | 
apply (frule conf_hext, assumption, assumption)  | 
|
668  | 
apply (frule eval_evals_exec_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]]])  | 
|
669  | 
apply (subgoal_tac "G \<turnstile> (gx s, (gh s, gl s)) -es[\<succ>]vs-> (gx s', (gh s', gl s'))")  | 
|
670  | 
apply assumption  | 
|
671  | 
apply (simp add: gx_def gh_def gl_def surjective_pairing [THEN sym])  | 
|
672  | 
apply (case_tac E)  | 
|
| 14045 | 673  | 
apply (simp add: gx_def gs_def gh_def gl_def surjective_pairing [THEN sym])  | 
| 13673 | 674  | 
done  | 
675  | 
||
676  | 
lemma eval_of_class: "\<lbrakk> G \<turnstile> s -e\<succ>a'-> s'; E \<turnstile> e :: Class C;  | 
|
| 14045 | 677  | 
wf_java_prog G; s::\<preceq>E; gx s'=None; a' \<noteq> Null; G=prg E\<rbrakk>  | 
| 13673 | 678  | 
\<Longrightarrow> (\<exists> lc. a' = Addr lc)"  | 
679  | 
apply (case_tac s, case_tac s', simp)  | 
|
| 14045 | 680  | 
apply (frule eval_type_sound, (simp add: gs_def)+)  | 
| 13673 | 681  | 
apply (case_tac a')  | 
682  | 
apply (auto simp: conf_def)  | 
|
683  | 
done  | 
|
684  | 
||
685  | 
||
686  | 
lemma dynT_subcls:  | 
|
687  | 
"\<lbrakk> a' \<noteq> Null; G,h\<turnstile>a'::\<preceq> Class C; dynT = fst (the (h (the_Addr a')));  | 
|
| 14045 | 688  | 
is_class G dynT; ws_prog G \<rbrakk> \<Longrightarrow> G\<turnstile>dynT \<preceq>C C"  | 
| 13673 | 689  | 
apply (case_tac "C = Object")  | 
690  | 
apply (simp, rule subcls_C_Object, assumption+)  | 
|
| 14045 | 691  | 
apply simp  | 
| 13673 | 692  | 
apply (frule non_np_objD, auto)  | 
693  | 
done  | 
|
694  | 
||
695  | 
||
696  | 
lemma method_defined: "\<lbrakk>  | 
|
697  | 
m = the (method (G, dynT) (mn, pTs));  | 
|
698  | 
dynT = fst (the (h a)); is_class G dynT; wf_java_prog G;  | 
|
699  | 
a' \<noteq> Null; G,h\<turnstile>a'::\<preceq> Class C; a = the_Addr a';  | 
|
700  | 
  \<exists>pTsa md rT. max_spec G C (mn, pTsa) = {((md, rT), pTs)} \<rbrakk>
 | 
|
701  | 
\<Longrightarrow> (method (G, dynT) (mn, pTs)) = Some m"  | 
|
702  | 
apply (erule exE)+  | 
|
703  | 
apply (drule singleton_in_set, drule max_spec2appl_meths)  | 
|
704  | 
apply (simp add: appl_methds_def)  | 
|
705  | 
apply ((erule exE)+, (erule conjE)+, (erule exE)+)  | 
|
706  | 
apply (drule widen_methd)  | 
|
707  | 
apply assumption  | 
|
| 14045 | 708  | 
apply (rule dynT_subcls) apply assumption+ apply simp apply simp  | 
| 13673 | 709  | 
apply (erule exE)+ apply simp  | 
710  | 
done  | 
|
711  | 
||
712  | 
||
713  | 
||
714  | 
(**********************************************************************)  | 
|
715  | 
||
716  | 
||
717  | 
(* 1. any difference between locvars_xstate \<dots> and L ??? *)  | 
|
718  | 
(* 2. possibly skip env_of_jmb ??? *)  | 
|
719  | 
theorem compiler_correctness:  | 
|
720  | 
"wf_java_prog G \<Longrightarrow>  | 
|
| 22271 | 721  | 
(G \<turnstile> xs -ex\<succ>val-> xs' \<longrightarrow>  | 
| 13673 | 722  | 
gx xs = None \<longrightarrow> gx xs' = None \<longrightarrow>  | 
723  | 
(\<forall> os CL S.  | 
|
724  | 
(class_sig_defined G CL S) \<longrightarrow>  | 
|
725  | 
(wtpd_expr (env_of_jmb G CL S) ex) \<longrightarrow>  | 
|
| 14045 | 726  | 
(xs ::\<preceq> (env_of_jmb G CL S)) \<longrightarrow>  | 
| 13673 | 727  | 
  ( {TranslComp.comp G, CL, S} \<turnstile>
 | 
728  | 
    {gh xs, os, (locvars_xstate G CL S xs)}
 | 
|
729  | 
>- (compExpr (gmb G CL S) ex) \<rightarrow>  | 
|
730  | 
    {gh xs', val#os, locvars_xstate G CL S xs'}))) \<and> 
 | 
|
731  | 
||
| 22271 | 732  | 
(G \<turnstile> xs -exs[\<succ>]vals-> xs' \<longrightarrow>  | 
| 13673 | 733  | 
gx xs = None \<longrightarrow> gx xs' = None \<longrightarrow>  | 
734  | 
(\<forall> os CL S.  | 
|
735  | 
(class_sig_defined G CL S) \<longrightarrow>  | 
|
736  | 
(wtpd_exprs (env_of_jmb G CL S) exs) \<longrightarrow>  | 
|
| 14045 | 737  | 
(xs::\<preceq>(env_of_jmb G CL S)) \<longrightarrow>  | 
| 13673 | 738  | 
  ( {TranslComp.comp G, CL, S} \<turnstile>
 | 
739  | 
    {gh xs, os, (locvars_xstate G CL S xs)}
 | 
|
740  | 
>- (compExprs (gmb G CL S) exs) \<rightarrow>  | 
|
741  | 
    {gh xs', (rev vals)@os, (locvars_xstate G CL S xs')}))) \<and> 
 | 
|
742  | 
||
| 22271 | 743  | 
(G \<turnstile> xs -st-> xs' \<longrightarrow>  | 
| 13673 | 744  | 
gx xs = None \<longrightarrow> gx xs' = None \<longrightarrow>  | 
745  | 
(\<forall> os CL S.  | 
|
746  | 
(class_sig_defined G CL S) \<longrightarrow>  | 
|
747  | 
(wtpd_stmt (env_of_jmb G CL S) st) \<longrightarrow>  | 
|
| 14045 | 748  | 
(xs::\<preceq>(env_of_jmb G CL S)) \<longrightarrow>  | 
| 13673 | 749  | 
  ( {TranslComp.comp G, CL, S} \<turnstile>
 | 
750  | 
    {gh xs, os, (locvars_xstate G CL S xs)}
 | 
|
751  | 
>- (compStmt (gmb G CL S) st) \<rightarrow>  | 
|
752  | 
    {gh xs', os, (locvars_xstate G CL S xs')})))"
 | 
|
753  | 
apply (rule Eval.eval_evals_exec.induct)  | 
|
754  | 
||
755  | 
(* case XcptE *)  | 
|
756  | 
apply simp  | 
|
757  | 
||
758  | 
(* case NewC *)  | 
|
| 14045 | 759  | 
apply clarify  | 
760  | 
apply (frule wf_prog_ws_prog [THEN wf_subcls1]) (* establish wf ((subcls1 G)^-1) *)  | 
|
| 13673 | 761  | 
apply (simp add: c_hupd_hp_invariant)  | 
762  | 
apply (rule progression_one_step)  | 
|
763  | 
apply (rotate_tac 1, drule sym) (* reverse equation (a, None) = new_Addr (fst s) *)  | 
|
764  | 
apply (simp add: locvars_xstate_def locvars_locals_def comp_fields)  | 
|
765  | 
||
766  | 
||
767  | 
(* case Cast *)  | 
|
768  | 
apply (intro allI impI)  | 
|
769  | 
apply simp  | 
|
770  | 
apply (frule raise_if_NoneD)  | 
|
771  | 
apply (frule wtpd_expr_cast)  | 
|
772  | 
apply simp  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
773  | 
apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, simp)  | 
| 13673 | 774  | 
apply blast  | 
775  | 
apply (rule progression_one_step)  | 
|
776  | 
apply (simp add: raise_system_xcpt_def gh_def comp_cast_ok)  | 
|
777  | 
||
778  | 
||
779  | 
(* case Lit *)  | 
|
780  | 
apply simp  | 
|
781  | 
apply (intro strip)  | 
|
782  | 
apply (rule progression_one_step)  | 
|
783  | 
apply simp  | 
|
784  | 
||
785  | 
||
786  | 
(* case BinOp *)  | 
|
787  | 
apply (intro allI impI)  | 
|
788  | 
apply (frule_tac xs=s1 in eval_xcpt, assumption) (* establish (gx s1 = None) *)  | 
|
789  | 
apply (frule wtpd_expr_binop)  | 
|
| 14045 | 790  | 
(* establish (s1::\<preceq> \<dots>) *)  | 
| 13673 | 791  | 
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)  | 
792  | 
||
793  | 
||
794  | 
apply (simp (no_asm_use) only: compExpr_compExprs.simps)  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
795  | 
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: 
14143 
diff
changeset
 | 
796  | 
apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e2" in progression_transitive, simp) apply blast  | 
| 13673 | 797  | 
apply (case_tac bop)  | 
798  | 
(*subcase bop = Eq *) apply simp apply (rule progression_Eq)  | 
|
799  | 
(*subcase bop = Add *) apply simp apply (rule progression_one_step) apply simp  | 
|
800  | 
||
801  | 
||
802  | 
(* case LAcc *)  | 
|
803  | 
apply simp  | 
|
804  | 
apply (intro strip)  | 
|
805  | 
apply (rule progression_one_step)  | 
|
806  | 
apply (simp add: locvars_xstate_def locvars_locals_def)  | 
|
807  | 
apply (frule wtpd_expr_lacc)  | 
|
808  | 
apply assumption  | 
|
809  | 
apply (simp add: gl_def)  | 
|
810  | 
apply (erule select_at_index)  | 
|
811  | 
||
812  | 
||
813  | 
(* case LAss *)  | 
|
814  | 
apply (intro allI impI)  | 
|
815  | 
apply (frule wtpd_expr_lass, erule conjE, erule conjE)  | 
|
816  | 
apply (simp add: compExpr_compExprs.simps)  | 
|
817  | 
||
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
818  | 
apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)  | 
| 13673 | 819  | 
apply blast  | 
820  | 
||
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
821  | 
apply (rule_tac ?instrs0.0 = "[Dup]" in progression_transitive, simp)  | 
| 13673 | 822  | 
apply (rule progression_one_step)  | 
823  | 
apply (simp add: gh_def)  | 
|
824  | 
apply (rule conjI, simp)+ apply simp  | 
|
825  | 
apply (rule progression_one_step)  | 
|
826  | 
apply (simp add: gh_def)  | 
|
827  | 
(* the following falls out of the general scheme *)  | 
|
828  | 
apply (frule wtpd_expr_lacc) apply assumption  | 
|
829  | 
apply (rule update_at_index)  | 
|
830  | 
apply (rule distinct_method_if_class_sig_defined) apply assumption  | 
|
831  | 
apply assumption apply simp apply assumption  | 
|
832  | 
||
833  | 
||
834  | 
(* case FAcc *)  | 
|
835  | 
apply (intro allI impI)  | 
|
836  | 
(* establish x1 = None \<and> a' \<noteq> Null *)  | 
|
837  | 
apply (simp (no_asm_use) only: gx_conv, frule np_NoneD)  | 
|
838  | 
apply (frule wtpd_expr_facc)  | 
|
839  | 
||
840  | 
apply (simp (no_asm_use) only: compExpr_compExprs.simps)  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
841  | 
apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)  | 
| 13673 | 842  | 
apply blast  | 
843  | 
apply (rule progression_one_step)  | 
|
844  | 
apply (simp add: gh_def)  | 
|
845  | 
apply (case_tac "(the (fst s1 (the_Addr a')))")  | 
|
846  | 
apply (simp add: raise_system_xcpt_def)  | 
|
847  | 
||
848  | 
||
849  | 
(* case FAss *)  | 
|
850  | 
apply (intro allI impI)  | 
|
851  | 
apply (frule wtpd_expr_fass) apply (erule conjE) apply (frule wtpd_expr_facc)  | 
|
852  | 
apply (simp only: c_hupd_xcpt_invariant) (* establish x2 = None *)  | 
|
853  | 
(* establish x1 = None and a' \<noteq> Null *)  | 
|
854  | 
apply (frule_tac xs="(np a' x1, s1)" in eval_xcpt)  | 
|
855  | 
apply (simp only: gx_conv, simp only: gx_conv, frule np_NoneD, erule conjE)  | 
|
856  | 
||
857  | 
||
| 14045 | 858  | 
(* establish ((Norm s1)::\<preceq> \<dots>) *)  | 
859  | 
apply (frule_tac e=e1 in state_ok_eval) apply (simp (no_asm_simp) only: env_of_jmb_fst)  | 
|
860  | 
apply assumption apply (simp (no_asm_use) only: env_of_jmb_fst)  | 
|
| 13673 | 861  | 
|
862  | 
apply (simp only: compExpr_compExprs.simps)  | 
|
863  | 
||
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
864  | 
apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e1)" in progression_transitive, rule HOL.refl)  | 
| 13673 | 865  | 
apply fast (* blast does not seem to work - why? *)  | 
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
866  | 
apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e2)" in progression_transitive, rule HOL.refl)  | 
| 13673 | 867  | 
apply fast  | 
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
868  | 
apply (rule_tac ?instrs0.0 = "[Dup_x1]" and ?instrs1.0 = "[Putfield fn T]" in progression_transitive, simp)  | 
| 13673 | 869  | 
|
870  | 
(* Dup_x1 *)  | 
|
871  | 
apply (rule progression_one_step)  | 
|
872  | 
apply (simp add: gh_def)  | 
|
873  | 
apply (rule conjI, simp)+ apply simp  | 
|
874  | 
||
875  | 
||
876  | 
(* Putfield \<longrightarrow> still looks nasty*)  | 
|
877  | 
apply (rule progression_one_step)  | 
|
878  | 
apply simp  | 
|
879  | 
apply (case_tac "(the (fst s2 (the_Addr a')))")  | 
|
880  | 
apply (simp add: c_hupd_hp_invariant)  | 
|
881  | 
apply (case_tac s2)  | 
|
882  | 
apply (simp add: c_hupd_conv raise_system_xcpt_def)  | 
|
883  | 
apply (rule locvars_xstate_par_dep, rule HOL.refl)  | 
|
884  | 
||
885  | 
defer (* method call *)  | 
|
886  | 
||
887  | 
(* case XcptEs *)  | 
|
888  | 
apply simp  | 
|
889  | 
||
890  | 
(* case Nil *)  | 
|
891  | 
apply (simp add: compExpr_compExprs.simps)  | 
|
892  | 
apply (intro strip)  | 
|
893  | 
apply (rule progression_refl)  | 
|
894  | 
||
895  | 
(* case Cons *)  | 
|
896  | 
apply (intro allI impI)  | 
|
897  | 
apply (frule_tac xs=s1 in evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *)  | 
|
898  | 
apply (frule wtpd_exprs_cons)  | 
|
| 14045 | 899  | 
(* establish ((Norm s0)::\<preceq> \<dots>) *)  | 
| 13673 | 900  | 
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)  | 
901  | 
||
902  | 
apply simp  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
903  | 
apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)  | 
| 13673 | 904  | 
apply fast  | 
905  | 
apply fast  | 
|
906  | 
||
907  | 
||
908  | 
(* case Statement: exception *)  | 
|
909  | 
apply simp  | 
|
910  | 
||
911  | 
(* case Skip *)  | 
|
912  | 
apply (intro allI impI)  | 
|
913  | 
apply simp  | 
|
914  | 
apply (rule progression_refl)  | 
|
915  | 
||
916  | 
(* case Expr *)  | 
|
917  | 
apply (intro allI impI)  | 
|
918  | 
apply (frule wtpd_stmt_expr)  | 
|
919  | 
apply simp  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
920  | 
apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)  | 
| 13673 | 921  | 
apply fast  | 
922  | 
apply (rule progression_one_step)  | 
|
923  | 
apply simp  | 
|
924  | 
||
925  | 
(* case Comp *)  | 
|
926  | 
apply (intro allI impI)  | 
|
927  | 
apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *)  | 
|
928  | 
apply (frule wtpd_stmt_comp)  | 
|
929  | 
||
| 14045 | 930  | 
(* establish (s1::\<preceq> \<dots>) *)  | 
| 13673 | 931  | 
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)  | 
932  | 
||
933  | 
apply simp  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
934  | 
apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, rule HOL.refl)  | 
| 13673 | 935  | 
apply fast  | 
936  | 
apply fast  | 
|
937  | 
||
938  | 
||
939  | 
(* case Cond *)  | 
|
940  | 
apply (intro allI impI)  | 
|
941  | 
apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *)  | 
|
942  | 
apply (frule wtpd_stmt_cond, (erule conjE)+)  | 
|
| 14045 | 943  | 
(* establish (s1::\<preceq> \<dots>) *)  | 
| 13673 | 944  | 
apply (frule_tac e=e in state_ok_eval)  | 
945  | 
apply (simp (no_asm_simp) only: env_of_jmb_fst)  | 
|
946  | 
apply assumption  | 
|
947  | 
apply (simp (no_asm_use) only: env_of_jmb_fst)  | 
|
948  | 
(* establish G,gh s1\<turnstile>v::\<preceq>PrimT Boolean *)  | 
|
949  | 
apply (frule eval_conf, assumption+, rule env_of_jmb_fst)  | 
|
950  | 
apply (frule conf_bool) (* establish \<exists>b. v = Bool b *)  | 
|
951  | 
apply (erule exE)  | 
|
952  | 
||
953  | 
apply simp  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
954  | 
apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp))  | 
| 13673 | 955  | 
apply (rule progression_one_step, simp)  | 
956  | 
apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl)  | 
|
957  | 
||
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
958  | 
apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl)  | 
| 13673 | 959  | 
apply fast  | 
960  | 
||
961  | 
apply (case_tac b)  | 
|
962  | 
(* case b= True *)  | 
|
963  | 
apply simp  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
964  | 
apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c1)))]" in progression_transitive, simp)  | 
| 13673 | 965  | 
apply (rule progression_one_step) apply simp  | 
966  | 
apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl)  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
967  | 
apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, simp)  | 
| 13673 | 968  | 
apply fast  | 
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
969  | 
apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression)  | 
| 13673 | 970  | 
apply (simp, rule conjI, (rule HOL.refl)+)  | 
| 
20432
 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 
webertj 
parents: 
20272 
diff
changeset
 | 
971  | 
apply simp apply (rule conjI, simp) apply (simp add: nat_add_distrib)  | 
| 13673 | 972  | 
apply (rule progression_refl)  | 
973  | 
||
974  | 
(* case b= False *)  | 
|
975  | 
apply simp  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
976  | 
apply (rule_tac ?instrs1.0 = "compStmt (gmb G CL S) c2" in jump_fwd_progression)  | 
| 13673 | 977  | 
apply (simp, rule conjI, (rule HOL.refl)+)  | 
| 
20432
 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 
webertj 
parents: 
20272 
diff
changeset
 | 
978  | 
apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib)  | 
| 13673 | 979  | 
apply fast  | 
980  | 
||
981  | 
(* case exit Loop *)  | 
|
982  | 
apply (intro allI impI)  | 
|
983  | 
apply (frule wtpd_stmt_loop, (erule conjE)+)  | 
|
984  | 
||
985  | 
(* establish G,gh s1\<turnstile>v::\<preceq>PrimT Boolean *)  | 
|
986  | 
apply (frule eval_conf, assumption+, rule env_of_jmb_fst)  | 
|
987  | 
apply (frule conf_bool) (* establish \<exists>b. v = Bool b *)  | 
|
988  | 
apply (erule exE)  | 
|
989  | 
apply (case_tac b)  | 
|
990  | 
||
991  | 
(* case b= True \<longrightarrow> contradiction *)  | 
|
992  | 
apply simp  | 
|
993  | 
||
994  | 
(* case b= False *)  | 
|
995  | 
apply simp  | 
|
996  | 
||
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
997  | 
apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp))  | 
| 13673 | 998  | 
apply (rule progression_one_step)  | 
999  | 
apply simp  | 
|
1000  | 
apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl)  | 
|
1001  | 
||
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
1002  | 
apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl)  | 
| 13673 | 1003  | 
apply fast  | 
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
1004  | 
apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression)  | 
| 13673 | 1005  | 
apply (simp, rule conjI, rule HOL.refl, rule HOL.refl)  | 
| 
20432
 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 
webertj 
parents: 
20272 
diff
changeset
 | 
1006  | 
apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib)  | 
| 13673 | 1007  | 
apply (rule progression_refl)  | 
1008  | 
||
1009  | 
||
1010  | 
(* case continue Loop *)  | 
|
1011  | 
apply (intro allI impI)  | 
|
1012  | 
apply (frule_tac xs=s2 in exec_xcpt, assumption) (* establish (gx s2 = None) *)  | 
|
1013  | 
apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *)  | 
|
1014  | 
apply (frule wtpd_stmt_loop, (erule conjE)+)  | 
|
1015  | 
||
| 14045 | 1016  | 
(* establish (s1::\<preceq> \<dots>) *)  | 
| 13673 | 1017  | 
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 | 1018  | 
(* establish (s2::\<preceq> \<dots>) *)  | 
| 13673 | 1019  | 
apply (frule_tac xs=s1 and st=c in state_ok_exec)  | 
1020  | 
apply (simp (no_asm_simp) only: env_of_jmb_fst) apply assumption apply (simp (no_asm_use) only: env_of_jmb_fst)  | 
|
1021  | 
||
1022  | 
(* establish G,gh s1\<turnstile>v::\<preceq>PrimT Boolean *)  | 
|
1023  | 
apply (frule eval_conf, assumption+, rule env_of_jmb_fst)  | 
|
1024  | 
apply (frule conf_bool) (* establish \<exists>b. v = Bool b *)  | 
|
1025  | 
apply (erule exE)  | 
|
1026  | 
||
1027  | 
apply simp  | 
|
1028  | 
apply (rule jump_bwd_progression)  | 
|
1029  | 
apply simp  | 
|
1030  | 
apply (rule conjI, (rule HOL.refl)+)  | 
|
1031  | 
||
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
1032  | 
apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp))  | 
| 13673 | 1033  | 
apply (rule progression_one_step)  | 
1034  | 
apply simp  | 
|
1035  | 
apply (rule conjI, simp)+ apply simp  | 
|
1036  | 
||
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
1037  | 
apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl)  | 
| 13673 | 1038  | 
apply fast  | 
1039  | 
||
1040  | 
apply (case_tac b)  | 
|
1041  | 
(* case b= True *)  | 
|
1042  | 
apply simp  | 
|
1043  | 
||
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
1044  | 
apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c)))]" in progression_transitive, simp)  | 
| 13673 | 1045  | 
apply (rule progression_one_step) apply simp  | 
1046  | 
apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl)  | 
|
1047  | 
apply fast  | 
|
1048  | 
||
1049  | 
(* case b= False \<longrightarrow> contradiction*)  | 
|
1050  | 
apply simp  | 
|
1051  | 
||
1052  | 
apply (rule jump_bwd_one_step)  | 
|
1053  | 
apply simp  | 
|
1054  | 
apply blast  | 
|
1055  | 
||
1056  | 
(*****)  | 
|
| 14045 | 1057  | 
(* case method call *)  | 
| 13673 | 1058  | 
|
1059  | 
apply (intro allI impI)  | 
|
1060  | 
||
1061  | 
apply (frule_tac xs=s3 in eval_xcpt, simp only: gx_conv) (* establish gx s3 = None *)  | 
|
1062  | 
apply (frule exec_xcpt, assumption, simp (no_asm_use) only: gx_conv, frule np_NoneD) (* establish x = None \<and> a' \<noteq> Null *)  | 
|
1063  | 
apply (frule evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *)  | 
|
1064  | 
||
1065  | 
apply (frule wtpd_expr_call, (erule conjE)+)  | 
|
1066  | 
||
1067  | 
||
1068  | 
(* assumptions about state_ok and is_class *)  | 
|
1069  | 
||
| 14045 | 1070  | 
(* establish s1::\<preceq> (env_of_jmb G CL S) *)  | 
| 13673 | 1071  | 
apply (frule_tac xs="Norm s0" and e=e in state_ok_eval)  | 
1072  | 
apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp (no_asm_use) only: env_of_jmb_fst)  | 
|
1073  | 
||
| 14045 | 1074  | 
(* establish (x, h, l)::\<preceq>(env_of_jmb G CL S) *)  | 
| 13673 | 1075  | 
apply (frule_tac xs=s1 and xs'="(x, h, l)" in state_ok_evals)  | 
1076  | 
apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp only: env_of_jmb_fst)  | 
|
1077  | 
||
1078  | 
(* establish \<exists> lc. a' = Addr lc *)  | 
|
1079  | 
apply (frule (5) eval_of_class, rule env_of_jmb_fst [THEN sym])  | 
|
1080  | 
apply (subgoal_tac "G,h \<turnstile> a' ::\<preceq> Class C")  | 
|
1081  | 
apply (subgoal_tac "is_class G dynT")  | 
|
1082  | 
||
1083  | 
(* establish method (G, dynT) (mn, pTs) = Some(md, rT, pns, lvars, blk, res) *)  | 
|
1084  | 
apply (drule method_defined, assumption+)  | 
|
1085  | 
apply (simp only: env_of_jmb_fst)  | 
|
1086  | 
apply ((erule exE)+, erule conjE, (rule exI)+, assumption)  | 
|
1087  | 
||
1088  | 
apply (subgoal_tac "is_class G md")  | 
|
1089  | 
apply (subgoal_tac "G\<turnstile>Class dynT \<preceq> Class md")  | 
|
1090  | 
apply (subgoal_tac " method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res)")  | 
|
1091  | 
apply (subgoal_tac "list_all2 (conf G h) pvs pTs")  | 
|
1092  | 
||
| 14045 | 1093  | 
(* establish (np a' x, h, init_vars lvars(pns[\<mapsto>]pvs)(This\<mapsto>a'))::\<preceq>(env_of_jmb G md (mn, pTs)) *)  | 
| 13673 | 1094  | 
apply (subgoal_tac "G,h \<turnstile> a' ::\<preceq> Class dynT")  | 
1095  | 
apply (frule (2) conf_widen)  | 
|
1096  | 
apply (frule state_ok_init, assumption+)  | 
|
1097  | 
||
1098  | 
apply (subgoal_tac "class_sig_defined G md (mn, pTs)")  | 
|
1099  | 
apply (frule wtpd_blk, assumption, assumption)  | 
|
1100  | 
apply (frule wtpd_res, assumption, assumption)  | 
|
| 14045 | 1101  | 
apply (subgoal_tac "s3::\<preceq>(env_of_jmb G md (mn, pTs))")  | 
| 13673 | 1102  | 
|
| 14045 | 1103  | 
apply (subgoal_tac "method (TranslComp.comp G, md) (mn, pTs) =  | 
1104  | 
Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))")  | 
|
1105  | 
prefer 2 apply (simp add: wf_prog_ws_prog [THEN comp_method])  | 
|
1106  | 
apply (subgoal_tac "method (TranslComp.comp G, dynT) (mn, pTs) =  | 
|
1107  | 
Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))")  | 
|
1108  | 
prefer 2 apply (simp add: wf_prog_ws_prog [THEN comp_method])  | 
|
| 13673 | 1109  | 
apply (simp only: fst_conv snd_conv)  | 
1110  | 
||
1111  | 
(* establish length pns = length pTs *)  | 
|
1112  | 
apply (frule method_preserves_length, assumption, assumption)  | 
|
1113  | 
(* establish length pvs = length ps *)  | 
|
1114  | 
apply (frule evals_preserves_length [THEN sym])  | 
|
1115  | 
||
1116  | 
(** start evaluating subexpressions **)  | 
|
1117  | 
apply (simp (no_asm_use) only: compExpr_compExprs.simps)  | 
|
1118  | 
||
1119  | 
(* evaluate e *)  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
1120  | 
apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)  | 
| 13673 | 1121  | 
apply fast  | 
1122  | 
||
1123  | 
(* evaluate parameters *)  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
1124  | 
apply (rule_tac ?instrs0.0 = "compExprs (gmb G CL S) ps" in progression_transitive, rule HOL.refl)  | 
| 13673 | 1125  | 
apply fast  | 
1126  | 
||
1127  | 
(* invokation *)  | 
|
1128  | 
apply (rule progression_call)  | 
|
1129  | 
apply (intro allI impI conjI)  | 
|
1130  | 
(* execute Invoke statement *)  | 
|
1131  | 
apply (simp (no_asm_use) only: exec_instr.simps)  | 
|
1132  | 
apply (erule thin_rl, erule thin_rl, erule thin_rl)  | 
|
1133  | 
apply (simp add: compMethod_def raise_system_xcpt_def)  | 
|
1134  | 
apply (rule conjI, simp)+ apply (rule HOL.refl)  | 
|
1135  | 
||
1136  | 
(* get instructions of invoked method *)  | 
|
1137  | 
apply (simp (no_asm_simp) add: gis_def gmb_def compMethod_def)  | 
|
1138  | 
||
1139  | 
(* var. initialization *)  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
1140  | 
apply (rule_tac ?instrs0.0 = "(compInitLvars (pns, lvars, blk, res) lvars)" in progression_transitive, rule HOL.refl)  | 
| 13673 | 1141  | 
apply (rule_tac C=md in progression_lvar_init, assumption, assumption, assumption)  | 
1142  | 
apply (simp (no_asm_simp)) (* length pns = length pvs *)  | 
|
1143  | 
apply (simp (no_asm_simp)) (* length lvars = length (replicate (length lvars) arbitrary) *)  | 
|
1144  | 
||
1145  | 
||
1146  | 
(* body statement *)  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14143 
diff
changeset
 | 
1147  | 
apply (rule_tac ?instrs0.0 = "compStmt (pns, lvars, blk, res) blk" in progression_transitive, rule HOL.refl)  | 
| 13673 | 1148  | 
apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)")  | 
1149  | 
apply (simp (no_asm_simp))  | 
|
| 15866 | 1150  | 
apply (simp only: gh_conv)  | 
1151  | 
apply (drule mp [OF _ TrueI])+  | 
|
1152  | 
apply (erule allE, erule allE, erule allE, erule impE, assumption)+  | 
|
1153  | 
apply ((erule impE, assumption)+, assumption)  | 
|
| 15860 | 1154  | 
|
| 13673 | 1155  | 
apply (simp (no_asm_use))  | 
1156  | 
apply (simp (no_asm_simp) add: gmb_def)  | 
|
1157  | 
||
1158  | 
(* return expression *)  | 
|
1159  | 
apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)")  | 
|
1160  | 
apply (simp (no_asm_simp))  | 
|
1161  | 
apply (simp only: gh_conv)  | 
|
1162  | 
apply ((drule mp, rule TrueI)+, (drule spec)+, (drule mp, assumption)+, assumption)  | 
|
1163  | 
apply (simp (no_asm_use))  | 
|
1164  | 
apply (simp (no_asm_simp) add: gmb_def)  | 
|
1165  | 
||
1166  | 
(* execute return statement *)  | 
|
1167  | 
apply (simp (no_asm_use) add: gh_def locvars_xstate_def gl_def del: drop_append)  | 
|
1168  | 
apply (subgoal_tac "rev pvs @ a' # os = (rev (a' # pvs)) @ os")  | 
|
1169  | 
apply (simp only: drop_append)  | 
|
1170  | 
apply (simp (no_asm_simp))  | 
|
1171  | 
apply (simp (no_asm_simp))  | 
|
1172  | 
||
| 14045 | 1173  | 
(* show s3::\<preceq>\<dots> *)  | 
| 13673 | 1174  | 
apply (rule_tac xs = "(np a' x, h, init_vars lvars(pns[\<mapsto>]pvs)(This\<mapsto>a'))" and st=blk in state_ok_exec)  | 
1175  | 
apply assumption apply (simp (no_asm_simp) only: env_of_jmb_fst)  | 
|
1176  | 
apply assumption apply (simp (no_asm_use) only: env_of_jmb_fst)  | 
|
1177  | 
||
1178  | 
(* show class_sig_defined G md (mn, pTs) *)  | 
|
1179  | 
apply (simp (no_asm_simp) add: class_sig_defined_def)  | 
|
1180  | 
||
1181  | 
(* show G,h \<turnstile> a' ::\<preceq> Class dynT *)  | 
|
1182  | 
apply (frule non_npD) apply assumption  | 
|
1183  | 
apply (erule exE)+ apply simp  | 
|
1184  | 
apply (rule conf_obj_AddrI) apply simp  | 
|
1185  | 
apply (rule conjI, (rule HOL.refl)+)  | 
|
1186  | 
apply (rule widen_Class_Class [THEN iffD1], rule widen.refl)  | 
|
1187  | 
||
1188  | 
||
1189  | 
(* show list_all2 (conf G h) pvs pTs *)  | 
|
1190  | 
apply (erule exE)+ apply (erule conjE)+  | 
|
1191  | 
apply (rule_tac Ts="pTsa" in conf_list_gext_widen) apply assumption  | 
|
| 22271 | 1192  | 
apply (subgoal_tac "G \<turnstile> (gx s1, gs s1) -ps[\<succ>]pvs-> (x, h, l)")  | 
| 13673 | 1193  | 
apply (frule_tac E="env_of_jmb G CL S" in evals_type_sound)  | 
| 14143 | 1194  | 
apply assumption+  | 
| 13673 | 1195  | 
apply (simp only: env_of_jmb_fst)  | 
1196  | 
apply (simp add: conforms_def xconf_def gs_def)  | 
|
1197  | 
apply simp  | 
|
1198  | 
apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [THEN sym])  | 
|
1199  | 
apply (simp (no_asm_use) only: ty_exprs_list_all2) apply simp  | 
|
| 14143 | 1200  | 
apply simp  | 
| 13673 | 1201  | 
apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [THEN sym])  | 
1202  | 
(* list_all2 (\<lambda>T T'. G \<turnstile> T \<preceq> T') pTsa pTs *)  | 
|
1203  | 
apply (rule max_spec_widen, simp only: env_of_jmb_fst)  | 
|
1204  | 
||
1205  | 
||
1206  | 
(* show method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res) *)  | 
|
| 14045 | 1207  | 
apply (frule wf_prog_ws_prog [THEN method_in_md [THEN conjunct2]], assumption+)  | 
| 13673 | 1208  | 
|
1209  | 
(* show G\<turnstile>Class dynT \<preceq> Class md *)  | 
|
1210  | 
apply (simp (no_asm_use) only: widen_Class_Class)  | 
|
1211  | 
apply (rule method_wf_mdecl [THEN conjunct1], assumption+)  | 
|
1212  | 
||
1213  | 
(* is_class G md *)  | 
|
| 14045 | 1214  | 
apply (rule wf_prog_ws_prog [THEN method_in_md [THEN conjunct1]], assumption+)  | 
| 13673 | 1215  | 
|
1216  | 
(* show is_class G dynT *)  | 
|
1217  | 
apply (frule non_npD) apply assumption  | 
|
1218  | 
apply (erule exE)+ apply (erule conjE)+  | 
|
1219  | 
apply simp  | 
|
1220  | 
apply (rule subcls_is_class2) apply assumption  | 
|
1221  | 
apply (frule class_expr_is_class) apply (simp only: env_of_jmb_fst)  | 
|
| 14045 | 1222  | 
apply (rule wf_prog_ws_prog, assumption)  | 
| 13673 | 1223  | 
apply (simp only: env_of_jmb_fst)  | 
1224  | 
||
1225  | 
(* show G,h \<turnstile> a' ::\<preceq> Class C *)  | 
|
1226  | 
apply (simp only: wtpd_exprs_def, erule exE)  | 
|
1227  | 
apply (frule evals_preserves_conf)  | 
|
1228  | 
apply (rule eval_conf, assumption+)  | 
|
1229  | 
apply (rule env_of_jmb_fst, assumption+)  | 
|
1230  | 
apply (rule env_of_jmb_fst)  | 
|
1231  | 
apply (simp only: gh_conv)  | 
|
1232  | 
done  | 
|
1233  | 
||
1234  | 
||
1235  | 
theorem compiler_correctness_eval: "  | 
|
1236  | 
\<lbrakk> G \<turnstile> (None,hp,loc) -ex \<succ> val-> (None,hp',loc');  | 
|
1237  | 
wf_java_prog G;  | 
|
1238  | 
class_sig_defined G C S;  | 
|
1239  | 
wtpd_expr (env_of_jmb G C S) ex;  | 
|
1240  | 
(None,hp,loc) ::\<preceq> (env_of_jmb G C S) \<rbrakk> \<Longrightarrow>  | 
|
1241  | 
  {(TranslComp.comp G), C, S} \<turnstile> 
 | 
|
1242  | 
    {hp, os, (locvars_locals G C S loc)}
 | 
|
1243  | 
>- (compExpr (gmb G C S) ex) \<rightarrow>  | 
|
1244  | 
    {hp', val#os, (locvars_locals G C S loc')}"
 | 
|
1245  | 
apply (frule compiler_correctness [THEN conjunct1])  | 
|
| 14045 | 1246  | 
apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def)  | 
| 13673 | 1247  | 
done  | 
1248  | 
||
1249  | 
theorem compiler_correctness_exec: "  | 
|
| 22271 | 1250  | 
\<lbrakk> G \<turnstile> Norm (hp, loc) -st-> Norm (hp', loc');  | 
| 13673 | 1251  | 
wf_java_prog G;  | 
1252  | 
class_sig_defined G C S;  | 
|
1253  | 
wtpd_stmt (env_of_jmb G C S) st;  | 
|
1254  | 
(None,hp,loc) ::\<preceq> (env_of_jmb G C S) \<rbrakk> \<Longrightarrow>  | 
|
1255  | 
  {(TranslComp.comp G), C, S} \<turnstile> 
 | 
|
1256  | 
    {hp, os, (locvars_locals G C S loc)}
 | 
|
1257  | 
>- (compStmt (gmb G C S) st) \<rightarrow>  | 
|
1258  | 
    {hp', os, (locvars_locals G C S loc')}"
 | 
|
1259  | 
apply (frule compiler_correctness [THEN conjunct2 [THEN conjunct2]])  | 
|
| 14045 | 1260  | 
apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def)  | 
| 13673 | 1261  | 
done  | 
1262  | 
||
1263  | 
(**********************************************************************)  | 
|
1264  | 
||
1265  | 
||
1266  | 
(* reinstall pair splits *)  | 
|
1267  | 
declare split_paired_All [simp] split_paired_Ex [simp]  | 
|
1268  | 
ML_setup {*
 | 
|
| 17876 | 1269  | 
change_simpset (fn ss => ss addloop ("split_all_tac", split_all_tac));
 | 
| 13673 | 1270  | 
*}  | 
1271  | 
||
| 14045 | 1272  | 
declare wf_prog_ws_prog [simp del]  | 
1273  | 
||
| 13673 | 1274  | 
end  |