author | haftmann |
Fri, 23 Mar 2007 09:40:47 +0100 | |
changeset 22505 | e2d378a97905 |
parent 18576 | 8d98b7711e47 |
child 25511 | 54db9b5080b8 |
permissions | -rw-r--r-- |
12857 | 1 |
(* Title: HOL/Bali/State.thy |
12854 | 2 |
ID: $Id$ |
3 |
Author: David von Oheimb |
|
4 |
*) |
|
5 |
header {* State for evaluation of Java expressions and statements *} |
|
6 |
||
16417 | 7 |
theory State imports DeclConcepts begin |
12854 | 8 |
|
9 |
text {* |
|
10 |
design issues: |
|
11 |
\begin{itemize} |
|
12 |
\item all kinds of objects (class instances, arrays, and class objects) |
|
13 |
are handeled via a general object abstraction |
|
14 |
\item the heap and the map for class objects are combined into a single table |
|
15 |
@{text "(recall (loc, obj) table \<times> (qtname, obj) table ~= (loc + qtname, obj) table)"} |
|
16 |
\end{itemize} |
|
17 |
*} |
|
18 |
||
19 |
section "objects" |
|
20 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
21 |
datatype obj_tag = --{* tag for generic object *} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
22 |
CInst qtname --{* class instance *} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
23 |
| Arr ty int --{* array with component type and length *} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
24 |
--{* | CStat qtname the tag is irrelevant for a class object, |
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
25 |
i.e. the static fields of a class, |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
26 |
since its type is given already by the reference to |
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
27 |
it (see below) *} |
12854 | 28 |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
29 |
types vn = "fspec + int" --{* variable name *} |
12854 | 30 |
record obj = |
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
31 |
tag :: "obj_tag" --{* generalized object *} |
12854 | 32 |
values :: "(vn, val) table" |
33 |
||
34 |
translations |
|
35 |
"fspec" <= (type) "vname \<times> qtname" |
|
36 |
"vn" <= (type) "fspec + int" |
|
37 |
"obj" <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option\<rparr>" |
|
38 |
"obj" <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option,\<dots>::'a\<rparr>" |
|
39 |
||
40 |
constdefs |
|
41 |
||
42 |
the_Arr :: "obj option \<Rightarrow> ty \<times> int \<times> (vn, val) table" |
|
14766 | 43 |
"the_Arr obj \<equiv> SOME (T,k,t). obj = Some \<lparr>tag=Arr T k,values=t\<rparr>" |
12854 | 44 |
|
45 |
lemma the_Arr_Arr [simp]: "the_Arr (Some \<lparr>tag=Arr T k,values=cs\<rparr>) = (T,k,cs)" |
|
46 |
apply (auto simp: the_Arr_def) |
|
47 |
done |
|
48 |
||
49 |
lemma the_Arr_Arr1 [simp,intro,dest]: |
|
50 |
"\<lbrakk>tag obj = Arr T k\<rbrakk> \<Longrightarrow> the_Arr (Some obj) = (T,k,values obj)" |
|
51 |
apply (auto simp add: the_Arr_def) |
|
52 |
done |
|
53 |
||
54 |
constdefs |
|
55 |
||
56 |
upd_obj :: "vn \<Rightarrow> val \<Rightarrow> obj \<Rightarrow> obj" |
|
57 |
"upd_obj n v \<equiv> \<lambda> obj . obj \<lparr>values:=(values obj)(n\<mapsto>v)\<rparr>" |
|
58 |
||
59 |
lemma upd_obj_def2 [simp]: |
|
60 |
"upd_obj n v obj = obj \<lparr>values:=(values obj)(n\<mapsto>v)\<rparr>" |
|
61 |
apply (auto simp: upd_obj_def) |
|
62 |
done |
|
63 |
||
64 |
constdefs |
|
65 |
obj_ty :: "obj \<Rightarrow> ty" |
|
66 |
"obj_ty obj \<equiv> case tag obj of |
|
67 |
CInst C \<Rightarrow> Class C |
|
68 |
| Arr T k \<Rightarrow> T.[]" |
|
69 |
||
70 |
lemma obj_ty_eq [intro!]: "obj_ty \<lparr>tag=oi,values=x\<rparr> = obj_ty \<lparr>tag=oi,values=y\<rparr>" |
|
71 |
by (simp add: obj_ty_def) |
|
72 |
||
73 |
||
74 |
lemma obj_ty_eq1 [intro!,dest]: |
|
75 |
"tag obj = tag obj' \<Longrightarrow> obj_ty obj = obj_ty obj'" |
|
76 |
by (simp add: obj_ty_def) |
|
77 |
||
78 |
lemma obj_ty_cong [simp]: |
|
79 |
"obj_ty (obj \<lparr>values:=vs\<rparr>) = obj_ty obj" |
|
80 |
by auto |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
81 |
|
12854 | 82 |
lemma obj_ty_CInst [simp]: |
83 |
"obj_ty \<lparr>tag=CInst C,values=vs\<rparr> = Class C" |
|
84 |
by (simp add: obj_ty_def) |
|
85 |
||
86 |
lemma obj_ty_CInst1 [simp,intro!,dest]: |
|
87 |
"\<lbrakk>tag obj = CInst C\<rbrakk> \<Longrightarrow> obj_ty obj = Class C" |
|
88 |
by (simp add: obj_ty_def) |
|
89 |
||
90 |
lemma obj_ty_Arr [simp]: |
|
91 |
"obj_ty \<lparr>tag=Arr T i,values=vs\<rparr> = T.[]" |
|
92 |
by (simp add: obj_ty_def) |
|
93 |
||
94 |
lemma obj_ty_Arr1 [simp,intro!,dest]: |
|
95 |
"\<lbrakk>tag obj = Arr T i\<rbrakk> \<Longrightarrow> obj_ty obj = T.[]" |
|
96 |
by (simp add: obj_ty_def) |
|
97 |
||
98 |
lemma obj_ty_widenD: |
|
99 |
"G\<turnstile>obj_ty obj\<preceq>RefT t \<Longrightarrow> (\<exists>C. tag obj = CInst C) \<or> (\<exists>T k. tag obj = Arr T k)" |
|
100 |
apply (unfold obj_ty_def) |
|
101 |
apply (auto split add: obj_tag.split_asm) |
|
102 |
done |
|
103 |
||
104 |
constdefs |
|
105 |
||
106 |
obj_class :: "obj \<Rightarrow> qtname" |
|
107 |
"obj_class obj \<equiv> case tag obj of |
|
108 |
CInst C \<Rightarrow> C |
|
109 |
| Arr T k \<Rightarrow> Object" |
|
110 |
||
111 |
||
112 |
lemma obj_class_CInst [simp]: "obj_class \<lparr>tag=CInst C,values=vs\<rparr> = C" |
|
113 |
by (auto simp: obj_class_def) |
|
114 |
||
115 |
lemma obj_class_CInst1 [simp,intro!,dest]: |
|
116 |
"tag obj = CInst C \<Longrightarrow> obj_class obj = C" |
|
117 |
by (auto simp: obj_class_def) |
|
118 |
||
119 |
lemma obj_class_Arr [simp]: "obj_class \<lparr>tag=Arr T k,values=vs\<rparr> = Object" |
|
120 |
by (auto simp: obj_class_def) |
|
121 |
||
122 |
lemma obj_class_Arr1 [simp,intro!,dest]: |
|
123 |
"tag obj = Arr T k \<Longrightarrow> obj_class obj = Object" |
|
124 |
by (auto simp: obj_class_def) |
|
125 |
||
126 |
lemma obj_ty_obj_class: "G\<turnstile>obj_ty obj\<preceq> Class statC = G\<turnstile>obj_class obj \<preceq>\<^sub>C statC" |
|
127 |
apply (case_tac "tag obj") |
|
128 |
apply (auto simp add: obj_ty_def obj_class_def) |
|
129 |
apply (case_tac "statC = Object") |
|
130 |
apply (auto dest: widen_Array_Class) |
|
131 |
done |
|
132 |
||
133 |
section "object references" |
|
134 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
135 |
types oref = "loc + qtname" --{* generalized object reference *} |
12854 | 136 |
syntax |
137 |
Heap :: "loc \<Rightarrow> oref" |
|
138 |
Stat :: "qtname \<Rightarrow> oref" |
|
139 |
||
140 |
translations |
|
141 |
"Heap" => "Inl" |
|
142 |
"Stat" => "Inr" |
|
143 |
"oref" <= (type) "loc + qtname" |
|
144 |
||
145 |
constdefs |
|
146 |
fields_table:: |
|
147 |
"prog \<Rightarrow> qtname \<Rightarrow> (fspec \<Rightarrow> field \<Rightarrow> bool) \<Rightarrow> (fspec, ty) table" |
|
148 |
"fields_table G C P |
|
149 |
\<equiv> option_map type \<circ> table_of (filter (split P) (DeclConcepts.fields G C))" |
|
150 |
||
151 |
lemma fields_table_SomeI: |
|
152 |
"\<lbrakk>table_of (DeclConcepts.fields G C) n = Some f; P n f\<rbrakk> |
|
153 |
\<Longrightarrow> fields_table G C P n = Some (type f)" |
|
154 |
apply (unfold fields_table_def) |
|
155 |
apply clarsimp |
|
156 |
apply (rule exI) |
|
157 |
apply (rule conjI) |
|
158 |
apply (erule map_of_filter_in) |
|
159 |
apply assumption |
|
160 |
apply simp |
|
161 |
done |
|
162 |
||
163 |
(* unused *) |
|
164 |
lemma fields_table_SomeD': "fields_table G C P fn = Some T \<Longrightarrow> |
|
165 |
\<exists>f. (fn,f)\<in>set(DeclConcepts.fields G C) \<and> type f = T" |
|
166 |
apply (unfold fields_table_def) |
|
167 |
apply clarsimp |
|
168 |
apply (drule map_of_SomeD) |
|
169 |
apply auto |
|
170 |
done |
|
171 |
||
172 |
lemma fields_table_SomeD: |
|
173 |
"\<lbrakk>fields_table G C P fn = Some T; unique (DeclConcepts.fields G C)\<rbrakk> \<Longrightarrow> |
|
174 |
\<exists>f. table_of (DeclConcepts.fields G C) fn = Some f \<and> type f = T" |
|
175 |
apply (unfold fields_table_def) |
|
176 |
apply clarsimp |
|
177 |
apply (rule exI) |
|
178 |
apply (rule conjI) |
|
179 |
apply (erule table_of_filter_unique_SomeD) |
|
180 |
apply assumption |
|
181 |
apply simp |
|
182 |
done |
|
183 |
||
184 |
constdefs |
|
185 |
in_bounds :: "int \<Rightarrow> int \<Rightarrow> bool" ("(_/ in'_bounds _)" [50, 51] 50) |
|
186 |
"i in_bounds k \<equiv> 0 \<le> i \<and> i < k" |
|
187 |
||
188 |
arr_comps :: "'a \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a option" |
|
189 |
"arr_comps T k \<equiv> \<lambda>i. if i in_bounds k then Some T else None" |
|
190 |
||
191 |
var_tys :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> (vn, ty) table" |
|
192 |
"var_tys G oi r |
|
193 |
\<equiv> case r of |
|
194 |
Heap a \<Rightarrow> (case oi of |
|
195 |
CInst C \<Rightarrow> fields_table G C (\<lambda>n f. \<not>static f) (+) empty |
|
196 |
| Arr T k \<Rightarrow> empty (+) arr_comps T k) |
|
197 |
| Stat C \<Rightarrow> fields_table G C (\<lambda>fn f. declclassf fn = C \<and> static f) |
|
198 |
(+) empty" |
|
199 |
||
200 |
lemma var_tys_Some_eq: |
|
201 |
"var_tys G oi r n = Some T |
|
202 |
= (case r of |
|
203 |
Inl a \<Rightarrow> (case oi of |
|
204 |
CInst C \<Rightarrow> (\<exists>nt. n = Inl nt \<and> fields_table G C (\<lambda>n f. |
|
205 |
\<not>static f) nt = Some T) |
|
206 |
| Arr t k \<Rightarrow> (\<exists> i. n = Inr i \<and> i in_bounds k \<and> t = T)) |
|
207 |
| Inr C \<Rightarrow> (\<exists>nt. n = Inl nt \<and> |
|
208 |
fields_table G C (\<lambda>fn f. declclassf fn = C \<and> static f) nt |
|
209 |
= Some T))" |
|
210 |
apply (unfold var_tys_def arr_comps_def) |
|
211 |
apply (force split add: sum.split_asm sum.split obj_tag.split) |
|
212 |
done |
|
213 |
||
214 |
||
215 |
section "stores" |
|
216 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
217 |
types globs --{* global variables: heap and static variables *} |
12854 | 218 |
= "(oref , obj) table" |
219 |
heap |
|
220 |
= "(loc , obj) table" |
|
13337
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
221 |
(* locals |
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
222 |
= "(lname, val) table" *) (* defined in Value.thy local variables *) |
12854 | 223 |
|
224 |
translations |
|
225 |
"globs" <= (type) "(oref , obj) table" |
|
226 |
"heap" <= (type) "(loc , obj) table" |
|
13337
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
227 |
(* "locals" <= (type) "(lname, val) table" *) |
12854 | 228 |
|
229 |
datatype st = (* pure state, i.e. contents of all variables *) |
|
230 |
st globs locals |
|
231 |
||
232 |
subsection "access" |
|
233 |
||
234 |
constdefs |
|
235 |
||
236 |
globs :: "st \<Rightarrow> globs" |
|
237 |
"globs \<equiv> st_case (\<lambda>g l. g)" |
|
238 |
||
239 |
locals :: "st \<Rightarrow> locals" |
|
240 |
"locals \<equiv> st_case (\<lambda>g l. l)" |
|
241 |
||
242 |
heap :: "st \<Rightarrow> heap" |
|
243 |
"heap s \<equiv> globs s \<circ> Heap" |
|
244 |
||
245 |
||
246 |
lemma globs_def2 [simp]: " globs (st g l) = g" |
|
247 |
by (simp add: globs_def) |
|
248 |
||
249 |
lemma locals_def2 [simp]: "locals (st g l) = l" |
|
250 |
by (simp add: locals_def) |
|
251 |
||
252 |
lemma heap_def2 [simp]: "heap s a=globs s (Heap a)" |
|
253 |
by (simp add: heap_def) |
|
254 |
||
255 |
||
256 |
syntax |
|
257 |
val_this :: "st \<Rightarrow> val" |
|
258 |
lookup_obj :: "st \<Rightarrow> val \<Rightarrow> obj" |
|
259 |
||
260 |
translations |
|
261 |
"val_this s" == "the (locals s This)" |
|
262 |
"lookup_obj s a'" == "the (heap s (the_Addr a'))" |
|
263 |
||
264 |
subsection "memory allocation" |
|
265 |
||
266 |
constdefs |
|
267 |
new_Addr :: "heap \<Rightarrow> loc option" |
|
14766 | 268 |
"new_Addr h \<equiv> if (\<forall>a. h a \<noteq> None) then None else Some (SOME a. h a = None)" |
12854 | 269 |
|
270 |
lemma new_AddrD: "new_Addr h = Some a \<Longrightarrow> h a = None" |
|
18576 | 271 |
apply (auto simp add: new_Addr_def) |
18447 | 272 |
apply (erule someI) |
12854 | 273 |
done |
274 |
||
275 |
lemma new_AddrD2: "new_Addr h = Some a \<Longrightarrow> \<forall>b. h b \<noteq> None \<longrightarrow> b \<noteq> a" |
|
276 |
apply (drule new_AddrD) |
|
277 |
apply auto |
|
278 |
done |
|
279 |
||
280 |
lemma new_Addr_SomeI: "h a = None \<Longrightarrow> \<exists>b. new_Addr h = Some b \<and> h b = None" |
|
18576 | 281 |
apply (simp add: new_Addr_def) |
18447 | 282 |
apply (fast intro: someI2) |
12854 | 283 |
done |
284 |
||
285 |
||
286 |
subsection "initialization" |
|
287 |
||
288 |
syntax |
|
289 |
||
290 |
init_vals :: "('a, ty) table \<Rightarrow> ('a, val) table" |
|
291 |
||
292 |
translations |
|
293 |
"init_vals vs" == "option_map default_val \<circ> vs" |
|
294 |
||
295 |
lemma init_arr_comps_base [simp]: "init_vals (arr_comps T 0) = empty" |
|
296 |
apply (unfold arr_comps_def in_bounds_def) |
|
297 |
apply (rule ext) |
|
298 |
apply auto |
|
299 |
done |
|
300 |
||
301 |
lemma init_arr_comps_step [simp]: |
|
302 |
"0 < j \<Longrightarrow> init_vals (arr_comps T j ) = |
|
303 |
init_vals (arr_comps T (j - 1))(j - 1\<mapsto>default_val T)" |
|
304 |
apply (unfold arr_comps_def in_bounds_def) |
|
305 |
apply (rule ext) |
|
306 |
apply auto |
|
307 |
done |
|
308 |
||
309 |
subsection "update" |
|
310 |
||
311 |
constdefs |
|
312 |
gupd :: "oref \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st" ("gupd'(_\<mapsto>_')"[10,10]1000) |
|
313 |
"gupd r obj \<equiv> st_case (\<lambda>g l. st (g(r\<mapsto>obj)) l)" |
|
314 |
||
315 |
lupd :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" ("lupd'(_\<mapsto>_')"[10,10]1000) |
|
316 |
"lupd vn v \<equiv> st_case (\<lambda>g l. st g (l(vn\<mapsto>v)))" |
|
317 |
||
318 |
upd_gobj :: "oref \<Rightarrow> vn \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" |
|
319 |
"upd_gobj r n v \<equiv> st_case (\<lambda>g l. st (chg_map (upd_obj n v) r g) l)" |
|
320 |
||
321 |
set_locals :: "locals \<Rightarrow> st \<Rightarrow> st" |
|
322 |
"set_locals l \<equiv> st_case (\<lambda>g l'. st g l)" |
|
323 |
||
324 |
init_obj :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> st \<Rightarrow> st" |
|
325 |
"init_obj G oi r \<equiv> gupd(r\<mapsto>\<lparr>tag=oi, values=init_vals (var_tys G oi r)\<rparr>)" |
|
326 |
||
327 |
syntax |
|
328 |
init_class_obj :: "prog \<Rightarrow> qtname \<Rightarrow> st \<Rightarrow> st" |
|
329 |
||
330 |
translations |
|
331 |
"init_class_obj G C" == "init_obj G arbitrary (Inr C)" |
|
332 |
||
333 |
lemma gupd_def2 [simp]: "gupd(r\<mapsto>obj) (st g l) = st (g(r\<mapsto>obj)) l" |
|
334 |
apply (unfold gupd_def) |
|
335 |
apply (simp (no_asm)) |
|
336 |
done |
|
337 |
||
338 |
lemma lupd_def2 [simp]: "lupd(vn\<mapsto>v) (st g l) = st g (l(vn\<mapsto>v))" |
|
339 |
apply (unfold lupd_def) |
|
340 |
apply (simp (no_asm)) |
|
341 |
done |
|
342 |
||
343 |
lemma globs_gupd [simp]: "globs (gupd(r\<mapsto>obj) s) = globs s(r\<mapsto>obj)" |
|
344 |
apply (induct "s") |
|
345 |
by (simp add: gupd_def) |
|
346 |
||
347 |
lemma globs_lupd [simp]: "globs (lupd(vn\<mapsto>v ) s) = globs s" |
|
348 |
apply (induct "s") |
|
349 |
by (simp add: lupd_def) |
|
350 |
||
351 |
lemma locals_gupd [simp]: "locals (gupd(r\<mapsto>obj) s) = locals s" |
|
352 |
apply (induct "s") |
|
353 |
by (simp add: gupd_def) |
|
354 |
||
355 |
lemma locals_lupd [simp]: "locals (lupd(vn\<mapsto>v ) s) = locals s(vn\<mapsto>v )" |
|
356 |
apply (induct "s") |
|
357 |
by (simp add: lupd_def) |
|
358 |
||
359 |
lemma globs_upd_gobj_new [rule_format (no_asm), simp]: |
|
360 |
"globs s r = None \<longrightarrow> globs (upd_gobj r n v s) = globs s" |
|
361 |
apply (unfold upd_gobj_def) |
|
362 |
apply (induct "s") |
|
363 |
apply auto |
|
364 |
done |
|
365 |
||
366 |
lemma globs_upd_gobj_upd [rule_format (no_asm), simp]: |
|
367 |
"globs s r=Some obj\<longrightarrow> globs (upd_gobj r n v s) = globs s(r\<mapsto>upd_obj n v obj)" |
|
368 |
apply (unfold upd_gobj_def) |
|
369 |
apply (induct "s") |
|
370 |
apply auto |
|
371 |
done |
|
372 |
||
373 |
lemma locals_upd_gobj [simp]: "locals (upd_gobj r n v s) = locals s" |
|
374 |
apply (induct "s") |
|
375 |
by (simp add: upd_gobj_def) |
|
376 |
||
377 |
||
378 |
lemma globs_init_obj [simp]: "globs (init_obj G oi r s) t = |
|
379 |
(if t=r then Some \<lparr>tag=oi,values=init_vals (var_tys G oi r)\<rparr> else globs s t)" |
|
380 |
apply (unfold init_obj_def) |
|
381 |
apply (simp (no_asm)) |
|
382 |
done |
|
383 |
||
384 |
lemma locals_init_obj [simp]: "locals (init_obj G oi r s) = locals s" |
|
385 |
by (simp add: init_obj_def) |
|
386 |
||
387 |
lemma surjective_st [simp]: "st (globs s) (locals s) = s" |
|
388 |
apply (induct "s") |
|
389 |
by auto |
|
390 |
||
391 |
lemma surjective_st_init_obj: |
|
392 |
"st (globs (init_obj G oi r s)) (locals s) = init_obj G oi r s" |
|
393 |
apply (subst locals_init_obj [THEN sym]) |
|
394 |
apply (rule surjective_st) |
|
395 |
done |
|
396 |
||
397 |
lemma heap_heap_upd [simp]: |
|
398 |
"heap (st (g(Inl a\<mapsto>obj)) l) = heap (st g l)(a\<mapsto>obj)" |
|
399 |
apply (rule ext) |
|
400 |
apply (simp (no_asm)) |
|
401 |
done |
|
402 |
lemma heap_stat_upd [simp]: "heap (st (g(Inr C\<mapsto>obj)) l) = heap (st g l)" |
|
403 |
apply (rule ext) |
|
404 |
apply (simp (no_asm)) |
|
405 |
done |
|
406 |
lemma heap_local_upd [simp]: "heap (st g (l(vn\<mapsto>v))) = heap (st g l)" |
|
407 |
apply (rule ext) |
|
408 |
apply (simp (no_asm)) |
|
409 |
done |
|
410 |
||
411 |
lemma heap_gupd_Heap [simp]: "heap (gupd(Heap a\<mapsto>obj) s) = heap s(a\<mapsto>obj)" |
|
412 |
apply (rule ext) |
|
413 |
apply (simp (no_asm)) |
|
414 |
done |
|
415 |
lemma heap_gupd_Stat [simp]: "heap (gupd(Stat C\<mapsto>obj) s) = heap s" |
|
416 |
apply (rule ext) |
|
417 |
apply (simp (no_asm)) |
|
418 |
done |
|
419 |
lemma heap_lupd [simp]: "heap (lupd(vn\<mapsto>v) s) = heap s" |
|
420 |
apply (rule ext) |
|
421 |
apply (simp (no_asm)) |
|
422 |
done |
|
423 |
||
424 |
lemma heap_upd_gobj_Stat [simp]: "heap (upd_gobj (Stat C) n v s) = heap s" |
|
425 |
apply (rule ext) |
|
426 |
apply (simp (no_asm)) |
|
427 |
apply (case_tac "globs s (Stat C)") |
|
428 |
apply auto |
|
429 |
done |
|
430 |
||
431 |
lemma set_locals_def2 [simp]: "set_locals l (st g l') = st g l" |
|
432 |
apply (unfold set_locals_def) |
|
433 |
apply (simp (no_asm)) |
|
434 |
done |
|
435 |
||
436 |
lemma set_locals_id [simp]: "set_locals (locals s) s = s" |
|
437 |
apply (unfold set_locals_def) |
|
438 |
apply (induct_tac "s") |
|
439 |
apply (simp (no_asm)) |
|
440 |
done |
|
441 |
||
442 |
lemma set_set_locals [simp]: "set_locals l (set_locals l' s) = set_locals l s" |
|
443 |
apply (unfold set_locals_def) |
|
444 |
apply (induct_tac "s") |
|
445 |
apply (simp (no_asm)) |
|
446 |
done |
|
447 |
||
448 |
lemma locals_set_locals [simp]: "locals (set_locals l s) = l" |
|
449 |
apply (unfold set_locals_def) |
|
450 |
apply (induct_tac "s") |
|
451 |
apply (simp (no_asm)) |
|
452 |
done |
|
453 |
||
454 |
lemma globs_set_locals [simp]: "globs (set_locals l s) = globs s" |
|
455 |
apply (unfold set_locals_def) |
|
456 |
apply (induct_tac "s") |
|
457 |
apply (simp (no_asm)) |
|
458 |
done |
|
459 |
||
460 |
lemma heap_set_locals [simp]: "heap (set_locals l s) = heap s" |
|
461 |
apply (unfold heap_def) |
|
462 |
apply (induct_tac "s") |
|
463 |
apply (simp (no_asm)) |
|
464 |
done |
|
465 |
||
466 |
||
467 |
section "abrupt completion" |
|
468 |
||
469 |
||
470 |
||
471 |
consts |
|
472 |
||
473 |
the_Xcpt :: "abrupt \<Rightarrow> xcpt" |
|
474 |
the_Jump :: "abrupt => jump" |
|
475 |
the_Loc :: "xcpt \<Rightarrow> loc" |
|
476 |
the_Std :: "xcpt \<Rightarrow> xname" |
|
477 |
||
478 |
primrec "the_Xcpt (Xcpt x) = x" |
|
479 |
primrec "the_Jump (Jump j) = j" |
|
480 |
primrec "the_Loc (Loc a) = a" |
|
481 |
primrec "the_Std (Std x) = x" |
|
482 |
||
13337
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
483 |
|
12854 | 484 |
|
485 |
||
486 |
constdefs |
|
487 |
abrupt_if :: "bool \<Rightarrow> abopt \<Rightarrow> abopt \<Rightarrow> abopt" |
|
488 |
"abrupt_if c x' x \<equiv> if c \<and> (x = None) then x' else x" |
|
489 |
||
490 |
lemma abrupt_if_True_None [simp]: "abrupt_if True x None = x" |
|
491 |
by (simp add: abrupt_if_def) |
|
492 |
||
493 |
lemma abrupt_if_True_not_None [simp]: "x \<noteq> None \<Longrightarrow> abrupt_if True x y \<noteq> None" |
|
494 |
by (simp add: abrupt_if_def) |
|
495 |
||
496 |
lemma abrupt_if_False [simp]: "abrupt_if False x y = y" |
|
497 |
by (simp add: abrupt_if_def) |
|
498 |
||
499 |
lemma abrupt_if_Some [simp]: "abrupt_if c x (Some y) = Some y" |
|
500 |
by (simp add: abrupt_if_def) |
|
501 |
||
502 |
lemma abrupt_if_not_None [simp]: "y \<noteq> None \<Longrightarrow> abrupt_if c x y = y" |
|
503 |
apply (simp add: abrupt_if_def) |
|
504 |
by auto |
|
505 |
||
506 |
||
507 |
lemma split_abrupt_if: |
|
508 |
"P (abrupt_if c x' x) = |
|
509 |
((c \<and> x = None \<longrightarrow> P x') \<and> (\<not> (c \<and> x = None) \<longrightarrow> P x))" |
|
510 |
apply (unfold abrupt_if_def) |
|
511 |
apply (split split_if) |
|
512 |
apply auto |
|
513 |
done |
|
514 |
||
515 |
syntax |
|
516 |
||
517 |
raise_if :: "bool \<Rightarrow> xname \<Rightarrow> abopt \<Rightarrow> abopt" |
|
518 |
np :: "val \<spacespace> \<Rightarrow> abopt \<Rightarrow> abopt" |
|
519 |
check_neg:: "val \<spacespace> \<Rightarrow> abopt \<Rightarrow> abopt" |
|
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
520 |
error_if :: "bool \<Rightarrow> error \<Rightarrow> abopt \<Rightarrow> abopt" |
12854 | 521 |
|
522 |
translations |
|
523 |
||
524 |
"raise_if c xn" == "abrupt_if c (Some (Xcpt (Std xn)))" |
|
525 |
"np v" == "raise_if (v = Null) NullPointer" |
|
526 |
"check_neg i'" == "raise_if (the_Intg i'<0) NegArrSize" |
|
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
527 |
"error_if c e" == "abrupt_if c (Some (Error e))" |
12854 | 528 |
|
529 |
lemma raise_if_None [simp]: "(raise_if c x y = None) = (\<not>c \<and> y = None)" |
|
530 |
apply (simp add: abrupt_if_def) |
|
531 |
by auto |
|
532 |
declare raise_if_None [THEN iffD1, dest!] |
|
533 |
||
534 |
lemma if_raise_if_None [simp]: |
|
535 |
"((if b then y else raise_if c x y) = None) = ((c \<longrightarrow> b) \<and> y = None)" |
|
536 |
apply (simp add: abrupt_if_def) |
|
537 |
apply auto |
|
538 |
done |
|
539 |
||
540 |
lemma raise_if_SomeD [dest!]: |
|
541 |
"raise_if c x y = Some z \<Longrightarrow> c \<and> z=(Xcpt (Std x)) \<and> y=None \<or> (y=Some z)" |
|
542 |
apply (case_tac y) |
|
543 |
apply (case_tac c) |
|
544 |
apply (simp add: abrupt_if_def) |
|
545 |
apply (simp add: abrupt_if_def) |
|
546 |
apply auto |
|
547 |
done |
|
548 |
||
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
549 |
lemma error_if_None [simp]: "(error_if c e y = None) = (\<not>c \<and> y = None)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
550 |
apply (simp add: abrupt_if_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
551 |
by auto |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
552 |
declare error_if_None [THEN iffD1, dest!] |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
553 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
554 |
lemma if_error_if_None [simp]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
555 |
"((if b then y else error_if c e y) = None) = ((c \<longrightarrow> b) \<and> y = None)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
556 |
apply (simp add: abrupt_if_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
557 |
apply auto |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
558 |
done |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
559 |
|
13524 | 560 |
lemma error_if_SomeD [dest!]: |
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
561 |
"error_if c e y = Some z \<Longrightarrow> c \<and> z=(Error e) \<and> y=None \<or> (y=Some z)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
562 |
apply (case_tac y) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
563 |
apply (case_tac c) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
564 |
apply (simp add: abrupt_if_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
565 |
apply (simp add: abrupt_if_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
566 |
apply auto |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
567 |
done |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
568 |
|
12854 | 569 |
constdefs |
570 |
absorb :: "jump \<Rightarrow> abopt \<Rightarrow> abopt" |
|
571 |
"absorb j a \<equiv> if a=Some (Jump j) then None else a" |
|
572 |
||
573 |
lemma absorb_SomeD [dest!]: "absorb j a = Some x \<Longrightarrow> a = Some x" |
|
574 |
by (auto simp add: absorb_def) |
|
575 |
||
576 |
lemma absorb_same [simp]: "absorb j (Some (Jump j)) = None" |
|
577 |
by (auto simp add: absorb_def) |
|
578 |
||
579 |
lemma absorb_other [simp]: "a \<noteq> Some (Jump j) \<Longrightarrow> absorb j a = a" |
|
580 |
by (auto simp add: absorb_def) |
|
581 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
582 |
lemma absorb_Some_NoneD: "absorb j (Some abr) = None \<Longrightarrow> abr = Jump j" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
583 |
by (simp add: absorb_def) |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
584 |
|
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
585 |
lemma absorb_Some_JumpD: "absorb j s = Some (Jump j') \<Longrightarrow> j'\<noteq>j" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
586 |
by (simp add: absorb_def) |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
587 |
|
12854 | 588 |
|
589 |
section "full program state" |
|
590 |
||
591 |
types |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
592 |
state = "abopt \<times> st" --{* state including abruption information *} |
12854 | 593 |
|
594 |
syntax |
|
595 |
Norm :: "st \<Rightarrow> state" |
|
596 |
abrupt :: "state \<Rightarrow> abopt" |
|
597 |
store :: "state \<Rightarrow> st" |
|
598 |
||
599 |
translations |
|
600 |
||
601 |
"Norm s" == "(None,s)" |
|
602 |
"abrupt" => "fst" |
|
603 |
"store" => "snd" |
|
604 |
"abopt" <= (type) "State.abrupt option" |
|
605 |
"abopt" <= (type) "abrupt option" |
|
606 |
"state" <= (type) "abopt \<times> State.st" |
|
607 |
"state" <= (type) "abopt \<times> st" |
|
608 |
||
609 |
||
610 |
||
611 |
lemma single_stateE: "\<forall>Z. Z = (s::state) \<Longrightarrow> False" |
|
612 |
apply (erule_tac x = "(Some k,y)" in all_dupE) |
|
613 |
apply (erule_tac x = "(None,y)" in allE) |
|
614 |
apply clarify |
|
615 |
done |
|
616 |
||
617 |
lemma state_not_single: "All (op = (x::state)) \<Longrightarrow> R" |
|
618 |
apply (drule_tac x = "(if abrupt x = None then Some ?x else None,?y)" in spec) |
|
619 |
apply clarsimp |
|
620 |
done |
|
621 |
||
622 |
constdefs |
|
623 |
||
624 |
normal :: "state \<Rightarrow> bool" |
|
625 |
"normal \<equiv> \<lambda>s. abrupt s = None" |
|
626 |
||
627 |
lemma normal_def2 [simp]: "normal s = (abrupt s = None)" |
|
628 |
apply (unfold normal_def) |
|
629 |
apply (simp (no_asm)) |
|
630 |
done |
|
631 |
||
632 |
constdefs |
|
633 |
heap_free :: "nat \<Rightarrow> state \<Rightarrow> bool" |
|
634 |
"heap_free n \<equiv> \<lambda>s. atleast_free (heap (store s)) n" |
|
635 |
||
636 |
lemma heap_free_def2 [simp]: "heap_free n s = atleast_free (heap (store s)) n" |
|
637 |
apply (unfold heap_free_def) |
|
638 |
apply simp |
|
639 |
done |
|
640 |
||
641 |
subsection "update" |
|
642 |
||
643 |
constdefs |
|
644 |
||
645 |
abupd :: "(abopt \<Rightarrow> abopt) \<Rightarrow> state \<Rightarrow> state" |
|
646 |
"abupd f \<equiv> prod_fun f id" |
|
647 |
||
648 |
supd :: "(st \<Rightarrow> st) \<Rightarrow> state \<Rightarrow> state" |
|
649 |
"supd \<equiv> prod_fun id" |
|
650 |
||
651 |
lemma abupd_def2 [simp]: "abupd f (x,s) = (f x,s)" |
|
652 |
by (simp add: abupd_def) |
|
653 |
||
654 |
lemma abupd_abrupt_if_False [simp]: "\<And> s. abupd (abrupt_if False xo) s = s" |
|
655 |
by simp |
|
656 |
||
657 |
lemma supd_def2 [simp]: "supd f (x,s) = (x,f s)" |
|
658 |
by (simp add: supd_def) |
|
659 |
||
660 |
lemma supd_lupd [simp]: |
|
661 |
"\<And> s. supd (lupd vn v ) s = (abrupt s,lupd vn v (store s))" |
|
662 |
apply (simp (no_asm_simp) only: split_tupled_all) |
|
663 |
apply (simp (no_asm)) |
|
664 |
done |
|
665 |
||
666 |
||
667 |
lemma supd_gupd [simp]: |
|
668 |
"\<And> s. supd (gupd r obj) s = (abrupt s,gupd r obj (store s))" |
|
669 |
apply (simp (no_asm_simp) only: split_tupled_all) |
|
670 |
apply (simp (no_asm)) |
|
671 |
done |
|
672 |
||
673 |
lemma supd_init_obj [simp]: |
|
674 |
"supd (init_obj G oi r) s = (abrupt s,init_obj G oi r (store s))" |
|
675 |
apply (unfold init_obj_def) |
|
676 |
apply (simp (no_asm)) |
|
677 |
done |
|
678 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
679 |
lemma abupd_store_invariant [simp]: "store (abupd f s) = store s" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
680 |
by (cases s) simp |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
681 |
|
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
682 |
lemma supd_abrupt_invariant [simp]: "abrupt (supd f s) = abrupt s" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
683 |
by (cases s) simp |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
684 |
|
12854 | 685 |
syntax |
686 |
||
687 |
set_lvars :: "locals \<Rightarrow> state \<Rightarrow> state" |
|
688 |
restore_lvars :: "state \<Rightarrow> state \<Rightarrow> state" |
|
689 |
||
690 |
translations |
|
691 |
||
692 |
"set_lvars l" == "supd (set_locals l)" |
|
693 |
"restore_lvars s' s" == "set_lvars (locals (store s')) s" |
|
694 |
||
695 |
lemma set_set_lvars [simp]: "\<And> s. set_lvars l (set_lvars l' s) = set_lvars l s" |
|
696 |
apply (simp (no_asm_simp) only: split_tupled_all) |
|
697 |
apply (simp (no_asm)) |
|
698 |
done |
|
699 |
||
700 |
lemma set_lvars_id [simp]: "\<And> s. set_lvars (locals (store s)) s = s" |
|
701 |
apply (simp (no_asm_simp) only: split_tupled_all) |
|
702 |
apply (simp (no_asm)) |
|
703 |
done |
|
704 |
||
705 |
section "initialisation test" |
|
706 |
||
707 |
constdefs |
|
708 |
||
709 |
inited :: "qtname \<Rightarrow> globs \<Rightarrow> bool" |
|
710 |
"inited C g \<equiv> g (Stat C) \<noteq> None" |
|
711 |
||
712 |
initd :: "qtname \<Rightarrow> state \<Rightarrow> bool" |
|
713 |
"initd C \<equiv> inited C \<circ> globs \<circ> store" |
|
714 |
||
715 |
lemma not_inited_empty [simp]: "\<not>inited C empty" |
|
716 |
apply (unfold inited_def) |
|
717 |
apply (simp (no_asm)) |
|
718 |
done |
|
719 |
||
720 |
lemma inited_gupdate [simp]: "inited C (g(r\<mapsto>obj)) = (inited C g \<or> r = Stat C)" |
|
721 |
apply (unfold inited_def) |
|
722 |
apply (auto split add: st.split) |
|
723 |
done |
|
724 |
||
725 |
lemma inited_init_class_obj [intro!]: "inited C (globs (init_class_obj G C s))" |
|
726 |
apply (unfold inited_def) |
|
727 |
apply (simp (no_asm)) |
|
728 |
done |
|
729 |
||
730 |
lemma not_initedD: "\<not> inited C g \<Longrightarrow> g (Stat C) = None" |
|
731 |
apply (unfold inited_def) |
|
732 |
apply (erule notnotD) |
|
733 |
done |
|
734 |
||
735 |
lemma initedD: "inited C g \<Longrightarrow> \<exists> obj. g (Stat C) = Some obj" |
|
736 |
apply (unfold inited_def) |
|
737 |
apply auto |
|
738 |
done |
|
739 |
||
740 |
lemma initd_def2 [simp]: "initd C s = inited C (globs (store s))" |
|
741 |
apply (unfold initd_def) |
|
742 |
apply (simp (no_asm)) |
|
743 |
done |
|
744 |
||
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
745 |
section {* @{text error_free} *} |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
746 |
constdefs error_free:: "state \<Rightarrow> bool" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
747 |
"error_free s \<equiv> \<not> (\<exists> err. abrupt s = Some (Error err))" |
12854 | 748 |
|
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
749 |
lemma error_free_Norm [simp,intro]: "error_free (Norm s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
750 |
by (simp add: error_free_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
751 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
752 |
lemma error_free_normal [simp,intro]: "normal s \<Longrightarrow> error_free s" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
753 |
by (simp add: error_free_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
754 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
755 |
lemma error_free_Xcpt [simp]: "error_free (Some (Xcpt x),s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
756 |
by (simp add: error_free_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
757 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
758 |
lemma error_free_Jump [simp,intro]: "error_free (Some (Jump j),s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
759 |
by (simp add: error_free_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
760 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
761 |
lemma error_free_Error [simp]: "error_free (Some (Error e),s) = False" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
762 |
by (simp add: error_free_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
763 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
764 |
lemma error_free_Some [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
765 |
"\<not> (\<exists> err. x=Error err) \<Longrightarrow> error_free ((Some x),s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
766 |
by (auto simp add: error_free_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
767 |
|
13524 | 768 |
lemma error_free_abupd_absorb [simp,intro]: |
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
769 |
"error_free s \<Longrightarrow> error_free (abupd (absorb j) s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
770 |
by (cases s) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
771 |
(auto simp add: error_free_def absorb_def |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
772 |
split: split_if_asm) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
773 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
774 |
lemma error_free_absorb [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
775 |
"error_free (a,s) \<Longrightarrow> error_free (absorb j a, s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
776 |
by (auto simp add: error_free_def absorb_def |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
777 |
split: split_if_asm) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
778 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
779 |
lemma error_free_abrupt_if [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
780 |
"\<lbrakk>error_free s; \<not> (\<exists> err. x=Error err)\<rbrakk> |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
781 |
\<Longrightarrow> error_free (abupd (abrupt_if p (Some x)) s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
782 |
by (cases s) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
783 |
(auto simp add: abrupt_if_def |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
784 |
split: split_if) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
785 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
786 |
lemma error_free_abrupt_if1 [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
787 |
"\<lbrakk>error_free (a,s); \<not> (\<exists> err. x=Error err)\<rbrakk> |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
788 |
\<Longrightarrow> error_free (abrupt_if p (Some x) a, s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
789 |
by (auto simp add: abrupt_if_def |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
790 |
split: split_if) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
791 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
792 |
lemma error_free_abrupt_if_Xcpt [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
793 |
"error_free s |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
794 |
\<Longrightarrow> error_free (abupd (abrupt_if p (Some (Xcpt x))) s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
795 |
by simp |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
796 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
797 |
lemma error_free_abrupt_if_Xcpt1 [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
798 |
"error_free (a,s) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
799 |
\<Longrightarrow> error_free (abrupt_if p (Some (Xcpt x)) a, s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
800 |
by simp |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
801 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
802 |
lemma error_free_abrupt_if_Jump [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
803 |
"error_free s |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
804 |
\<Longrightarrow> error_free (abupd (abrupt_if p (Some (Jump j))) s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
805 |
by simp |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
806 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
807 |
lemma error_free_abrupt_if_Jump1 [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
808 |
"error_free (a,s) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
809 |
\<Longrightarrow> error_free (abrupt_if p (Some (Jump j)) a, s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
810 |
by simp |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
811 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
812 |
lemma error_free_raise_if [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
813 |
"error_free s \<Longrightarrow> error_free (abupd (raise_if p x) s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
814 |
by simp |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
815 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
816 |
lemma error_free_raise_if1 [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
817 |
"error_free (a,s) \<Longrightarrow> error_free ((raise_if p x a), s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
818 |
by simp |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
819 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
820 |
lemma error_free_supd [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
821 |
"error_free s \<Longrightarrow> error_free (supd f s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
822 |
by (cases s) (simp add: error_free_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
823 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
824 |
lemma error_free_supd1 [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
825 |
"error_free (a,s) \<Longrightarrow> error_free (a,f s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
826 |
by (simp add: error_free_def) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
827 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
828 |
lemma error_free_set_lvars [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
829 |
"error_free s \<Longrightarrow> error_free ((set_lvars l) s)" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
830 |
by (cases s) simp |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
831 |
|
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
832 |
lemma error_free_set_locals [simp,intro]: |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
833 |
"error_free (x, s) |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
834 |
\<Longrightarrow> error_free (x, set_locals l s')" |
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12858
diff
changeset
|
835 |
by (simp add: error_free_def) |
12854 | 836 |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13524
diff
changeset
|
837 |
|
12854 | 838 |
end |
839 |