author | wenzelm |
Thu, 11 Feb 2010 00:45:02 +0100 | |
changeset 35102 | cc7a0b9f938c |
parent 28524 | 644b62cf678f |
child 33128 | 1f990689349f |
child 35416 | d8d7d1b785af |
permissions | -rwxr-xr-x |
8011 | 1 |
(* Title: HOL/MicroJava/J/Eval.thy |
2 |
ID: $Id$ |
|
3 |
Author: David von Oheimb |
|
4 |
Copyright 1999 Technische Universitaet Muenchen |
|
11070 | 5 |
*) |
8011 | 6 |
|
12911 | 7 |
header {* \isaheader{Operational Evaluation (big step) Semantics} *} |
8011 | 8 |
|
16417 | 9 |
theory Eval imports State WellType begin |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
10 |
|
13672 | 11 |
|
12 |
-- "Auxiliary notions" |
|
13 |
||
14 |
constdefs |
|
15 |
fits :: "java_mb prog \<Rightarrow> state \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60) |
|
16 |
"G,s\<turnstile>a' fits T \<equiv> case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T" |
|
17 |
||
18 |
constdefs |
|
19 |
catch ::"java_mb prog \<Rightarrow> xstate \<Rightarrow> cname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60) |
|
20 |
"G,s\<turnstile>catch C\<equiv> case abrupt s of None \<Rightarrow> False | Some a \<Rightarrow> G,store s\<turnstile> a fits Class C" |
|
21 |
||
22 |
||
23 |
||
24 |
constdefs |
|
25 |
lupd :: "vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state" ("lupd'(_\<mapsto>_')"[10,10]1000) |
|
26 |
"lupd vn v \<equiv> \<lambda> (hp,loc). (hp, (loc(vn\<mapsto>v)))" |
|
27 |
||
28 |
constdefs |
|
29 |
new_xcpt_var :: "vname \<Rightarrow> xstate \<Rightarrow> xstate" |
|
30 |
"new_xcpt_var vn \<equiv> \<lambda>(x,s). Norm (lupd(vn\<mapsto>the x) s)" |
|
31 |
||
32 |
||
33 |
-- "Evaluation relations" |
|
34 |
||
23757 | 35 |
inductive |
10056 | 36 |
eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
37 |
("_ \<turnstile> _ -_\<succ>_-> _" [51,82,60,82,82] 81) |
22271 | 38 |
and evals :: "[java_mb prog,xstate,expr list, |
10056 | 39 |
val list,xstate] => bool " |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
40 |
("_ \<turnstile> _ -_[\<succ>]_-> _" [51,82,60,51,82] 81) |
22271 | 41 |
and exec :: "[java_mb prog,xstate,stmt, xstate] => bool " |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
42 |
("_ \<turnstile> _ -_-> _" [51,82,60,82] 81) |
22271 | 43 |
for G :: "java_mb prog" |
44 |
where |
|
8011 | 45 |
|
22271 | 46 |
-- "evaluation of expressions" |
8011 | 47 |
|
28524 | 48 |
XcptE:"G\<turnstile>(Some xc,s) -e\<succ>undefined-> (Some xc,s)" -- "cf. 15.5" |
8011 | 49 |
|
12517 | 50 |
-- "cf. 15.8.1" |
22271 | 51 |
| NewC: "[| h = heap s; (a,x) = new_Addr h; |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
52 |
h'= h(a\<mapsto>(C,init_vars (fields (G,C)))) |] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
53 |
G\<turnstile>Norm s -NewC C\<succ>Addr a-> c_hupd h' (x,s)" |
8011 | 54 |
|
12517 | 55 |
-- "cf. 15.15" |
22271 | 56 |
| Cast: "[| G\<turnstile>Norm s0 -e\<succ>v-> (x1,s1); |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
57 |
x2 = raise_if (\<not> cast_ok G C (heap s1) v) ClassCast x1 |] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
58 |
G\<turnstile>Norm s0 -Cast C e\<succ>v-> (x2,s1)" |
8011 | 59 |
|
12517 | 60 |
-- "cf. 15.7.1" |
22271 | 61 |
| Lit: "G\<turnstile>Norm s -Lit v\<succ>v-> Norm s" |
8011 | 62 |
|
22271 | 63 |
| BinOp:"[| G\<turnstile>Norm s -e1\<succ>v1-> s1; |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
64 |
G\<turnstile>s1 -e2\<succ>v2-> s2; |
10056 | 65 |
v = (case bop of Eq => Bool (v1 = v2) |
66 |
| Add => Intg (the_Intg v1 + the_Intg v2)) |] ==> |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
67 |
G\<turnstile>Norm s -BinOp bop e1 e2\<succ>v-> s2" |
9240 | 68 |
|
12517 | 69 |
-- "cf. 15.13.1, 15.2" |
22271 | 70 |
| LAcc: "G\<turnstile>Norm s -LAcc v\<succ>the (locals s v)-> Norm s" |
8011 | 71 |
|
12517 | 72 |
-- "cf. 15.25.1" |
22271 | 73 |
| LAss: "[| G\<turnstile>Norm s -e\<succ>v-> (x,(h,l)); |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
74 |
l' = (if x = None then l(va\<mapsto>v) else l) |] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
75 |
G\<turnstile>Norm s -va::=e\<succ>v-> (x,(h,l'))" |
8011 | 76 |
|
12517 | 77 |
-- "cf. 15.10.1, 15.2" |
22271 | 78 |
| FAcc: "[| G\<turnstile>Norm s0 -e\<succ>a'-> (x1,s1); |
10056 | 79 |
v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==> |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
80 |
G\<turnstile>Norm s0 -{T}e..fn\<succ>v-> (np a' x1,s1)" |
8011 | 81 |
|
12517 | 82 |
-- "cf. 15.25.1" |
22271 | 83 |
| FAss: "[| G\<turnstile> Norm s0 -e1\<succ>a'-> (x1,s1); a = the_Addr a'; |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
84 |
G\<turnstile>(np a' x1,s1) -e2\<succ>v -> (x2,s2); |
10056 | 85 |
h = heap s2; (c,fs) = the (h a); |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
86 |
h' = h(a\<mapsto>(c,(fs((fn,T)\<mapsto>v)))) |] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
87 |
G\<turnstile>Norm s0 -{T}e1..fn:=e2\<succ>v-> c_hupd h' (x2,s2)" |
8011 | 88 |
|
12517 | 89 |
-- "cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15" |
22271 | 90 |
| Call: "[| G\<turnstile>Norm s0 -e\<succ>a'-> s1; a = the_Addr a'; |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
91 |
G\<turnstile>s1 -ps[\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a)); |
10056 | 92 |
(md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs)); |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
93 |
G\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\<mapsto>]pvs)(This\<mapsto>a'))) -blk-> s3; |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
94 |
G\<turnstile> s3 -res\<succ>v -> (x4,s4) |] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
95 |
G\<turnstile>Norm s0 -{C}e..mn({pTs}ps)\<succ>v-> (x4,(heap s4,l))" |
8011 | 96 |
|
97 |
||
12517 | 98 |
-- "evaluation of expression lists" |
99 |
||
100 |
-- "cf. 15.5" |
|
28524 | 101 |
| XcptEs:"G\<turnstile>(Some xc,s) -e[\<succ>]undefined-> (Some xc,s)" |
8011 | 102 |
|
12517 | 103 |
-- "cf. 15.11.???" |
22271 | 104 |
| Nil: "G\<turnstile>Norm s0 -[][\<succ>][]-> Norm s0" |
8011 | 105 |
|
12517 | 106 |
-- "cf. 15.6.4" |
22271 | 107 |
| Cons: "[| G\<turnstile>Norm s0 -e \<succ> v -> s1; |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
108 |
G\<turnstile> s1 -es[\<succ>]vs-> s2 |] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
109 |
G\<turnstile>Norm s0 -e#es[\<succ>]v#vs-> s2" |
8011 | 110 |
|
111 |
||
12517 | 112 |
-- "execution of statements" |
113 |
||
114 |
-- "cf. 14.1" |
|
22271 | 115 |
| XcptS:"G\<turnstile>(Some xc,s) -c-> (Some xc,s)" |
8011 | 116 |
|
12517 | 117 |
-- "cf. 14.5" |
22271 | 118 |
| Skip: "G\<turnstile>Norm s -Skip-> Norm s" |
8011 | 119 |
|
12517 | 120 |
-- "cf. 14.7" |
22271 | 121 |
| Expr: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1 |] ==> |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
122 |
G\<turnstile>Norm s0 -Expr e-> s1" |
8011 | 123 |
|
12517 | 124 |
-- "cf. 14.2" |
22271 | 125 |
| Comp: "[| G\<turnstile>Norm s0 -c1-> s1; |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
126 |
G\<turnstile> s1 -c2-> s2|] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
127 |
G\<turnstile>Norm s0 -c1;; c2-> s2" |
8011 | 128 |
|
12517 | 129 |
-- "cf. 14.8.2" |
22271 | 130 |
| Cond: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1; |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
131 |
G\<turnstile> s1 -(if the_Bool v then c1 else c2)-> s2|] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
132 |
G\<turnstile>Norm s0 -If(e) c1 Else c2-> s2" |
8011 | 133 |
|
12517 | 134 |
-- "cf. 14.10, 14.10.1" |
22271 | 135 |
| LoopF:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; \<not>the_Bool v |] ==> |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
136 |
G\<turnstile>Norm s0 -While(e) c-> s1" |
22271 | 137 |
| LoopT:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; the_Bool v; |
12517 | 138 |
G\<turnstile>s1 -c-> s2; G\<turnstile>s2 -While(e) c-> s3 |] ==> |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
139 |
G\<turnstile>Norm s0 -While(e) c-> s3" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
140 |
|
12517 | 141 |
|
11040 | 142 |
lemmas eval_evals_exec_induct = eval_evals_exec.induct [split_format (complete)] |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
143 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
144 |
lemma NewCI: "[|new_Addr (heap s) = (a,x); |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
145 |
s' = c_hupd (heap s(a\<mapsto>(C,init_vars (fields (G,C))))) (x,s)|] ==> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
146 |
G\<turnstile>Norm s -NewC C\<succ>Addr a-> s'" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
147 |
apply (simp (no_asm_simp)) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
148 |
apply (rule eval_evals_exec.NewC) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
149 |
apply auto |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
150 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
151 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
152 |
lemma eval_evals_exec_no_xcpt: |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
153 |
"!!s s'. (G\<turnstile>(x,s) -e \<succ> v -> (x',s') --> x'=None --> x=None) \<and> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
154 |
(G\<turnstile>(x,s) -es[\<succ>]vs-> (x',s') --> x'=None --> x=None) \<and> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
155 |
(G\<turnstile>(x,s) -c -> (x',s') --> x'=None --> x=None)" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
156 |
apply(simp (no_asm_simp) only: split_tupled_all) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
157 |
apply(rule eval_evals_exec_induct) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
158 |
apply(unfold c_hupd_def) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
159 |
apply(simp_all) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
160 |
done |
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
161 |
|
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
162 |
lemma eval_no_xcpt: "G\<turnstile>(x,s) -e\<succ>v-> (None,s') ==> x=None" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
163 |
apply (drule eval_evals_exec_no_xcpt [THEN conjunct1, THEN mp]) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
164 |
apply (fast) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
165 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
166 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
167 |
lemma evals_no_xcpt: "G\<turnstile>(x,s) -e[\<succ>]v-> (None,s') ==> x=None" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
168 |
apply (drule eval_evals_exec_no_xcpt [THEN conjunct2, THEN conjunct1, THEN mp]) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
169 |
apply (fast) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
170 |
done |
8011 | 171 |
|
14141 | 172 |
lemma exec_no_xcpt: "G \<turnstile> (x, s) -c-> (None, s') |
173 |
\<Longrightarrow> x = None" |
|
174 |
apply (drule eval_evals_exec_no_xcpt [THEN conjunct2 [THEN conjunct2], rule_format]) |
|
175 |
apply simp+ |
|
176 |
done |
|
177 |
||
178 |
||
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
179 |
lemma eval_evals_exec_xcpt: |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
180 |
"!!s s'. (G\<turnstile>(x,s) -e \<succ> v -> (x',s') --> x=Some xc --> x'=Some xc \<and> s'=s) \<and> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
181 |
(G\<turnstile>(x,s) -es[\<succ>]vs-> (x',s') --> x=Some xc --> x'=Some xc \<and> s'=s) \<and> |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
182 |
(G\<turnstile>(x,s) -c -> (x',s') --> x=Some xc --> x'=Some xc \<and> s'=s)" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
183 |
apply (simp (no_asm_simp) only: split_tupled_all) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
184 |
apply (rule eval_evals_exec_induct) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
185 |
apply (unfold c_hupd_def) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
186 |
apply (simp_all) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
187 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
188 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
189 |
lemma eval_xcpt: "G\<turnstile>(Some xc,s) -e\<succ>v-> (x',s') ==> x'=Some xc \<and> s'=s" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
190 |
apply (drule eval_evals_exec_xcpt [THEN conjunct1, THEN mp]) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
191 |
apply (fast) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
192 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
193 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
194 |
lemma exec_xcpt: "G\<turnstile>(Some xc,s) -s0-> (x',s') ==> x'=Some xc \<and> s'=s" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
195 |
apply (drule eval_evals_exec_xcpt [THEN conjunct2, THEN conjunct2, THEN mp]) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
196 |
apply (fast) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
197 |
done |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10927
diff
changeset
|
198 |
|
11642 | 199 |
end |