8011
|
1 |
(* Title: HOL/MicroJava/J/Conform.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: David von Oheimb
|
|
4 |
Copyright 1999 Technische Universitaet Muenchen
|
|
5 |
*)
|
|
6 |
|
|
7 |
section "hext";
|
|
8 |
|
10042
|
9 |
val hextI = prove_goalw thy [hext_def] "!!h. \
|
|
10 |
\ \\<forall>a C fs . h a = Some (C,fs) --> \
|
|
11 |
\ (\\<exists>fs'. h' a = Some (C,fs')) ==> h\\<le>|h'" (K [Auto_tac ]);
|
8011
|
12 |
|
|
13 |
val hext_objD = prove_goalw thy [hext_def]
|
10042
|
14 |
"!!h. [|h\\<le>|h'; h a = Some (C,fs) |] ==> \\<exists>fs'. h' a = Some (C,fs')"
|
8011
|
15 |
(K [Force_tac 1]);
|
|
16 |
|
|
17 |
val hext_refl = prove_goal thy "h\\<le>|h" (K [
|
|
18 |
rtac hextI 1,
|
|
19 |
Fast_tac 1]);
|
|
20 |
|
10042
|
21 |
val hext_new = prove_goal thy "!!h. h a = None ==> h\\<le>|h(a\\<mapsto>x)" (K [
|
8011
|
22 |
rtac hextI 1,
|
|
23 |
safe_tac HOL_cs,
|
|
24 |
ALLGOALS (case_tac "aa = a"),
|
|
25 |
Auto_tac]);
|
|
26 |
|
10042
|
27 |
val hext_trans = prove_goal thy "!!h. [|h\\<le>|h'; h'\\<le>|h''|] ==> h\\<le>|h''" (K [
|
8011
|
28 |
rtac hextI 1,
|
|
29 |
safe_tac HOL_cs,
|
|
30 |
fast_tac (HOL_cs addDs [hext_objD]) 1]);
|
|
31 |
|
|
32 |
Addsimps [hext_refl, hext_new];
|
|
33 |
|
|
34 |
val hext_upd_obj = prove_goal thy
|
10042
|
35 |
"!!h. h a = Some (C,fs) ==> h\\<le>|h(a\\<mapsto>(C,fs'))" (K [
|
8011
|
36 |
rtac hextI 1,
|
|
37 |
safe_tac HOL_cs,
|
|
38 |
ALLGOALS (case_tac "aa = a"),
|
|
39 |
Auto_tac]);
|
|
40 |
|
|
41 |
|
|
42 |
section "conf";
|
|
43 |
|
|
44 |
val conf_Null = prove_goalw thy [conf_def]
|
10042
|
45 |
"G,h\\<turnstile>Null::\\<preceq>T = G\\<turnstile>RefT NullT\\<preceq>T" (K [Simp_tac 1]);
|
8011
|
46 |
Addsimps [conf_Null];
|
|
47 |
|
|
48 |
val conf_litval = prove_goalw thy [conf_def]
|
10042
|
49 |
"typeof (\\<lambda>v. None) v = Some T --> G,h\\<turnstile>v::\\<preceq>T" (K [
|
8011
|
50 |
rtac val_.induct 1,
|
|
51 |
Auto_tac]) RS mp;
|
|
52 |
|
10042
|
53 |
Goalw [conf_def] "G,s\\<turnstile>Unit::\\<preceq>PrimT Void";
|
9240
|
54 |
by( Simp_tac 1);
|
|
55 |
qed "conf_VoidI";
|
|
56 |
|
10042
|
57 |
Goalw [conf_def] "G,s\\<turnstile>Bool b::\\<preceq>PrimT Boolean";
|
9240
|
58 |
by( Simp_tac 1);
|
|
59 |
qed "conf_BooleanI";
|
|
60 |
|
10042
|
61 |
Goalw [conf_def] "G,s\\<turnstile>Intg i::\\<preceq>PrimT Integer";
|
9240
|
62 |
by( Simp_tac 1);
|
|
63 |
qed "conf_IntegerI";
|
|
64 |
|
|
65 |
Addsimps [conf_VoidI, conf_BooleanI, conf_IntegerI];
|
|
66 |
|
8011
|
67 |
val conf_AddrI = prove_goalw thy [conf_def]
|
10042
|
68 |
"!!G. [|h a = Some obj; G\\<turnstile>obj_ty obj\\<preceq>T|] ==> G,h\\<turnstile>Addr a::\\<preceq>T"
|
8011
|
69 |
(K [Asm_full_simp_tac 1]);
|
|
70 |
|
|
71 |
val conf_obj_AddrI = prove_goalw thy [conf_def]
|
10042
|
72 |
"!!G. [|h a = Some (C,fs); G\\<turnstile>C\\<preceq>C D|] ==> G,h\\<turnstile>Addr a::\\<preceq> Class D"
|
8011
|
73 |
(K [Asm_full_simp_tac 1]);
|
|
74 |
|
10042
|
75 |
Goalw [conf_def] "is_type G T --> G,h\\<turnstile>default_val T::\\<preceq>T";
|
8011
|
76 |
by (res_inst_tac [("y","T")] ty.exhaust 1);
|
|
77 |
by (etac ssubst 1);
|
|
78 |
by (res_inst_tac [("y","prim_ty")] prim_ty.exhaust 1);
|
|
79 |
by (auto_tac (claset(), simpset() addsimps [widen.null]));
|
|
80 |
qed_spec_mp "defval_conf";
|
|
81 |
|
|
82 |
val conf_upd_obj = prove_goalw thy [conf_def]
|
10042
|
83 |
"h a = Some (C,fs) --> (G,h(a\\<mapsto>(C,fs'))\\<turnstile>x::\\<preceq>T) = (G,h\\<turnstile>x::\\<preceq>T)" (fn _ => [
|
8011
|
84 |
rtac impI 1,
|
|
85 |
rtac val_.induct 1,
|
|
86 |
ALLGOALS Simp_tac,
|
|
87 |
case_tac "loc = a" 1,
|
|
88 |
ALLGOALS Asm_simp_tac]) RS mp;
|
|
89 |
|
|
90 |
val conf_widen = prove_goalw thy [conf_def]
|
10042
|
91 |
"!!G. wf_prog wf_mb G ==> G,h\\<turnstile>x::\\<preceq>T --> G\\<turnstile>T\\<preceq>T' --> G,h\\<turnstile>x::\\<preceq>T'" (K [
|
8011
|
92 |
rtac val_.induct 1,
|
|
93 |
ALLGOALS Simp_tac,
|
|
94 |
ALLGOALS (fast_tac (HOL_cs addIs [widen_trans]))]) RS mp RS mp;
|
9758
|
95 |
bind_thm ("conf_widen", conf_widen);
|
8011
|
96 |
|
|
97 |
val conf_hext' = prove_goalw thy [conf_def]
|
10042
|
98 |
"!!h. h\\<le>|h' ==> (\\<forall>v T. G,h\\<turnstile>v::\\<preceq>T --> G,h'\\<turnstile>v::\\<preceq>T)" (K [
|
8011
|
99 |
REPEAT (rtac allI 1),
|
|
100 |
rtac val_.induct 1,
|
|
101 |
ALLGOALS Simp_tac,
|
|
102 |
safe_tac (HOL_cs addSDs [option_map_SomeD]),
|
|
103 |
rewtac option_map_def,
|
|
104 |
dtac hext_objD 1,
|
|
105 |
Auto_tac]);
|
|
106 |
val conf_hext = conf_hext' RS spec RS spec RS mp;
|
9758
|
107 |
bind_thm ("conf_hext", conf_hext);
|
8011
|
108 |
|
|
109 |
val new_locD = prove_goalw thy [conf_def]
|
10042
|
110 |
"[|h a = None; G,h\\<turnstile>Addr t::\\<preceq>T|] ==> t\\<noteq>a" (fn prems => [
|
8011
|
111 |
cut_facts_tac prems 1,
|
|
112 |
Full_simp_tac 1,
|
|
113 |
safe_tac HOL_cs,
|
|
114 |
Asm_full_simp_tac 1]);
|
|
115 |
|
8185
|
116 |
Goalw [conf_def]
|
10042
|
117 |
"G,h\\<turnstile>a'::\\<preceq>RefT T --> a' = Null | \
|
8185
|
118 |
\ (\\<exists>a obj T'. a' = Addr a \\<and> h a = Some obj \\<and> obj_ty obj = T' \\<and> G\\<turnstile>T'\\<preceq>RefT T)";
|
|
119 |
by(induct_tac "a'" 1);
|
|
120 |
by(Auto_tac);
|
|
121 |
qed_spec_mp "conf_RefTD";
|
8011
|
122 |
|
10042
|
123 |
val conf_NullTD = prove_goal thy "!!G. G,h\\<turnstile>a'::\\<preceq>RefT NullT ==> a' = Null" (K [
|
8011
|
124 |
dtac conf_RefTD 1,
|
|
125 |
Step_tac 1,
|
8185
|
126 |
Auto_tac]);
|
8011
|
127 |
|
10042
|
128 |
val non_npD = prove_goal thy "!!G. [|a' \\<noteq> Null; G,h\\<turnstile>a'::\\<preceq>RefT t|] ==> \
|
8011
|
129 |
\ \\<exists>a C fs. a' = Addr a \\<and> h a = Some (C,fs) \\<and> G\\<turnstile>Class C\\<preceq>RefT t" (K [
|
|
130 |
dtac conf_RefTD 1,
|
|
131 |
Step_tac 1,
|
|
132 |
Auto_tac]);
|
|
133 |
|
10042
|
134 |
val non_np_objD = prove_goal thy "!!G. [|a' \\<noteq> Null; G,h\\<turnstile>a'::\\<preceq> Class C; C \\<noteq> Object|] ==> \
|
8185
|
135 |
\ (\\<exists>a C' fs. a' = Addr a \\<and> h a = Some (C',fs) \\<and> G\\<turnstile>C'\\<preceq>C C)"
|
|
136 |
(K[fast_tac (claset() addDs [non_npD]) 1]);
|
8011
|
137 |
|
10042
|
138 |
Goal "a' \\<noteq> Null --> wf_prog wf_mb G --> G,h\\<turnstile>a'::\\<preceq>RefT t -->\
|
|
139 |
\ (\\<forall>C. t = ClassT C --> C \\<noteq> Object) --> \
|
8011
|
140 |
\ (\\<exists>a C fs. a' = Addr a \\<and> h a = Some (C,fs) \\<and> G\\<turnstile>Class C\\<preceq>RefT t)";
|
8185
|
141 |
by(rtac impI 1);
|
|
142 |
by(rtac impI 1);
|
|
143 |
by(res_inst_tac [("y","t")] ref_ty.exhaust 1);
|
|
144 |
by(Safe_tac);
|
|
145 |
by(dtac conf_NullTD 1);
|
|
146 |
by(contr_tac 1);
|
|
147 |
by(dtac non_np_objD 1);
|
|
148 |
by(atac 1);
|
|
149 |
by(Fast_tac 1);
|
|
150 |
by(Fast_tac 1);
|
8011
|
151 |
qed_spec_mp "non_np_objD'";
|
|
152 |
|
10042
|
153 |
Goal "wf_prog wf_mb G ==> \\<forall>Ts Ts'. list_all2 (conf G h) vs Ts --> list_all2 (\\<lambda>T T'. G\\<turnstile>T\\<preceq>T') Ts Ts' --> list_all2 (conf G h) vs Ts'";
|
8116
|
154 |
by(induct_tac "vs" 1);
|
|
155 |
by(ALLGOALS Clarsimp_tac);
|
|
156 |
by(forward_tac [list_all2_lengthD RS sym] 1);
|
|
157 |
by(full_simp_tac (simpset()addsimps[length_Suc_conv]) 1);
|
|
158 |
by(Safe_tac);
|
|
159 |
by(forward_tac [list_all2_lengthD RS sym] 1);
|
|
160 |
by(full_simp_tac (simpset()addsimps[length_Suc_conv]) 1);
|
|
161 |
by(Clarify_tac 1);
|
|
162 |
by(fast_tac (claset() addEs [conf_widen]) 1);
|
8011
|
163 |
qed_spec_mp "conf_list_gext_widen";
|
|
164 |
|
|
165 |
|
|
166 |
section "lconf";
|
|
167 |
|
|
168 |
val lconfD = prove_goalw thy [lconf_def]
|
10042
|
169 |
"!!X. [| G,h\\<turnstile>vs[::\\<preceq>]Ts; Ts n = Some T |] ==> G,h\\<turnstile>(the (vs n))::\\<preceq>T"
|
8011
|
170 |
(K [Force_tac 1]);
|
|
171 |
|
|
172 |
val lconf_hext = prove_goalw thy [lconf_def]
|
10042
|
173 |
"!!X. [| G,h\\<turnstile>l[::\\<preceq>]L; h\\<le>|h' |] ==> G,h'\\<turnstile>l[::\\<preceq>]L" (K [
|
8011
|
174 |
fast_tac (claset() addEs [conf_hext]) 1]);
|
|
175 |
AddEs [lconf_hext];
|
|
176 |
|
10042
|
177 |
Goalw [lconf_def] "!!X. [| G,h\\<turnstile>l[::\\<preceq>]lT; \
|
|
178 |
\ G,h\\<turnstile>v::\\<preceq>T; lT va = Some T |] ==> G,h\\<turnstile>l(va\\<mapsto>v)[::\\<preceq>]lT";
|
8011
|
179 |
by( Clarify_tac 1);
|
|
180 |
by( case_tac "n = va" 1);
|
|
181 |
by Auto_tac;
|
|
182 |
qed "lconf_upd";
|
|
183 |
|
10042
|
184 |
Goal "\\<forall>x. P x --> R (dv x) x ==> (\\<forall>x. map_of fs f = Some x --> P x) --> \
|
|
185 |
\ (\\<forall>T. map_of fs f = Some T --> \
|
8011
|
186 |
\ (\\<exists>v. map_of (map (\\<lambda>(f,ft). (f, dv ft)) fs) f = Some v \\<and> R v T))";
|
|
187 |
by( induct_tac "fs" 1);
|
|
188 |
by Auto_tac;
|
|
189 |
qed_spec_mp "lconf_init_vars_lemma";
|
|
190 |
|
|
191 |
Goalw [lconf_def, init_vars_def]
|
10042
|
192 |
"\\<forall>n. \\<forall>T. map_of fs n = Some T --> is_type G T ==> G,h\\<turnstile>init_vars fs[::\\<preceq>]map_of fs";
|
8011
|
193 |
by Auto_tac;
|
|
194 |
by( rtac lconf_init_vars_lemma 1);
|
|
195 |
by( atac 3);
|
|
196 |
by( strip_tac 1);
|
|
197 |
by( etac defval_conf 1);
|
|
198 |
by Auto_tac;
|
|
199 |
qed "lconf_init_vars";
|
|
200 |
AddSIs [lconf_init_vars];
|
|
201 |
|
|
202 |
val lconf_ext = prove_goalw thy [lconf_def]
|
10042
|
203 |
"!!X. [|G,s\\<turnstile>l[::\\<preceq>]L; G,s\\<turnstile>v::\\<preceq>T|] ==> G,s\\<turnstile>l(vn\\<mapsto>v)[::\\<preceq>]L(vn\\<mapsto>T)"
|
8011
|
204 |
(K [Auto_tac]);
|
|
205 |
|
10042
|
206 |
Goalw [lconf_def] "G,h\\<turnstile>l[::\\<preceq>]L ==> \\<forall>vs Ts. nodups vns --> length Ts = length vns --> list_all2 (\\<lambda>v T. G,h\\<turnstile>v::\\<preceq>T) vs Ts --> G,h\\<turnstile>l(vns[\\<mapsto>]vs)[::\\<preceq>]L(vns[\\<mapsto>]Ts)";
|
8011
|
207 |
by( induct_tac "vns" 1);
|
|
208 |
by( ALLGOALS Clarsimp_tac);
|
|
209 |
by( forward_tac [list_all2_lengthD] 1);
|
|
210 |
by( auto_tac (claset(), simpset() addsimps [length_Suc_conv]));
|
|
211 |
qed_spec_mp "lconf_ext_list";
|
|
212 |
|
|
213 |
|
|
214 |
section "oconf";
|
|
215 |
|
|
216 |
val oconf_hext = prove_goalw thy [oconf_def]
|
10042
|
217 |
"!!X. G,h\\<turnstile>obj\\<surd> ==> h\\<le>|h' ==> G,h'\\<turnstile>obj\\<surd>" (K [Fast_tac 1]);
|
8011
|
218 |
|
|
219 |
val oconf_obj = prove_goalw thy [oconf_def,lconf_def] "G,h\\<turnstile>(C,fs)\\<surd> = \
|
10042
|
220 |
\ (\\<forall>T f. map_of(fields (G,C)) f = Some T --> (\\<exists>v. fs f = Some v \\<and> G,h\\<turnstile>v::\\<preceq>T))"(K [
|
8011
|
221 |
Auto_tac]);
|
|
222 |
|
|
223 |
val oconf_objD = oconf_obj RS iffD1 RS spec RS spec RS mp;
|
|
224 |
|
|
225 |
|
|
226 |
section "hconf";
|
|
227 |
|
10042
|
228 |
Goalw [hconf_def] "[|G\\<turnstile>h h\\<surd>; h a = Some obj|] ==> G,h\\<turnstile>obj\\<surd>";
|
8011
|
229 |
by (Fast_tac 1);
|
|
230 |
qed "hconfD";
|
|
231 |
|
10042
|
232 |
Goalw [hconf_def] "\\<forall>a obj. h a=Some obj --> G,h\\<turnstile>obj\\<surd> ==> G\\<turnstile>h h\\<surd>";
|
8011
|
233 |
by (Fast_tac 1);
|
|
234 |
qed "hconfI";
|
|
235 |
|
|
236 |
|
|
237 |
section "conforms";
|
|
238 |
|
|
239 |
val conforms_heapD = prove_goalw thy [conforms_def]
|
10042
|
240 |
"(h, l)::\\<preceq>(G, lT) ==> G\\<turnstile>h h\\<surd>"
|
8011
|
241 |
(fn prems => [cut_facts_tac prems 1, Asm_full_simp_tac 1]);
|
|
242 |
|
|
243 |
val conforms_localD = prove_goalw thy [conforms_def]
|
10042
|
244 |
"(h, l)::\\<preceq>(G, lT) ==> G,h\\<turnstile>l[::\\<preceq>]lT" (fn prems => [
|
8011
|
245 |
cut_facts_tac prems 1, Asm_full_simp_tac 1]);
|
|
246 |
|
|
247 |
val conformsI = prove_goalw thy [conforms_def]
|
10042
|
248 |
"[|G\\<turnstile>h h\\<surd>; G,h\\<turnstile>l[::\\<preceq>]lT|] ==> (h, l)::\\<preceq>(G, lT)" (fn prems => [
|
8011
|
249 |
cut_facts_tac prems 1,
|
|
250 |
Simp_tac 1,
|
|
251 |
Auto_tac]);
|
|
252 |
|
10042
|
253 |
Goal "[|(h,l)::\\<preceq>(G,lT); h\\<le>|h'; G\\<turnstile>h h'\\<surd> |] ==> (h',l)::\\<preceq>(G,lT)";
|
8011
|
254 |
by( fast_tac (HOL_cs addDs [conforms_localD]
|
|
255 |
addSEs [conformsI, lconf_hext]) 1);
|
|
256 |
qed "conforms_hext";
|
|
257 |
|
10042
|
258 |
Goal "[|(h,l)::\\<preceq>(G, lT); G,h(a\\<mapsto>obj)\\<turnstile>obj\\<surd>; h\\<le>|h(a\\<mapsto>obj)|] ==> (h(a\\<mapsto>obj),l)::\\<preceq>(G, lT)";
|
8011
|
259 |
by( rtac conforms_hext 1);
|
|
260 |
by Auto_tac;
|
|
261 |
by( rtac hconfI 1);
|
|
262 |
by( dtac conforms_heapD 1);
|
|
263 |
by( (auto_tac (HOL_cs addEs [oconf_hext] addDs [hconfD],
|
|
264 |
simpset()delsimps[split_paired_All])));
|
|
265 |
qed "conforms_upd_obj";
|
|
266 |
|
|
267 |
Goalw [conforms_def]
|
10042
|
268 |
"[|(h, l)::\\<preceq>(G, lT); G,h\\<turnstile>v::\\<preceq>T; lT va = Some T|] ==> \
|
|
269 |
\ (h, l(va\\<mapsto>v))::\\<preceq>(G, lT)";
|
8011
|
270 |
by( auto_tac (claset() addEs [lconf_upd], simpset()));
|
|
271 |
qed "conforms_upd_local";
|