author | oheimb |
Tue, 09 Jan 2001 13:54:44 +0100 | |
changeset 10828 | b207d6d1bedc |
parent 10763 | 08e1610c1dcb |
child 10927 | 33e290a70445 |
permissions | -rw-r--r-- |
8011 | 1 |
(* Title: HOL/MicroJava/J/Eval.thy |
2 |
ID: $Id$ |
|
3 |
Author: David von Oheimb |
|
4 |
Copyright 1999 Technische Universitaet Muenchen |
|
5 |
||
6 |
Operational evaluation (big-step) semantics of the |
|
7 |
execution of Java expressions and statements |
|
8 |
*) |
|
9 |
||
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
9240
diff
changeset
|
10 |
Eval = State + WellType + |
8011 | 11 |
consts |
10042 | 12 |
eval :: "java_mb prog => (xstate \\<times> expr \\<times> val \\<times> xstate) set" |
13 |
evals :: "java_mb prog => (xstate \\<times> expr list \\<times> val list \\<times> xstate) set" |
|
14 |
exec :: "java_mb prog => (xstate \\<times> stmt \\<times> xstate) set" |
|
8011 | 15 |
|
16 |
syntax |
|
10056 | 17 |
eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " |
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
18 |
("_ \\<turnstile> _ -_\\<succ>_-> _" [51,82,60,82,82] 81) |
8082 | 19 |
evals:: "[java_mb prog,xstate,expr list, |
10056 | 20 |
val list,xstate] => bool " |
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
21 |
("_ \\<turnstile> _ -_[\\<succ>]_-> _" [51,82,60,51,82] 81) |
10056 | 22 |
exec :: "[java_mb prog,xstate,stmt, xstate] => bool " |
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
23 |
("_ \\<turnstile> _ -_-> _" [51,82,60,82] 81) |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10056
diff
changeset
|
24 |
|
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10056
diff
changeset
|
25 |
syntax (HTML) |
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10056
diff
changeset
|
26 |
eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " |
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
27 |
("_ |- _ -_>_-> _" [51,82,60,82,82] 81) |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10056
diff
changeset
|
28 |
evals:: "[java_mb prog,xstate,expr list, |
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10056
diff
changeset
|
29 |
val list,xstate] => bool " |
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
30 |
("_ |- _ -_[>]_-> _" [51,82,60,51,82] 81) |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10056
diff
changeset
|
31 |
exec :: "[java_mb prog,xstate,stmt, xstate] => bool " |
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
32 |
("_ |- _ -_-> _" [51,82,60,82] 81) |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10056
diff
changeset
|
33 |
|
8011 | 34 |
|
35 |
translations |
|
10042 | 36 |
"G\\<turnstile>s -e \\<succ> v-> (x,s')" <= "(s, e, v, x, s') \\<in> eval G" |
10056 | 37 |
"G\\<turnstile>s -e \\<succ> v-> s' " == "(s, e, v, s') \\<in> eval G" |
10042 | 38 |
"G\\<turnstile>s -e[\\<succ>]v-> (x,s')" <= "(s, e, v, x, s') \\<in> evals G" |
10056 | 39 |
"G\\<turnstile>s -e[\\<succ>]v-> s' " == "(s, e, v, s') \\<in> evals G" |
40 |
"G\\<turnstile>s -c -> (x,s')" <= "(s, c, x, s') \\<in> exec G" |
|
41 |
"G\\<turnstile>s -c -> s' " == "(s, c, s') \\<in> exec G" |
|
8011 | 42 |
|
43 |
inductive "eval G" "evals G" "exec G" intrs |
|
44 |
||
45 |
(* evaluation of expressions *) |
|
46 |
||
47 |
(* cf. 15.5 *) |
|
10056 | 48 |
XcptE "G\\<turnstile>(Some xc,s) -e\\<succ>arbitrary-> (Some xc,s)" |
8011 | 49 |
|
50 |
(* cf. 15.8.1 *) |
|
10056 | 51 |
NewC "[| h = heap s; (a,x) = new_Addr h; |
52 |
h'= h(a\\<mapsto>(C,init_vars (fields (G,C)))) |] ==> |
|
53 |
G\\<turnstile>Norm s -NewC C\\<succ>Addr a-> c_hupd h' (x,s)" |
|
8011 | 54 |
|
55 |
(* cf. 15.15 *) |
|
10056 | 56 |
Cast "[| G\\<turnstile>Norm s0 -e\\<succ>v-> (x1,s1); |
57 |
x2 = raise_if (\\<not> cast_ok G C (heap s1) v) ClassCast x1 |] ==> |
|
58 |
G\\<turnstile>Norm s0 -Cast C e\\<succ>v-> (x2,s1)" |
|
8011 | 59 |
|
60 |
(* cf. 15.7.1 *) |
|
10056 | 61 |
Lit "G\\<turnstile>Norm s -Lit v\\<succ>v-> Norm s" |
8011 | 62 |
|
10056 | 63 |
BinOp "[| G\\<turnstile>Norm s -e1\\<succ>v1-> s1; |
64 |
G\\<turnstile>s1 -e2\\<succ>v2-> s2; |
|
65 |
v = (case bop of Eq => Bool (v1 = v2) |
|
66 |
| Add => Intg (the_Intg v1 + the_Intg v2)) |] ==> |
|
67 |
G\\<turnstile>Norm s -BinOp bop e1 e2\\<succ>v-> s2" |
|
9240 | 68 |
|
8011 | 69 |
(* cf. 15.13.1, 15.2 *) |
10056 | 70 |
LAcc "G\\<turnstile>Norm s -LAcc v\\<succ>the (locals s v)-> Norm s" |
8011 | 71 |
|
72 |
(* cf. 15.25.1 *) |
|
10056 | 73 |
LAss "[| G\\<turnstile>Norm s -e\\<succ>v-> (x,(h,l)); |
74 |
l' = (if x = None then l(va\\<mapsto>v) else l) |] ==> |
|
75 |
G\\<turnstile>Norm s -va::=e\\<succ>v-> (x,(h,l'))" |
|
8011 | 76 |
|
77 |
||
78 |
(* cf. 15.10.1, 15.2 *) |
|
10056 | 79 |
FAcc "[| G\\<turnstile>Norm s0 -e\\<succ>a'-> (x1,s1); |
80 |
v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==> |
|
81 |
G\\<turnstile>Norm s0 -{T}e..fn\\<succ>v-> (np a' x1,s1)" |
|
8011 | 82 |
|
83 |
(* cf. 15.25.1 *) |
|
10056 | 84 |
FAss "[| G\\<turnstile> Norm s0 -e1\\<succ>a'-> (x1,s1); a = the_Addr a'; |
85 |
G\\<turnstile>(np a' x1,s1) -e2\\<succ>v -> (x2,s2); |
|
86 |
h = heap s2; (c,fs) = the (h a); |
|
87 |
h' = h(a\\<mapsto>(c,(fs((fn,T)\\<mapsto>v)))) |] ==> |
|
88 |
G\\<turnstile>Norm s0 -{T}e1..fn:=e2\\<succ>v-> c_hupd h' (x2,s2)" |
|
8011 | 89 |
|
90 |
(* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *) |
|
10056 | 91 |
Call "[| G\\<turnstile>Norm s0 -e\\<succ>a'-> s1; a = the_Addr a'; |
92 |
G\\<turnstile>s1 -ps[\\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a)); |
|
93 |
(md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs)); |
|
94 |
G\\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\\<mapsto>]pvs)(This\\<mapsto>a'))) -blk-> s3; |
|
95 |
G\\<turnstile> s3 -res\\<succ>v -> (x4,s4) |] ==> |
|
10763 | 96 |
G\\<turnstile>Norm s0 -{C}e..mn({pTs}ps)\\<succ>v-> (x4,(heap s4,l))" |
8011 | 97 |
|
98 |
||
99 |
(* evaluation of expression lists *) |
|
100 |
||
101 |
(* cf. 15.5 *) |
|
10056 | 102 |
XcptEs "G\\<turnstile>(Some xc,s) -e[\\<succ>]arbitrary-> (Some xc,s)" |
8011 | 103 |
|
104 |
(* cf. 15.11.??? *) |
|
10056 | 105 |
Nil "G\\<turnstile>Norm s0 -[][\\<succ>][]-> Norm s0" |
8011 | 106 |
|
107 |
(* cf. 15.6.4 *) |
|
10056 | 108 |
Cons "[| G\\<turnstile>Norm s0 -e \\<succ> v -> s1; |
109 |
G\\<turnstile> s1 -es[\\<succ>]vs-> s2 |] ==> |
|
110 |
G\\<turnstile>Norm s0 -e#es[\\<succ>]v#vs-> s2" |
|
8011 | 111 |
|
112 |
(* execution of statements *) |
|
113 |
||
114 |
(* cf. 14.1 *) |
|
10056 | 115 |
XcptS "G\\<turnstile>(Some xc,s) -s0-> (Some xc,s)" |
8011 | 116 |
|
117 |
(* cf. 14.5 *) |
|
10056 | 118 |
Skip "G\\<turnstile>Norm s -Skip-> Norm s" |
8011 | 119 |
|
120 |
(* cf. 14.7 *) |
|
10056 | 121 |
Expr "[| G\\<turnstile>Norm s0 -e\\<succ>v-> s1 |] ==> |
122 |
G\\<turnstile>Norm s0 -Expr e-> s1" |
|
8011 | 123 |
|
124 |
(* cf. 14.2 *) |
|
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
125 |
Comp "[| G\\<turnstile>Norm s0 -s-> s1; |
10056 | 126 |
G\\<turnstile> s1 -t -> s2|] ==> |
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
127 |
G\\<turnstile>Norm s0 -s;; t-> s2" |
8011 | 128 |
|
129 |
(* cf. 14.8.2 *) |
|
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
130 |
Cond "[| G\\<turnstile>Norm s0 -e\\<succ>v-> s1; |
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
131 |
G\\<turnstile> s1 -(if the_Bool v then s else t)-> s2|] ==> |
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
132 |
G\\<turnstile>Norm s0 -If(e) s Else t-> s2" |
8011 | 133 |
|
134 |
(* cf. 14.10, 14.10.1 *) |
|
10828
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
135 |
Loop "[| G\\<turnstile>Norm s0 -e\\<succ>v-> s1; |
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
136 |
if the_Bool v then G\\<turnstile>s1 -s-> s2 \\<and> G\\<turnstile>s2 -While(e) s-> s3 |
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
137 |
else s3 = s1 |] ==> |
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
138 |
G\\<turnstile>Norm s0 -While(e) s-> s3" |
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
139 |
|
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
140 |
monos |
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
oheimb
parents:
10763
diff
changeset
|
141 |
if_def2 |
8011 | 142 |
|
143 |
end |