author | wenzelm |
Fri, 21 Apr 2017 21:36:49 +0200 | |
changeset 65548 | b7caa2b8bdbf |
parent 61361 | 8b5f00202e1a |
child 80914 | d97fdabd9e2b |
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 |
||
61361 | 6 |
section \<open>BV Type Safety Invariant\<close> |
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" |
60773 | 45 |
("_,_ \<turnstile>JVM _ \<surd>" [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 |
|
11252 | 62 |
lemma sup_ty_opt_OK: |
63 |
"(G \<turnstile> X <=o (OK T')) = (\<exists>T. X = OK T \<and> G \<turnstile> T \<preceq> T')" |
|
52998 | 64 |
by (cases X) auto |
11252 | 65 |
|
66 |
||
61361 | 67 |
subsection \<open>approx-val\<close> |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
68 |
|
11252 | 69 |
lemma approx_val_Err [simp,intro!]: |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
70 |
"approx_val G hp x Err" |
11252 | 71 |
by (simp add: approx_val_def) |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
72 |
|
11252 | 73 |
lemma approx_val_OK [iff]: |
74 |
"approx_val G hp x (OK T) = (G,hp \<turnstile> x ::\<preceq> T)" |
|
75 |
by (simp add: approx_val_def) |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
76 |
|
11252 | 77 |
lemma approx_val_Null [simp,intro!]: |
78 |
"approx_val G hp Null (OK (RefT x))" |
|
79 |
by (auto simp add: approx_val_def) |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
80 |
|
11252 | 81 |
lemma approx_val_sup_heap: |
82 |
"\<lbrakk> approx_val G hp v T; hp \<le>| hp' \<rbrakk> \<Longrightarrow> approx_val G hp' v T" |
|
83 |
by (cases T) (blast intro: conf_hext)+ |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
84 |
|
11252 | 85 |
lemma approx_val_heap_update: |
13006 | 86 |
"\<lbrakk> hp a = Some obj'; G,hp\<turnstile> v::\<preceq>T; obj_ty obj = obj_ty obj'\<rbrakk> |
87 |
\<Longrightarrow> G,hp(a\<mapsto>obj)\<turnstile> v::\<preceq>T" |
|
52998 | 88 |
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
|
89 |
|
11252 | 90 |
lemma approx_val_widen: |
91 |
"\<lbrakk> approx_val G hp v T; G \<turnstile> T <=o T'; wf_prog wt G \<rbrakk> |
|
92 |
\<Longrightarrow> approx_val G hp v T'" |
|
52998 | 93 |
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
|
94 |
|
61361 | 95 |
subsection \<open>approx-loc\<close> |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
96 |
|
11252 | 97 |
lemma approx_loc_Nil [simp,intro!]: |
98 |
"approx_loc G hp [] []" |
|
99 |
by (simp add: approx_loc_def) |
|
100 |
||
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
101 |
lemma approx_loc_Cons [iff]: |
11252 | 102 |
"approx_loc G hp (l#ls) (L#LT) = |
103 |
(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
|
104 |
by (simp add: approx_loc_def) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
105 |
|
11252 | 106 |
lemma approx_loc_nth: |
107 |
"\<lbrakk> approx_loc G hp loc LT; n < length LT \<rbrakk> |
|
108 |
\<Longrightarrow> approx_val G hp (loc!n) (LT!n)" |
|
109 |
by (simp add: approx_loc_def list_all2_conv_all_nth) |
|
110 |
||
111 |
lemma approx_loc_imp_approx_val_sup: |
|
112 |
"\<lbrakk>approx_loc G hp loc LT; n < length LT; LT ! n = OK T; G \<turnstile> T \<preceq> T'; wf_prog wt G\<rbrakk> |
|
113 |
\<Longrightarrow> G,hp \<turnstile> (loc!n) ::\<preceq> T'" |
|
114 |
apply (drule approx_loc_nth, assumption) |
|
115 |
apply simp |
|
116 |
apply (erule conf_widen, assumption+) |
|
117 |
done |
|
118 |
||
119 |
lemma approx_loc_conv_all_nth: |
|
120 |
"approx_loc G hp loc LT = |
|
121 |
(length loc = length LT \<and> (\<forall>n < length loc. approx_val G hp (loc!n) (LT!n)))" |
|
122 |
by (simp add: approx_loc_def list_all2_conv_all_nth) |
|
123 |
||
124 |
lemma approx_loc_sup_heap: |
|
125 |
"\<lbrakk> approx_loc G hp loc LT; hp \<le>| hp' \<rbrakk> |
|
126 |
\<Longrightarrow> approx_loc G hp' loc LT" |
|
127 |
apply (clarsimp simp add: approx_loc_conv_all_nth) |
|
128 |
apply (blast intro: approx_val_sup_heap) |
|
129 |
done |
|
130 |
||
131 |
lemma approx_loc_widen: |
|
132 |
"\<lbrakk> approx_loc G hp loc LT; G \<turnstile> LT <=l LT'; wf_prog wt G \<rbrakk> |
|
133 |
\<Longrightarrow> approx_loc G hp loc LT'" |
|
134 |
apply (unfold Listn.le_def lesub_def sup_loc_def) |
|
135 |
apply (simp (no_asm_use) only: list_all2_conv_all_nth approx_loc_conv_all_nth) |
|
136 |
apply (simp (no_asm_simp)) |
|
137 |
apply clarify |
|
138 |
apply (erule allE, erule impE) |
|
139 |
apply simp |
|
140 |
apply (erule approx_val_widen) |
|
141 |
apply simp |
|
142 |
apply assumption |
|
10056 | 143 |
done |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
144 |
|
13052 | 145 |
lemma loc_widen_Err [dest]: |
146 |
"\<And>XT. G \<turnstile> replicate n Err <=l XT \<Longrightarrow> XT = replicate n Err" |
|
147 |
by (induct n) auto |
|
148 |
||
149 |
lemma approx_loc_Err [iff]: |
|
150 |
"approx_loc G hp (replicate n v) (replicate n Err)" |
|
151 |
by (induct n) auto |
|
152 |
||
11252 | 153 |
lemma approx_loc_subst: |
154 |
"\<lbrakk> approx_loc G hp loc LT; approx_val G hp x X \<rbrakk> |
|
155 |
\<Longrightarrow> approx_loc G hp (loc[idx:=x]) (LT[idx:=X])" |
|
55524
f41ef840f09d
folded 'list_all2' with the relator generated by 'datatype_new'
blanchet
parents:
52998
diff
changeset
|
156 |
apply (unfold approx_loc_def list_all2_iff) |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
157 |
apply (auto dest: subsetD [OF set_update_subset_insert] simp add: zip_update) |
10056 | 158 |
done |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
159 |
|
11252 | 160 |
lemma approx_loc_append: |
161 |
"length l1=length L1 \<Longrightarrow> |
|
10056 | 162 |
approx_loc G hp (l1@l2) (L1@L2) = |
163 |
(approx_loc G hp l1 L1 \<and> approx_loc G hp l2 L2)" |
|
55524
f41ef840f09d
folded 'list_all2' with the relator generated by 'datatype_new'
blanchet
parents:
52998
diff
changeset
|
164 |
apply (unfold approx_loc_def list_all2_iff) |
11252 | 165 |
apply (simp cong: conj_cong) |
166 |
apply blast |
|
167 |
done |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
168 |
|
61361 | 169 |
subsection \<open>approx-stk\<close> |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
170 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
171 |
lemma approx_stk_rev_lem: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
172 |
"approx_stk G hp (rev s) (rev t) = approx_stk G hp s t" |
11252 | 173 |
apply (unfold approx_stk_def approx_loc_def) |
59199 | 174 |
apply (simp add: rev_map [symmetric]) |
11252 | 175 |
done |
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: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
178 |
"approx_stk G hp (rev s) t = approx_stk G hp s (rev t)" |
11252 | 179 |
by (auto intro: subst [OF approx_stk_rev_lem]) |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
180 |
|
11252 | 181 |
lemma approx_stk_sup_heap: |
182 |
"\<lbrakk> approx_stk G hp stk ST; hp \<le>| hp' \<rbrakk> \<Longrightarrow> approx_stk G hp' stk ST" |
|
183 |
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
|
184 |
|
11252 | 185 |
lemma approx_stk_widen: |
186 |
"\<lbrakk> approx_stk G hp stk ST; G \<turnstile> map OK ST <=l map OK ST'; wf_prog wt G \<rbrakk> |
|
187 |
\<Longrightarrow> approx_stk G hp stk ST'" |
|
188 |
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
|
189 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
190 |
lemma approx_stk_Nil [iff]: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
191 |
"approx_stk G hp [] []" |
11252 | 192 |
by (simp add: approx_stk_def) |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
193 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
194 |
lemma approx_stk_Cons [iff]: |
11252 | 195 |
"approx_stk G hp (x#stk) (S#ST) = |
196 |
(approx_val G hp x (OK S) \<and> approx_stk G hp stk ST)" |
|
197 |
by (simp add: approx_stk_def) |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
198 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
199 |
lemma approx_stk_Cons_lemma [iff]: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
200 |
"approx_stk G hp stk (S#ST') = |
10496 | 201 |
(\<exists>s stk'. stk = s#stk' \<and> approx_val G hp s (OK S) \<and> approx_stk G hp stk' ST')" |
11252 | 202 |
by (simp add: list_all2_Cons2 approx_stk_def approx_loc_def) |
203 |
||
204 |
lemma approx_stk_append: |
|
205 |
"approx_stk G hp stk (S@S') \<Longrightarrow> |
|
206 |
(\<exists>s stk'. stk = s@stk' \<and> length s = length S \<and> length stk' = length S' \<and> |
|
207 |
approx_stk G hp s S \<and> approx_stk G hp stk' S')" |
|
208 |
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
|
209 |
|
11252 | 210 |
lemma approx_stk_all_widen: |
22271 | 211 |
"\<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 | 212 |
\<Longrightarrow> approx_stk G hp stk ST'" |
213 |
apply (unfold approx_stk_def) |
|
214 |
apply (clarsimp simp add: approx_loc_conv_all_nth all_set_conv_all_nth) |
|
215 |
apply (erule allE, erule impE, assumption) |
|
216 |
apply (erule allE, erule impE, assumption) |
|
217 |
apply (erule conf_widen, assumption+) |
|
218 |
done |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
219 |
|
61361 | 220 |
subsection \<open>oconf\<close> |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
221 |
|
11252 | 222 |
lemma oconf_field_update: |
13006 | 223 |
"\<lbrakk>map_of (fields (G, oT)) FD = Some T; G,hp\<turnstile>v::\<preceq>T; G,hp\<turnstile>(oT,fs)\<surd> \<rbrakk> |
224 |
\<Longrightarrow> G,hp\<turnstile>(oT, fs(FD\<mapsto>v))\<surd>" |
|
11252 | 225 |
by (simp add: oconf_def lconf_def) |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
226 |
|
11252 | 227 |
lemma oconf_newref: |
228 |
"\<lbrakk>hp oref = None; G,hp \<turnstile> obj \<surd>; G,hp \<turnstile> obj' \<surd>\<rbrakk> \<Longrightarrow> G,hp(oref\<mapsto>obj') \<turnstile> obj \<surd>" |
|
229 |
apply (unfold oconf_def lconf_def) |
|
230 |
apply simp |
|
231 |
apply (blast intro: conf_hext hext_new) |
|
232 |
done |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
233 |
|
11252 | 234 |
lemma oconf_heap_update: |
235 |
"\<lbrakk> hp a = Some obj'; obj_ty obj' = obj_ty obj''; G,hp\<turnstile>obj\<surd> \<rbrakk> |
|
236 |
\<Longrightarrow> G,hp(a\<mapsto>obj'')\<turnstile>obj\<surd>" |
|
237 |
apply (unfold oconf_def lconf_def) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
35417
diff
changeset
|
238 |
apply (fastforce intro: approx_val_heap_update) |
11252 | 239 |
done |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
240 |
|
61361 | 241 |
subsection \<open>hconf\<close> |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
242 |
|
11252 | 243 |
lemma hconf_newref: |
244 |
"\<lbrakk> hp oref = None; G\<turnstile>h hp\<surd>; G,hp\<turnstile>obj\<surd> \<rbrakk> \<Longrightarrow> G\<turnstile>h hp(oref\<mapsto>obj)\<surd>" |
|
245 |
apply (simp add: hconf_def) |
|
246 |
apply (fast intro: oconf_newref) |
|
247 |
done |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
248 |
|
11252 | 249 |
lemma hconf_field_update: |
250 |
"\<lbrakk> map_of (fields (G, oT)) X = Some T; hp a = Some(oT,fs); |
|
251 |
G,hp\<turnstile>v::\<preceq>T; G\<turnstile>h hp\<surd> \<rbrakk> |
|
252 |
\<Longrightarrow> G\<turnstile>h hp(a \<mapsto> (oT, fs(X\<mapsto>v)))\<surd>" |
|
253 |
apply (simp add: hconf_def) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
35417
diff
changeset
|
254 |
apply (fastforce intro: oconf_heap_update oconf_field_update |
11252 | 255 |
simp add: obj_ty_def) |
256 |
done |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
257 |
|
61361 | 258 |
subsection \<open>preallocated\<close> |
12545
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
259 |
|
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
260 |
lemma preallocated_field_update: |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
261 |
"\<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
|
262 |
G\<turnstile>h hp\<surd>; preallocated hp \<rbrakk> |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
263 |
\<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
|
264 |
apply (unfold preallocated_def) |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
265 |
apply (rule allI) |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
266 |
apply (erule_tac x=x in allE) |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
267 |
apply simp |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
268 |
apply (rule ccontr) |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
269 |
apply (unfold hconf_def) |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
270 |
apply (erule allE, erule allE, erule impE, assumption) |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
271 |
apply (unfold oconf_def lconf_def) |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
272 |
apply (simp del: split_paired_All) |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
273 |
done |
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
274 |
|
7319d384d0d3
removed preallocated heaps axiom (now in type safety invariant)
kleing
parents:
12516
diff
changeset
|
275 |
|
13052 | 276 |
lemma |
277 |
assumes none: "hp oref = None" and alloc: "preallocated hp" |
|
278 |
shows preallocated_newref: "preallocated (hp(oref\<mapsto>obj))" |
|
279 |
proof (cases oref) |
|
280 |
case (XcptRef x) |
|
52998 | 281 |
with none alloc have False by (auto elim: preallocatedE [of _ x]) |
13052 | 282 |
thus ?thesis .. |
283 |
next |
|
284 |
case (Loc l) |
|
285 |
with alloc show ?thesis by (simp add: preallocated_def) |
|
286 |
qed |
|
287 |
||
61361 | 288 |
subsection \<open>correct-frames\<close> |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
289 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
290 |
lemmas [simp del] = fun_upd_apply |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
291 |
|
11252 | 292 |
lemma correct_frames_field_update [rule_format]: |
293 |
"\<forall>rT C sig. |
|
13006 | 294 |
correct_frames G hp phi rT sig frs \<longrightarrow> |
295 |
hp a = Some (C,fs) \<longrightarrow> |
|
296 |
map_of (fields (G, C)) fl = Some fd \<longrightarrow> |
|
10042 | 297 |
G,hp\<turnstile>v::\<preceq>fd |
58860 | 298 |
\<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
|
299 |
apply (induct frs) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
300 |
apply simp |
10920 | 301 |
apply clarify |
11252 | 302 |
apply (simp (no_asm_use)) |
10920 | 303 |
apply clarify |
304 |
apply (unfold correct_frame_def) |
|
305 |
apply (simp (no_asm_use)) |
|
11252 | 306 |
apply clarify |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
307 |
apply (intro exI conjI) |
11252 | 308 |
apply assumption+ |
309 |
apply (erule approx_stk_sup_heap) |
|
310 |
apply (erule hext_upd_obj) |
|
311 |
apply (erule approx_loc_sup_heap) |
|
312 |
apply (erule hext_upd_obj) |
|
313 |
apply assumption+ |
|
314 |
apply blast |
|
10056 | 315 |
done |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
316 |
|
11252 | 317 |
lemma correct_frames_newref [rule_format]: |
318 |
"\<forall>rT C sig. |
|
319 |
hp x = None \<longrightarrow> |
|
320 |
correct_frames G hp phi rT sig frs \<longrightarrow> |
|
13681 | 321 |
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
|
322 |
apply (induct frs) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
323 |
apply simp |
11252 | 324 |
apply clarify |
325 |
apply (simp (no_asm_use)) |
|
326 |
apply clarify |
|
327 |
apply (unfold correct_frame_def) |
|
328 |
apply (simp (no_asm_use)) |
|
329 |
apply clarify |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9549
diff
changeset
|
330 |
apply (intro exI conjI) |
11252 | 331 |
apply assumption+ |
332 |
apply (erule approx_stk_sup_heap) |
|
333 |
apply (erule hext_new) |
|
334 |
apply (erule approx_loc_sup_heap) |
|
335 |
apply (erule hext_new) |
|
336 |
apply assumption+ |
|
337 |
apply blast |
|
10056 | 338 |
done |
8011 | 339 |
|
340 |
end |