| author | wenzelm | 
| Mon, 16 Jul 2012 15:57:21 +0200 | |
| changeset 48271 | b28defd0b5a5 | 
| parent 44890 | 22f665a2e91c | 
| child 52998 | 3295927cf777 | 
| permissions | -rw-r--r-- | 
| 8011 | 1  | 
(* Title: HOL/MicroJava/BV/Correct.thy  | 
| 12516 | 2  | 
Author: Cornelia Pusch, Gerwin Klein  | 
| 8011 | 3  | 
Copyright 1999 Technische Universitaet Muenchen  | 
4  | 
*)  | 
|
5  | 
||
| 12911 | 6  | 
header {* \isaheader{BV Type Safety Invariant} *}
 | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
7  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
26438 
diff
changeset
 | 
8  | 
theory Correct  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
26438 
diff
changeset
 | 
9  | 
imports BVSpec "../JVM/JVMExec"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
26438 
diff
changeset
 | 
10  | 
begin  | 
| 8011 | 11  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
12  | 
definition approx_val :: "[jvm_prog,aheap,val,ty err] \<Rightarrow> bool" where  | 
| 13006 | 13  | 
"approx_val G h v any == case any of Err \<Rightarrow> True | OK T \<Rightarrow> G,h\<turnstile>v::\<preceq>T"  | 
| 8011 | 14  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
15  | 
definition approx_loc :: "[jvm_prog,aheap,val list,locvars_type] \<Rightarrow> bool" where  | 
| 10056 | 16  | 
"approx_loc G hp loc LT == list_all2 (approx_val G hp) loc LT"  | 
17  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
18  | 
definition approx_stk :: "[jvm_prog,aheap,opstack,opstack_type] \<Rightarrow> bool" where  | 
| 10496 | 19  | 
"approx_stk G hp stk ST == approx_loc G hp stk (map OK ST)"  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
20  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
21  | 
definition correct_frame :: "[jvm_prog,aheap,state_type,nat,bytecode] \<Rightarrow> frame \<Rightarrow> bool" where  | 
| 10056 | 22  | 
"correct_frame G hp == \<lambda>(ST,LT) maxl ins (stk,loc,C,sig,pc).  | 
23  | 
approx_stk G hp stk ST \<and> approx_loc G hp loc LT \<and>  | 
|
24  | 
pc < length ins \<and> length loc=length(snd sig)+maxl+1"  | 
|
| 8011 | 25  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
26  | 
primrec correct_frames :: "[jvm_prog,aheap,prog_type,ty,sig,frame list] \<Rightarrow> bool" where  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
27  | 
"correct_frames G hp phi rT0 sig0 [] = True"  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
28  | 
| "correct_frames G hp phi rT0 sig0 (f#frs) =  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
29  | 
(let (stk,loc,C,sig,pc) = f in  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
30  | 
(\<exists>ST LT rT maxs maxl ins et.  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
31  | 
phi C sig ! pc = Some (ST,LT) \<and> is_class G C \<and>  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
32  | 
method (G,C) sig = Some(C,rT,(maxs,maxl,ins,et)) \<and>  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
33  | 
(\<exists>C' mn pTs. ins!pc = (Invoke C' mn pTs) \<and>  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
34  | 
(mn,pTs) = sig0 \<and>  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
35  | 
(\<exists>apTs D ST' LT'.  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
36  | 
(phi C sig)!pc = Some ((rev apTs) @ (Class D) # ST', LT') \<and>  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
37  | 
length apTs = length pTs \<and>  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
38  | 
(\<exists>D' rT' maxs' maxl' ins' et'.  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
39  | 
method (G,D) sig0 = Some(D',rT',(maxs',maxl',ins',et')) \<and>  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
40  | 
G \<turnstile> rT0 \<preceq> rT') \<and>  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
41  | 
correct_frame G hp (ST, LT) maxl ins f \<and>  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
42  | 
correct_frames G hp phi rT sig frs))))"  | 
| 8011 | 43  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
44  | 
definition correct_state :: "[jvm_prog,prog_type,jvm_state] \<Rightarrow> bool"  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
45  | 
                  ("_,_ |-JVM _ [ok]"  [51,51] 50) where
 | 
| 10042 | 46  | 
"correct_state G phi == \<lambda>(xp,hp,frs).  | 
| 8011 | 47  | 
case xp of  | 
| 13006 | 48  | 
None \<Rightarrow> (case frs of  | 
49  | 
[] \<Rightarrow> True  | 
|
50  | 
| (f#fs) \<Rightarrow> G\<turnstile>h hp\<surd> \<and> preallocated hp \<and>  | 
|
| 12516 | 51  | 
(let (stk,loc,C,sig,pc) = f  | 
52  | 
in  | 
|
53  | 
\<exists>rT maxs maxl ins et s.  | 
|
| 
10625
 
022c6b2bcd6b
strengthened invariant: current class must be defined
 
kleing 
parents: 
10612 
diff
changeset
 | 
54  | 
is_class G C \<and>  | 
| 12516 | 55  | 
method (G,C) sig = Some(C,rT,(maxs,maxl,ins,et)) \<and>  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
56  | 
phi C sig ! pc = Some s \<and>  | 
| 12516 | 57  | 
correct_frame G hp s maxl ins f \<and>  | 
58  | 
correct_frames G hp phi rT sig fs))  | 
|
| 13006 | 59  | 
| Some x \<Rightarrow> frs = []"  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
60  | 
|
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
61  | 
|
| 
35355
 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 
wenzelm 
parents: 
33954 
diff
changeset
 | 
62  | 
notation (xsymbols)  | 
| 
 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 
wenzelm 
parents: 
33954 
diff
changeset
 | 
63  | 
 correct_state  ("_,_ \<turnstile>JVM _ \<surd>"  [51,51] 50)
 | 
| 10060 | 64  | 
|
| 11252 | 65  | 
|
66  | 
lemma sup_ty_opt_OK:  | 
|
67  | 
"(G \<turnstile> X <=o (OK T')) = (\<exists>T. X = OK T \<and> G \<turnstile> T \<preceq> T')"  | 
|
68  | 
apply (cases X)  | 
|
69  | 
apply auto  | 
|
70  | 
done  | 
|
71  | 
||
72  | 
||
| 11085 | 73  | 
section {* approx-val *}
 | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
74  | 
|
| 11252 | 75  | 
lemma approx_val_Err [simp,intro!]:  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
76  | 
"approx_val G hp x Err"  | 
| 11252 | 77  | 
by (simp add: approx_val_def)  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
78  | 
|
| 11252 | 79  | 
lemma approx_val_OK [iff]:  | 
80  | 
"approx_val G hp x (OK T) = (G,hp \<turnstile> x ::\<preceq> T)"  | 
|
81  | 
by (simp add: approx_val_def)  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
82  | 
|
| 11252 | 83  | 
lemma approx_val_Null [simp,intro!]:  | 
84  | 
"approx_val G hp Null (OK (RefT x))"  | 
|
85  | 
by (auto simp add: approx_val_def)  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
86  | 
|
| 11252 | 87  | 
lemma approx_val_sup_heap:  | 
88  | 
"\<lbrakk> approx_val G hp v T; hp \<le>| hp' \<rbrakk> \<Longrightarrow> approx_val G hp' v T"  | 
|
89  | 
by (cases T) (blast intro: conf_hext)+  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
90  | 
|
| 11252 | 91  | 
lemma approx_val_heap_update:  | 
| 13006 | 92  | 
"\<lbrakk> hp a = Some obj'; G,hp\<turnstile> v::\<preceq>T; obj_ty obj = obj_ty obj'\<rbrakk>  | 
93  | 
\<Longrightarrow> G,hp(a\<mapsto>obj)\<turnstile> v::\<preceq>T"  | 
|
| 11252 | 94  | 
by (cases v, auto simp add: obj_ty_def conf_def)  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
95  | 
|
| 11252 | 96  | 
lemma approx_val_widen:  | 
97  | 
"\<lbrakk> approx_val G hp v T; G \<turnstile> T <=o T'; wf_prog wt G \<rbrakk>  | 
|
98  | 
\<Longrightarrow> approx_val G hp v T'"  | 
|
99  | 
by (cases T', auto simp add: sup_ty_opt_OK intro: conf_widen)  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
100  | 
|
| 11085 | 101  | 
section {* approx-loc *}
 | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
102  | 
|
| 11252 | 103  | 
lemma approx_loc_Nil [simp,intro!]:  | 
104  | 
"approx_loc G hp [] []"  | 
|
105  | 
by (simp add: approx_loc_def)  | 
|
106  | 
||
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
107  | 
lemma approx_loc_Cons [iff]:  | 
| 11252 | 108  | 
"approx_loc G hp (l#ls) (L#LT) =  | 
109  | 
(approx_val G hp l L \<and> approx_loc G hp ls LT)"  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
110  | 
by (simp add: approx_loc_def)  | 
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
111  | 
|
| 11252 | 112  | 
lemma approx_loc_nth:  | 
113  | 
"\<lbrakk> approx_loc G hp loc LT; n < length LT \<rbrakk>  | 
|
114  | 
\<Longrightarrow> approx_val G hp (loc!n) (LT!n)"  | 
|
115  | 
by (simp add: approx_loc_def list_all2_conv_all_nth)  | 
|
116  | 
||
117  | 
lemma approx_loc_imp_approx_val_sup:  | 
|
118  | 
"\<lbrakk>approx_loc G hp loc LT; n < length LT; LT ! n = OK T; G \<turnstile> T \<preceq> T'; wf_prog wt G\<rbrakk>  | 
|
119  | 
\<Longrightarrow> G,hp \<turnstile> (loc!n) ::\<preceq> T'"  | 
|
120  | 
apply (drule approx_loc_nth, assumption)  | 
|
121  | 
apply simp  | 
|
122  | 
apply (erule conf_widen, assumption+)  | 
|
123  | 
done  | 
|
124  | 
||
125  | 
lemma approx_loc_conv_all_nth:  | 
|
126  | 
"approx_loc G hp loc LT =  | 
|
127  | 
(length loc = length LT \<and> (\<forall>n < length loc. approx_val G hp (loc!n) (LT!n)))"  | 
|
128  | 
by (simp add: approx_loc_def list_all2_conv_all_nth)  | 
|
129  | 
||
130  | 
lemma approx_loc_sup_heap:  | 
|
131  | 
"\<lbrakk> approx_loc G hp loc LT; hp \<le>| hp' \<rbrakk>  | 
|
132  | 
\<Longrightarrow> approx_loc G hp' loc LT"  | 
|
133  | 
apply (clarsimp simp add: approx_loc_conv_all_nth)  | 
|
134  | 
apply (blast intro: approx_val_sup_heap)  | 
|
135  | 
done  | 
|
136  | 
||
137  | 
lemma approx_loc_widen:  | 
|
138  | 
"\<lbrakk> approx_loc G hp loc LT; G \<turnstile> LT <=l LT'; wf_prog wt G \<rbrakk>  | 
|
139  | 
\<Longrightarrow> approx_loc G hp loc LT'"  | 
|
140  | 
apply (unfold Listn.le_def lesub_def sup_loc_def)  | 
|
141  | 
apply (simp (no_asm_use) only: list_all2_conv_all_nth approx_loc_conv_all_nth)  | 
|
142  | 
apply (simp (no_asm_simp))  | 
|
143  | 
apply clarify  | 
|
144  | 
apply (erule allE, erule impE)  | 
|
145  | 
apply simp  | 
|
146  | 
apply (erule approx_val_widen)  | 
|
147  | 
apply simp  | 
|
148  | 
apply assumption  | 
|
| 10056 | 149  | 
done  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
150  | 
|
| 13052 | 151  | 
lemma loc_widen_Err [dest]:  | 
152  | 
"\<And>XT. G \<turnstile> replicate n Err <=l XT \<Longrightarrow> XT = replicate n Err"  | 
|
153  | 
by (induct n) auto  | 
|
154  | 
||
155  | 
lemma approx_loc_Err [iff]:  | 
|
156  | 
"approx_loc G hp (replicate n v) (replicate n Err)"  | 
|
157  | 
by (induct n) auto  | 
|
158  | 
||
| 11252 | 159  | 
lemma approx_loc_subst:  | 
160  | 
"\<lbrakk> approx_loc G hp loc LT; approx_val G hp x X \<rbrakk>  | 
|
161  | 
\<Longrightarrow> approx_loc G hp (loc[idx:=x]) (LT[idx:=X])"  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
162  | 
apply (unfold approx_loc_def list_all2_def)  | 
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
163  | 
apply (auto dest: subsetD [OF set_update_subset_insert] simp add: zip_update)  | 
| 10056 | 164  | 
done  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
165  | 
|
| 11252 | 166  | 
lemma approx_loc_append:  | 
167  | 
"length l1=length L1 \<Longrightarrow>  | 
|
| 10056 | 168  | 
approx_loc G hp (l1@l2) (L1@L2) =  | 
169  | 
(approx_loc G hp l1 L1 \<and> approx_loc G hp l2 L2)"  | 
|
| 11252 | 170  | 
apply (unfold approx_loc_def list_all2_def)  | 
171  | 
apply (simp cong: conj_cong)  | 
|
172  | 
apply blast  | 
|
173  | 
done  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
174  | 
|
| 11085 | 175  | 
section {* approx-stk *}
 | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
176  | 
|
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
177  | 
lemma approx_stk_rev_lem:  | 
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
178  | 
"approx_stk G hp (rev s) (rev t) = approx_stk G hp s t"  | 
| 11252 | 179  | 
apply (unfold approx_stk_def approx_loc_def)  | 
180  | 
apply (simp add: rev_map [THEN sym])  | 
|
181  | 
done  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
182  | 
|
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
183  | 
lemma approx_stk_rev:  | 
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
184  | 
"approx_stk G hp (rev s) t = approx_stk G hp s (rev t)"  | 
| 11252 | 185  | 
by (auto intro: subst [OF approx_stk_rev_lem])  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
186  | 
|
| 11252 | 187  | 
lemma approx_stk_sup_heap:  | 
188  | 
"\<lbrakk> approx_stk G hp stk ST; hp \<le>| hp' \<rbrakk> \<Longrightarrow> approx_stk G hp' stk ST"  | 
|
189  | 
by (auto intro: approx_loc_sup_heap simp add: approx_stk_def)  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
190  | 
|
| 11252 | 191  | 
lemma approx_stk_widen:  | 
192  | 
"\<lbrakk> approx_stk G hp stk ST; G \<turnstile> map OK ST <=l map OK ST'; wf_prog wt G \<rbrakk>  | 
|
193  | 
\<Longrightarrow> approx_stk G hp stk ST'"  | 
|
194  | 
by (auto elim: approx_loc_widen simp add: approx_stk_def)  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
195  | 
|
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
196  | 
lemma approx_stk_Nil [iff]:  | 
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
197  | 
"approx_stk G hp [] []"  | 
| 11252 | 198  | 
by (simp add: approx_stk_def)  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
199  | 
|
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
200  | 
lemma approx_stk_Cons [iff]:  | 
| 11252 | 201  | 
"approx_stk G hp (x#stk) (S#ST) =  | 
202  | 
(approx_val G hp x (OK S) \<and> approx_stk G hp stk ST)"  | 
|
203  | 
by (simp add: approx_stk_def)  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
204  | 
|
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
205  | 
lemma approx_stk_Cons_lemma [iff]:  | 
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
206  | 
"approx_stk G hp stk (S#ST') =  | 
| 10496 | 207  | 
(\<exists>s stk'. stk = s#stk' \<and> approx_val G hp s (OK S) \<and> approx_stk G hp stk' ST')"  | 
| 11252 | 208  | 
by (simp add: list_all2_Cons2 approx_stk_def approx_loc_def)  | 
209  | 
||
210  | 
lemma approx_stk_append:  | 
|
211  | 
"approx_stk G hp stk (S@S') \<Longrightarrow>  | 
|
212  | 
(\<exists>s stk'. stk = s@stk' \<and> length s = length S \<and> length stk' = length S' \<and>  | 
|
213  | 
approx_stk G hp s S \<and> approx_stk G hp stk' S')"  | 
|
214  | 
by (simp add: list_all2_append2 approx_stk_def approx_loc_def)  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
215  | 
|
| 11252 | 216  | 
lemma approx_stk_all_widen:  | 
| 22271 | 217  | 
"\<lbrakk> approx_stk G hp stk ST; \<forall>(x, y) \<in> set (zip ST ST'). G \<turnstile> x \<preceq> y; length ST = length ST'; wf_prog wt G \<rbrakk>  | 
| 11252 | 218  | 
\<Longrightarrow> approx_stk G hp stk ST'"  | 
219  | 
apply (unfold approx_stk_def)  | 
|
220  | 
apply (clarsimp simp add: approx_loc_conv_all_nth all_set_conv_all_nth)  | 
|
221  | 
apply (erule allE, erule impE, assumption)  | 
|
222  | 
apply (erule allE, erule impE, assumption)  | 
|
223  | 
apply (erule conf_widen, assumption+)  | 
|
224  | 
done  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
225  | 
|
| 11085 | 226  | 
section {* oconf *}
 | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
227  | 
|
| 11252 | 228  | 
lemma oconf_field_update:  | 
| 13006 | 229  | 
"\<lbrakk>map_of (fields (G, oT)) FD = Some T; G,hp\<turnstile>v::\<preceq>T; G,hp\<turnstile>(oT,fs)\<surd> \<rbrakk>  | 
230  | 
\<Longrightarrow> G,hp\<turnstile>(oT, fs(FD\<mapsto>v))\<surd>"  | 
|
| 11252 | 231  | 
by (simp add: oconf_def lconf_def)  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
232  | 
|
| 11252 | 233  | 
lemma oconf_newref:  | 
234  | 
"\<lbrakk>hp oref = None; G,hp \<turnstile> obj \<surd>; G,hp \<turnstile> obj' \<surd>\<rbrakk> \<Longrightarrow> G,hp(oref\<mapsto>obj') \<turnstile> obj \<surd>"  | 
|
235  | 
apply (unfold oconf_def lconf_def)  | 
|
236  | 
apply simp  | 
|
237  | 
apply (blast intro: conf_hext hext_new)  | 
|
238  | 
done  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
239  | 
|
| 11252 | 240  | 
lemma oconf_heap_update:  | 
241  | 
"\<lbrakk> hp a = Some obj'; obj_ty obj' = obj_ty obj''; G,hp\<turnstile>obj\<surd> \<rbrakk>  | 
|
242  | 
\<Longrightarrow> G,hp(a\<mapsto>obj'')\<turnstile>obj\<surd>"  | 
|
243  | 
apply (unfold oconf_def lconf_def)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
35417 
diff
changeset
 | 
244  | 
apply (fastforce intro: approx_val_heap_update)  | 
| 11252 | 245  | 
done  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
246  | 
|
| 11085 | 247  | 
section {* hconf *}
 | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
248  | 
|
| 11252 | 249  | 
lemma hconf_newref:  | 
250  | 
"\<lbrakk> hp oref = None; G\<turnstile>h hp\<surd>; G,hp\<turnstile>obj\<surd> \<rbrakk> \<Longrightarrow> G\<turnstile>h hp(oref\<mapsto>obj)\<surd>"  | 
|
251  | 
apply (simp add: hconf_def)  | 
|
252  | 
apply (fast intro: oconf_newref)  | 
|
253  | 
done  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
254  | 
|
| 11252 | 255  | 
lemma hconf_field_update:  | 
256  | 
"\<lbrakk> map_of (fields (G, oT)) X = Some T; hp a = Some(oT,fs);  | 
|
257  | 
G,hp\<turnstile>v::\<preceq>T; G\<turnstile>h hp\<surd> \<rbrakk>  | 
|
258  | 
\<Longrightarrow> G\<turnstile>h hp(a \<mapsto> (oT, fs(X\<mapsto>v)))\<surd>"  | 
|
259  | 
apply (simp add: hconf_def)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
35417 
diff
changeset
 | 
260  | 
apply (fastforce intro: oconf_heap_update oconf_field_update  | 
| 11252 | 261  | 
simp add: obj_ty_def)  | 
262  | 
done  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
263  | 
|
| 
12545
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
264  | 
section {* preallocated *}
 | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
265  | 
|
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
266  | 
lemma preallocated_field_update:  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
267  | 
"\<lbrakk> map_of (fields (G, oT)) X = Some T; hp a = Some(oT,fs);  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
268  | 
G\<turnstile>h hp\<surd>; preallocated hp \<rbrakk>  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
269  | 
\<Longrightarrow> preallocated (hp(a \<mapsto> (oT, fs(X\<mapsto>v))))"  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
270  | 
apply (unfold preallocated_def)  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
271  | 
apply (rule allI)  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
272  | 
apply (erule_tac x=x in allE)  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
273  | 
apply simp  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
274  | 
apply (rule ccontr)  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
275  | 
apply (unfold hconf_def)  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
276  | 
apply (erule allE, erule allE, erule impE, assumption)  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
277  | 
apply (unfold oconf_def lconf_def)  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
278  | 
apply (simp del: split_paired_All)  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
279  | 
done  | 
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
280  | 
|
| 
 
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
 
kleing 
parents: 
12516 
diff
changeset
 | 
281  | 
|
| 13052 | 282  | 
lemma  | 
283  | 
assumes none: "hp oref = None" and alloc: "preallocated hp"  | 
|
284  | 
shows preallocated_newref: "preallocated (hp(oref\<mapsto>obj))"  | 
|
285  | 
proof (cases oref)  | 
|
286  | 
case (XcptRef x)  | 
|
287  | 
with none alloc have "False" by (auto elim: preallocatedE [of _ x])  | 
|
288  | 
thus ?thesis ..  | 
|
289  | 
next  | 
|
290  | 
case (Loc l)  | 
|
291  | 
with alloc show ?thesis by (simp add: preallocated_def)  | 
|
292  | 
qed  | 
|
293  | 
||
| 11085 | 294  | 
section {* correct-frames *}
 | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
295  | 
|
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
296  | 
lemmas [simp del] = fun_upd_apply  | 
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
297  | 
|
| 11252 | 298  | 
lemma correct_frames_field_update [rule_format]:  | 
299  | 
"\<forall>rT C sig.  | 
|
| 13006 | 300  | 
correct_frames G hp phi rT sig frs \<longrightarrow>  | 
301  | 
hp a = Some (C,fs) \<longrightarrow>  | 
|
302  | 
map_of (fields (G, C)) fl = Some fd \<longrightarrow>  | 
|
| 10042 | 303  | 
G,hp\<turnstile>v::\<preceq>fd  | 
| 13006 | 304  | 
\<longrightarrow> correct_frames G (hp(a \<mapsto> (C, fs(fl\<mapsto>v)))) phi rT sig frs";  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
305  | 
apply (induct frs)  | 
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
306  | 
apply simp  | 
| 10920 | 307  | 
apply clarify  | 
| 11252 | 308  | 
apply (simp (no_asm_use))  | 
| 10920 | 309  | 
apply clarify  | 
310  | 
apply (unfold correct_frame_def)  | 
|
311  | 
apply (simp (no_asm_use))  | 
|
| 11252 | 312  | 
apply clarify  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
313  | 
apply (intro exI conjI)  | 
| 11252 | 314  | 
apply assumption+  | 
315  | 
apply (erule approx_stk_sup_heap)  | 
|
316  | 
apply (erule hext_upd_obj)  | 
|
317  | 
apply (erule approx_loc_sup_heap)  | 
|
318  | 
apply (erule hext_upd_obj)  | 
|
319  | 
apply assumption+  | 
|
320  | 
apply blast  | 
|
| 10056 | 321  | 
done  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
322  | 
|
| 11252 | 323  | 
lemma correct_frames_newref [rule_format]:  | 
324  | 
"\<forall>rT C sig.  | 
|
325  | 
hp x = None \<longrightarrow>  | 
|
326  | 
correct_frames G hp phi rT sig frs \<longrightarrow>  | 
|
| 13681 | 327  | 
correct_frames G (hp(x \<mapsto> obj)) phi rT sig frs"  | 
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
328  | 
apply (induct frs)  | 
| 
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
329  | 
apply simp  | 
| 11252 | 330  | 
apply clarify  | 
331  | 
apply (simp (no_asm_use))  | 
|
332  | 
apply clarify  | 
|
333  | 
apply (unfold correct_frame_def)  | 
|
334  | 
apply (simp (no_asm_use))  | 
|
335  | 
apply clarify  | 
|
| 
9757
 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 
kleing 
parents: 
9549 
diff
changeset
 | 
336  | 
apply (intro exI conjI)  | 
| 11252 | 337  | 
apply assumption+  | 
338  | 
apply (erule approx_stk_sup_heap)  | 
|
339  | 
apply (erule hext_new)  | 
|
340  | 
apply (erule approx_loc_sup_heap)  | 
|
341  | 
apply (erule hext_new)  | 
|
342  | 
apply assumption+  | 
|
343  | 
apply blast  | 
|
| 10056 | 344  | 
done  | 
| 8011 | 345  | 
|
346  | 
end  |