30 "new_xcpt_var vn \<equiv> \<lambda>(x,s). Norm (lupd(vn\<mapsto>the x) s)" |
30 "new_xcpt_var vn \<equiv> \<lambda>(x,s). Norm (lupd(vn\<mapsto>the x) s)" |
31 |
31 |
32 |
32 |
33 -- "Evaluation relations" |
33 -- "Evaluation relations" |
34 |
34 |
35 consts |
35 inductive2 |
36 eval :: "java_mb prog => (xstate \<times> expr \<times> val \<times> xstate) set" |
|
37 evals :: "java_mb prog => (xstate \<times> expr list \<times> val list \<times> xstate) set" |
|
38 exec :: "java_mb prog => (xstate \<times> stmt \<times> xstate) set" |
|
39 |
|
40 syntax (xsymbols) |
|
41 eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " |
36 eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " |
42 ("_ \<turnstile> _ -_\<succ>_-> _" [51,82,60,82,82] 81) |
37 ("_ \<turnstile> _ -_\<succ>_-> _" [51,82,60,82,82] 81) |
43 evals:: "[java_mb prog,xstate,expr list, |
38 and evals :: "[java_mb prog,xstate,expr list, |
44 val list,xstate] => bool " |
39 val list,xstate] => bool " |
45 ("_ \<turnstile> _ -_[\<succ>]_-> _" [51,82,60,51,82] 81) |
40 ("_ \<turnstile> _ -_[\<succ>]_-> _" [51,82,60,51,82] 81) |
46 exec :: "[java_mb prog,xstate,stmt, xstate] => bool " |
41 and exec :: "[java_mb prog,xstate,stmt, xstate] => bool " |
47 ("_ \<turnstile> _ -_-> _" [51,82,60,82] 81) |
42 ("_ \<turnstile> _ -_-> _" [51,82,60,82] 81) |
|
43 for G :: "java_mb prog" |
|
44 where |
48 |
45 |
49 syntax |
46 -- "evaluation of expressions" |
50 eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " |
|
51 ("_ |- _ -_>_-> _" [51,82,60,82,82] 81) |
|
52 evals:: "[java_mb prog,xstate,expr list, |
|
53 val list,xstate] => bool " |
|
54 ("_ |- _ -_[>]_-> _" [51,82,60,51,82] 81) |
|
55 exec :: "[java_mb prog,xstate,stmt, xstate] => bool " |
|
56 ("_ |- _ -_-> _" [51,82,60,82] 81) |
|
57 |
|
58 |
|
59 translations |
|
60 "G\<turnstile>s -e \<succ> v-> (x,s')" <= "(s, e, v, x, s') \<in> eval G" |
|
61 "G\<turnstile>s -e \<succ> v-> s' " == "(s, e, v, s') \<in> eval G" |
|
62 "G\<turnstile>s -e[\<succ>]v-> (x,s')" <= "(s, e, v, x, s') \<in> evals G" |
|
63 "G\<turnstile>s -e[\<succ>]v-> s' " == "(s, e, v, s') \<in> evals G" |
|
64 "G\<turnstile>s -c -> (x,s')" <= "(s, c, x, s') \<in> exec G" |
|
65 "G\<turnstile>s -c -> s' " == "(s, c, s') \<in> exec G" |
|
66 |
|
67 |
|
68 inductive "eval G" "evals G" "exec G" intros |
|
69 |
|
70 (* evaluation of expressions *) |
|
71 |
47 |
72 XcptE:"G\<turnstile>(Some xc,s) -e\<succ>arbitrary-> (Some xc,s)" -- "cf. 15.5" |
48 XcptE:"G\<turnstile>(Some xc,s) -e\<succ>arbitrary-> (Some xc,s)" -- "cf. 15.5" |
73 |
49 |
74 -- "cf. 15.8.1" |
50 -- "cf. 15.8.1" |
75 NewC: "[| h = heap s; (a,x) = new_Addr h; |
51 | NewC: "[| h = heap s; (a,x) = new_Addr h; |
76 h'= h(a\<mapsto>(C,init_vars (fields (G,C)))) |] ==> |
52 h'= h(a\<mapsto>(C,init_vars (fields (G,C)))) |] ==> |
77 G\<turnstile>Norm s -NewC C\<succ>Addr a-> c_hupd h' (x,s)" |
53 G\<turnstile>Norm s -NewC C\<succ>Addr a-> c_hupd h' (x,s)" |
78 |
54 |
79 -- "cf. 15.15" |
55 -- "cf. 15.15" |
80 Cast: "[| G\<turnstile>Norm s0 -e\<succ>v-> (x1,s1); |
56 | Cast: "[| G\<turnstile>Norm s0 -e\<succ>v-> (x1,s1); |
81 x2 = raise_if (\<not> cast_ok G C (heap s1) v) ClassCast x1 |] ==> |
57 x2 = raise_if (\<not> cast_ok G C (heap s1) v) ClassCast x1 |] ==> |
82 G\<turnstile>Norm s0 -Cast C e\<succ>v-> (x2,s1)" |
58 G\<turnstile>Norm s0 -Cast C e\<succ>v-> (x2,s1)" |
83 |
59 |
84 -- "cf. 15.7.1" |
60 -- "cf. 15.7.1" |
85 Lit: "G\<turnstile>Norm s -Lit v\<succ>v-> Norm s" |
61 | Lit: "G\<turnstile>Norm s -Lit v\<succ>v-> Norm s" |
86 |
62 |
87 BinOp:"[| G\<turnstile>Norm s -e1\<succ>v1-> s1; |
63 | BinOp:"[| G\<turnstile>Norm s -e1\<succ>v1-> s1; |
88 G\<turnstile>s1 -e2\<succ>v2-> s2; |
64 G\<turnstile>s1 -e2\<succ>v2-> s2; |
89 v = (case bop of Eq => Bool (v1 = v2) |
65 v = (case bop of Eq => Bool (v1 = v2) |
90 | Add => Intg (the_Intg v1 + the_Intg v2)) |] ==> |
66 | Add => Intg (the_Intg v1 + the_Intg v2)) |] ==> |
91 G\<turnstile>Norm s -BinOp bop e1 e2\<succ>v-> s2" |
67 G\<turnstile>Norm s -BinOp bop e1 e2\<succ>v-> s2" |
92 |
68 |
93 -- "cf. 15.13.1, 15.2" |
69 -- "cf. 15.13.1, 15.2" |
94 LAcc: "G\<turnstile>Norm s -LAcc v\<succ>the (locals s v)-> Norm s" |
70 | LAcc: "G\<turnstile>Norm s -LAcc v\<succ>the (locals s v)-> Norm s" |
95 |
71 |
96 -- "cf. 15.25.1" |
72 -- "cf. 15.25.1" |
97 LAss: "[| G\<turnstile>Norm s -e\<succ>v-> (x,(h,l)); |
73 | LAss: "[| G\<turnstile>Norm s -e\<succ>v-> (x,(h,l)); |
98 l' = (if x = None then l(va\<mapsto>v) else l) |] ==> |
74 l' = (if x = None then l(va\<mapsto>v) else l) |] ==> |
99 G\<turnstile>Norm s -va::=e\<succ>v-> (x,(h,l'))" |
75 G\<turnstile>Norm s -va::=e\<succ>v-> (x,(h,l'))" |
100 |
76 |
101 -- "cf. 15.10.1, 15.2" |
77 -- "cf. 15.10.1, 15.2" |
102 FAcc: "[| G\<turnstile>Norm s0 -e\<succ>a'-> (x1,s1); |
78 | FAcc: "[| G\<turnstile>Norm s0 -e\<succ>a'-> (x1,s1); |
103 v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==> |
79 v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==> |
104 G\<turnstile>Norm s0 -{T}e..fn\<succ>v-> (np a' x1,s1)" |
80 G\<turnstile>Norm s0 -{T}e..fn\<succ>v-> (np a' x1,s1)" |
105 |
81 |
106 -- "cf. 15.25.1" |
82 -- "cf. 15.25.1" |
107 FAss: "[| G\<turnstile> Norm s0 -e1\<succ>a'-> (x1,s1); a = the_Addr a'; |
83 | FAss: "[| G\<turnstile> Norm s0 -e1\<succ>a'-> (x1,s1); a = the_Addr a'; |
108 G\<turnstile>(np a' x1,s1) -e2\<succ>v -> (x2,s2); |
84 G\<turnstile>(np a' x1,s1) -e2\<succ>v -> (x2,s2); |
109 h = heap s2; (c,fs) = the (h a); |
85 h = heap s2; (c,fs) = the (h a); |
110 h' = h(a\<mapsto>(c,(fs((fn,T)\<mapsto>v)))) |] ==> |
86 h' = h(a\<mapsto>(c,(fs((fn,T)\<mapsto>v)))) |] ==> |
111 G\<turnstile>Norm s0 -{T}e1..fn:=e2\<succ>v-> c_hupd h' (x2,s2)" |
87 G\<turnstile>Norm s0 -{T}e1..fn:=e2\<succ>v-> c_hupd h' (x2,s2)" |
112 |
88 |
113 -- "cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15" |
89 -- "cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15" |
114 Call: "[| G\<turnstile>Norm s0 -e\<succ>a'-> s1; a = the_Addr a'; |
90 | Call: "[| G\<turnstile>Norm s0 -e\<succ>a'-> s1; a = the_Addr a'; |
115 G\<turnstile>s1 -ps[\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a)); |
91 G\<turnstile>s1 -ps[\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a)); |
116 (md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs)); |
92 (md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs)); |
117 G\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\<mapsto>]pvs)(This\<mapsto>a'))) -blk-> s3; |
93 G\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\<mapsto>]pvs)(This\<mapsto>a'))) -blk-> s3; |
118 G\<turnstile> s3 -res\<succ>v -> (x4,s4) |] ==> |
94 G\<turnstile> s3 -res\<succ>v -> (x4,s4) |] ==> |
119 G\<turnstile>Norm s0 -{C}e..mn({pTs}ps)\<succ>v-> (x4,(heap s4,l))" |
95 G\<turnstile>Norm s0 -{C}e..mn({pTs}ps)\<succ>v-> (x4,(heap s4,l))" |
120 |
96 |
121 |
97 |
122 -- "evaluation of expression lists" |
98 -- "evaluation of expression lists" |
123 |
99 |
124 -- "cf. 15.5" |
100 -- "cf. 15.5" |
125 XcptEs:"G\<turnstile>(Some xc,s) -e[\<succ>]arbitrary-> (Some xc,s)" |
101 | XcptEs:"G\<turnstile>(Some xc,s) -e[\<succ>]arbitrary-> (Some xc,s)" |
126 |
102 |
127 -- "cf. 15.11.???" |
103 -- "cf. 15.11.???" |
128 Nil: "G\<turnstile>Norm s0 -[][\<succ>][]-> Norm s0" |
104 | Nil: "G\<turnstile>Norm s0 -[][\<succ>][]-> Norm s0" |
129 |
105 |
130 -- "cf. 15.6.4" |
106 -- "cf. 15.6.4" |
131 Cons: "[| G\<turnstile>Norm s0 -e \<succ> v -> s1; |
107 | Cons: "[| G\<turnstile>Norm s0 -e \<succ> v -> s1; |
132 G\<turnstile> s1 -es[\<succ>]vs-> s2 |] ==> |
108 G\<turnstile> s1 -es[\<succ>]vs-> s2 |] ==> |
133 G\<turnstile>Norm s0 -e#es[\<succ>]v#vs-> s2" |
109 G\<turnstile>Norm s0 -e#es[\<succ>]v#vs-> s2" |
134 |
110 |
135 |
111 |
136 -- "execution of statements" |
112 -- "execution of statements" |
137 |
113 |
138 -- "cf. 14.1" |
114 -- "cf. 14.1" |
139 XcptS:"G\<turnstile>(Some xc,s) -c-> (Some xc,s)" |
115 | XcptS:"G\<turnstile>(Some xc,s) -c-> (Some xc,s)" |
140 |
116 |
141 -- "cf. 14.5" |
117 -- "cf. 14.5" |
142 Skip: "G\<turnstile>Norm s -Skip-> Norm s" |
118 | Skip: "G\<turnstile>Norm s -Skip-> Norm s" |
143 |
119 |
144 -- "cf. 14.7" |
120 -- "cf. 14.7" |
145 Expr: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1 |] ==> |
121 | Expr: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1 |] ==> |
146 G\<turnstile>Norm s0 -Expr e-> s1" |
122 G\<turnstile>Norm s0 -Expr e-> s1" |
147 |
123 |
148 -- "cf. 14.2" |
124 -- "cf. 14.2" |
149 Comp: "[| G\<turnstile>Norm s0 -c1-> s1; |
125 | Comp: "[| G\<turnstile>Norm s0 -c1-> s1; |
150 G\<turnstile> s1 -c2-> s2|] ==> |
126 G\<turnstile> s1 -c2-> s2|] ==> |
151 G\<turnstile>Norm s0 -c1;; c2-> s2" |
127 G\<turnstile>Norm s0 -c1;; c2-> s2" |
152 |
128 |
153 -- "cf. 14.8.2" |
129 -- "cf. 14.8.2" |
154 Cond: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1; |
130 | Cond: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1; |
155 G\<turnstile> s1 -(if the_Bool v then c1 else c2)-> s2|] ==> |
131 G\<turnstile> s1 -(if the_Bool v then c1 else c2)-> s2|] ==> |
156 G\<turnstile>Norm s0 -If(e) c1 Else c2-> s2" |
132 G\<turnstile>Norm s0 -If(e) c1 Else c2-> s2" |
157 |
133 |
158 -- "cf. 14.10, 14.10.1" |
134 -- "cf. 14.10, 14.10.1" |
159 LoopF:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; \<not>the_Bool v |] ==> |
135 | LoopF:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; \<not>the_Bool v |] ==> |
160 G\<turnstile>Norm s0 -While(e) c-> s1" |
136 G\<turnstile>Norm s0 -While(e) c-> s1" |
161 LoopT:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; the_Bool v; |
137 | LoopT:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; the_Bool v; |
162 G\<turnstile>s1 -c-> s2; G\<turnstile>s2 -While(e) c-> s3 |] ==> |
138 G\<turnstile>s1 -c-> s2; G\<turnstile>s2 -While(e) c-> s3 |] ==> |
163 G\<turnstile>Norm s0 -While(e) c-> s3" |
139 G\<turnstile>Norm s0 -While(e) c-> s3" |
164 |
140 |
165 |
141 |
166 lemmas eval_evals_exec_induct = eval_evals_exec.induct [split_format (complete)] |
142 lemmas eval_evals_exec_induct = eval_evals_exec.induct [split_format (complete)] |