author | wenzelm |
Wed, 06 Dec 2000 21:32:25 +0100 | |
changeset 10618 | 5b96bc5fbec3 |
parent 10613 | 78b1d6c3ee9c |
permissions | -rw-r--r-- |
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 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
9 |
Goalw [hext_def] |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
10 |
" \\<forall>a C fs . h a = Some (C,fs) --> \ |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
11 |
\ (\\<exists>fs'. h' a = Some (C,fs')) ==> h\\<le>|h'"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
12 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
13 |
qed "hextI"; |
8011 | 14 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
15 |
Goalw [hext_def] "[|h\\<le>|h'; h a = Some (C,fs) |] ==> \\<exists>fs'. h' a = Some (C,fs')"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
16 |
by (Force_tac 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
17 |
qed "hext_objD"; |
8011 | 18 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
19 |
Goal "h\\<le>|h"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
20 |
by (rtac hextI 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
21 |
by (Fast_tac 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
22 |
qed "hext_refl"; |
8011 | 23 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
24 |
Goal "h a = None ==> h\\<le>|h(a\\<mapsto>x)"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
25 |
by (rtac hextI 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
26 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
27 |
qed "hext_new"; |
8011 | 28 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
29 |
Goal "[|h\\<le>|h'; h'\\<le>|h''|] ==> h\\<le>|h''"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
30 |
by (rtac hextI 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
31 |
by (fast_tac (HOL_cs addDs [hext_objD]) 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
32 |
qed "hext_trans"; |
8011 | 33 |
|
34 |
Addsimps [hext_refl, hext_new]; |
|
35 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
36 |
Goal "h a = Some (C,fs) ==> h\\<le>|h(a\\<mapsto>(C,fs'))"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
37 |
by (rtac hextI 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
38 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
39 |
qed "hext_upd_obj"; |
8011 | 40 |
|
41 |
||
42 |
section "conf"; |
|
43 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
44 |
Goalw [conf_def] "G,h\\<turnstile>Null::\\<preceq>T = G\\<turnstile>RefT NullT\\<preceq>T"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
45 |
by (Simp_tac 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
46 |
qed "conf_Null"; |
8011 | 47 |
Addsimps [conf_Null]; |
48 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
49 |
Goalw [conf_def] "typeof (\\<lambda>v. None) v = Some T --> G,h\\<turnstile>v::\\<preceq>T"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
50 |
by (rtac val_.induct 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
51 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
52 |
qed_spec_mp "conf_litval"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
53 |
Addsimps[conf_litval]; |
9240 | 54 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
55 |
Goalw [conf_def] "[|h a = Some obj; G\\<turnstile>obj_ty obj\\<preceq>T|] ==> G,h\\<turnstile>Addr a::\\<preceq>T"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
56 |
by (Asm_full_simp_tac 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
57 |
qed "conf_AddrI"; |
9240 | 58 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
59 |
Goalw [conf_def] "[|h a = Some (C,fs); G\\<turnstile>C\\<preceq>C D|] ==> G,h\\<turnstile>Addr a::\\<preceq> Class D"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
60 |
by (Asm_full_simp_tac 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
61 |
qed "conf_obj_AddrI"; |
8011 | 62 |
|
10042 | 63 |
Goalw [conf_def] "is_type G T --> G,h\\<turnstile>default_val T::\\<preceq>T"; |
8011 | 64 |
by (res_inst_tac [("y","T")] ty.exhaust 1); |
65 |
by (etac ssubst 1); |
|
66 |
by (res_inst_tac [("y","prim_ty")] prim_ty.exhaust 1); |
|
67 |
by (auto_tac (claset(), simpset() addsimps [widen.null])); |
|
68 |
qed_spec_mp "defval_conf"; |
|
69 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
70 |
Goalw [conf_def] |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
71 |
"h a = Some (C,fs) ==> (G,h(a\\<mapsto>(C,fs'))\\<turnstile>x::\\<preceq>T) = (G,h\\<turnstile>x::\\<preceq>T)"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
72 |
by (rtac val_.induct 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
73 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
74 |
qed "conf_upd_obj"; |
8011 | 75 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
76 |
Goalw [conf_def] "wf_prog wf_mb G ==> G,h\\<turnstile>x::\\<preceq>T --> G\\<turnstile>T\\<preceq>T' --> G,h\\<turnstile>x::\\<preceq>T'"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
77 |
by (rtac val_.induct 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
78 |
by (auto_tac (claset() addIs [widen_trans], simpset())); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
79 |
qed_spec_mp "conf_widen"; |
8011 | 80 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
81 |
Goalw [conf_def] "h\\<le>|h' ==> G,h\\<turnstile>v::\\<preceq>T --> G,h'\\<turnstile>v::\\<preceq>T"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
82 |
by (rtac val_.induct 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
83 |
by (auto_tac (claset() addDs [hext_objD], simpset())); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
84 |
qed_spec_mp "conf_hext"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
85 |
|
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
86 |
Goalw [conf_def] "[|h a = None; G,h\\<turnstile>Addr t::\\<preceq>T|] ==> t\\<noteq>a"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
87 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
88 |
qed "new_locD"; |
8011 | 89 |
|
8185 | 90 |
Goalw [conf_def] |
10042 | 91 |
"G,h\\<turnstile>a'::\\<preceq>RefT T --> a' = Null | \ |
8185 | 92 |
\ (\\<exists>a obj T'. a' = Addr a \\<and> h a = Some obj \\<and> obj_ty obj = T' \\<and> G\\<turnstile>T'\\<preceq>RefT T)"; |
93 |
by(induct_tac "a'" 1); |
|
94 |
by(Auto_tac); |
|
95 |
qed_spec_mp "conf_RefTD"; |
|
8011 | 96 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
97 |
Goal "G,h\\<turnstile>a'::\\<preceq>RefT NullT ==> a' = Null"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
98 |
by (dtac conf_RefTD 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
99 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
100 |
qed "conf_NullTD"; |
8011 | 101 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
102 |
Goal "[|a' \\<noteq> Null; G,h\\<turnstile>a'::\\<preceq>RefT t|] ==> \ |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
103 |
\ \\<exists>a C fs. a' = Addr a \\<and> h a = Some (C,fs) \\<and> G\\<turnstile>Class C\\<preceq>RefT t"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
104 |
by (dtac conf_RefTD 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
105 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
106 |
qed "non_npD"; |
8011 | 107 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
108 |
Goal "!!G. [|a' \\<noteq> Null; G,h\\<turnstile>a'::\\<preceq> Class C; C \\<noteq> Object|] ==> \ |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
109 |
\ (\\<exists>a C' fs. a' = Addr a \\<and> h a = Some (C',fs) \\<and> G\\<turnstile>C'\\<preceq>C C)"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
110 |
by (fast_tac (claset() addDs [non_npD]) 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
111 |
qed "non_np_objD"; |
8011 | 112 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
113 |
Goal "a' \\<noteq> Null ==> wf_prog wf_mb G ==> G,h\\<turnstile>a'::\\<preceq>RefT t -->\ |
10042 | 114 |
\ (\\<forall>C. t = ClassT C --> C \\<noteq> Object) --> \ |
8011 | 115 |
\ (\\<exists>a C fs. a' = Addr a \\<and> h a = Some (C,fs) \\<and> G\\<turnstile>Class C\\<preceq>RefT t)"; |
8185 | 116 |
by(res_inst_tac [("y","t")] ref_ty.exhaust 1); |
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
117 |
by (fast_tac (claset() addDs [conf_NullTD]) 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
118 |
by (fast_tac (claset() addDs [non_np_objD]) 1); |
8011 | 119 |
qed_spec_mp "non_np_objD'"; |
120 |
||
10042 | 121 |
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 | 122 |
by(induct_tac "vs" 1); |
123 |
by(ALLGOALS Clarsimp_tac); |
|
124 |
by(forward_tac [list_all2_lengthD RS sym] 1); |
|
125 |
by(full_simp_tac (simpset()addsimps[length_Suc_conv]) 1); |
|
126 |
by(Safe_tac); |
|
127 |
by(forward_tac [list_all2_lengthD RS sym] 1); |
|
128 |
by(full_simp_tac (simpset()addsimps[length_Suc_conv]) 1); |
|
129 |
by(Clarify_tac 1); |
|
130 |
by(fast_tac (claset() addEs [conf_widen]) 1); |
|
8011 | 131 |
qed_spec_mp "conf_list_gext_widen"; |
132 |
||
133 |
||
134 |
section "lconf"; |
|
135 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
136 |
Goalw[lconf_def] "[| G,h\\<turnstile>vs[::\\<preceq>]Ts; Ts n = Some T |] ==> G,h\\<turnstile>(the (vs n))::\\<preceq>T"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
137 |
by (Force_tac 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
138 |
qed "lconfD"; |
8011 | 139 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
140 |
Goalw [lconf_def] "[| G,h\\<turnstile>l[::\\<preceq>]L; h\\<le>|h' |] ==> G,h'\\<turnstile>l[::\\<preceq>]L"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
141 |
by (fast_tac (claset() addEs [conf_hext]) 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
142 |
qed "lconf_hext"; |
8011 | 143 |
AddEs [lconf_hext]; |
144 |
||
10042 | 145 |
Goalw [lconf_def] "!!X. [| G,h\\<turnstile>l[::\\<preceq>]lT; \ |
146 |
\ G,h\\<turnstile>v::\\<preceq>T; lT va = Some T |] ==> G,h\\<turnstile>l(va\\<mapsto>v)[::\\<preceq>]lT"; |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
147 |
by Auto_tac; |
8011 | 148 |
qed "lconf_upd"; |
149 |
||
10042 | 150 |
Goal "\\<forall>x. P x --> R (dv x) x ==> (\\<forall>x. map_of fs f = Some x --> P x) --> \ |
151 |
\ (\\<forall>T. map_of fs f = Some T --> \ |
|
8011 | 152 |
\ (\\<exists>v. map_of (map (\\<lambda>(f,ft). (f, dv ft)) fs) f = Some v \\<and> R v T))"; |
153 |
by( induct_tac "fs" 1); |
|
154 |
by Auto_tac; |
|
155 |
qed_spec_mp "lconf_init_vars_lemma"; |
|
156 |
||
157 |
Goalw [lconf_def, init_vars_def] |
|
10042 | 158 |
"\\<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 | 159 |
by Auto_tac; |
160 |
by( rtac lconf_init_vars_lemma 1); |
|
161 |
by( atac 3); |
|
162 |
by( strip_tac 1); |
|
163 |
by( etac defval_conf 1); |
|
164 |
by Auto_tac; |
|
165 |
qed "lconf_init_vars"; |
|
166 |
AddSIs [lconf_init_vars]; |
|
167 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
168 |
Goalw [lconf_def] "[|G,s\\<turnstile>l[::\\<preceq>]L; G,s\\<turnstile>v::\\<preceq>T|] ==> G,s\\<turnstile>l(vn\\<mapsto>v)[::\\<preceq>]L(vn\\<mapsto>T)"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
169 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
170 |
qed "lconf_ext"; |
8011 | 171 |
|
10042 | 172 |
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 | 173 |
by( induct_tac "vns" 1); |
174 |
by( ALLGOALS Clarsimp_tac); |
|
175 |
by( forward_tac [list_all2_lengthD] 1); |
|
176 |
by( auto_tac (claset(), simpset() addsimps [length_Suc_conv])); |
|
177 |
qed_spec_mp "lconf_ext_list"; |
|
178 |
||
179 |
||
180 |
section "oconf"; |
|
181 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
182 |
Goalw [oconf_def] "G,h\\<turnstile>obj\\<surd> ==> h\\<le>|h' ==> G,h'\\<turnstile>obj\\<surd>"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
183 |
by (Fast_tac 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
184 |
qed "oconf_hext"; |
8011 | 185 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
186 |
Goalw [oconf_def,lconf_def] "G,h\\<turnstile>(C,fs)\\<surd> = \ |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
187 |
\ (\\<forall>T f. map_of(fields (G,C)) f = Some T --> (\\<exists>v. fs f = Some v \\<and> G,h\\<turnstile>v::\\<preceq>T))"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
188 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
189 |
qed "oconf_obj"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
190 |
bind_thm ("oconf_objD", oconf_obj RS iffD1 RS spec RS spec RS mp); |
8011 | 191 |
|
192 |
||
193 |
section "hconf"; |
|
194 |
||
10042 | 195 |
Goalw [hconf_def] "[|G\\<turnstile>h h\\<surd>; h a = Some obj|] ==> G,h\\<turnstile>obj\\<surd>"; |
8011 | 196 |
by (Fast_tac 1); |
197 |
qed "hconfD"; |
|
198 |
||
10042 | 199 |
Goalw [hconf_def] "\\<forall>a obj. h a=Some obj --> G,h\\<turnstile>obj\\<surd> ==> G\\<turnstile>h h\\<surd>"; |
8011 | 200 |
by (Fast_tac 1); |
201 |
qed "hconfI"; |
|
202 |
||
203 |
||
204 |
section "conforms"; |
|
205 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
206 |
Goalw [conforms_def] "(h, l)::\\<preceq>(G, lT) ==> G\\<turnstile>h h\\<surd>"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
207 |
by (Asm_full_simp_tac 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
208 |
qed "conforms_heapD"; |
8011 | 209 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
210 |
Goalw [conforms_def] "(h, l)::\\<preceq>(G, lT) ==> G,h\\<turnstile>l[::\\<preceq>]lT"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
211 |
by (Asm_full_simp_tac 1); |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
212 |
qed "conforms_localD"; |
8011 | 213 |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
214 |
Goalw [conforms_def] "[|G\\<turnstile>h h\\<surd>; G,h\\<turnstile>l[::\\<preceq>]lT|] ==> (h, l)::\\<preceq>(G, lT)"; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
215 |
by Auto_tac; |
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
216 |
qed "conformsI"; |
8011 | 217 |
|
10042 | 218 |
Goal "[|(h,l)::\\<preceq>(G,lT); h\\<le>|h'; G\\<turnstile>h h'\\<surd> |] ==> (h',l)::\\<preceq>(G,lT)"; |
8011 | 219 |
by( fast_tac (HOL_cs addDs [conforms_localD] |
220 |
addSEs [conformsI, lconf_hext]) 1); |
|
221 |
qed "conforms_hext"; |
|
222 |
||
10042 | 223 |
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 | 224 |
by( rtac conforms_hext 1); |
225 |
by Auto_tac; |
|
226 |
by( rtac hconfI 1); |
|
227 |
by( dtac conforms_heapD 1); |
|
228 |
by( (auto_tac (HOL_cs addEs [oconf_hext] addDs [hconfD], |
|
229 |
simpset()delsimps[split_paired_All]))); |
|
230 |
qed "conforms_upd_obj"; |
|
231 |
||
232 |
Goalw [conforms_def] |
|
10042 | 233 |
"[|(h, l)::\\<preceq>(G, lT); G,h\\<turnstile>v::\\<preceq>T; lT va = Some T|] ==> \ |
234 |
\ (h, l(va\\<mapsto>v))::\\<preceq>(G, lT)"; |
|
8011 | 235 |
by( auto_tac (claset() addEs [lconf_upd], simpset())); |
236 |
qed "conforms_upd_local"; |