author | kleing |
Thu, 06 Jul 2000 12:15:05 +0200 | |
changeset 9260 | 678e718a5a86 |
parent 9240 | f4d76cb26433 |
child 9346 | 297dcbf64526 |
permissions | -rw-r--r-- |
8011 | 1 |
(* Title: HOL/MicroJava/J/State.thy |
2 |
ID: $Id$ |
|
3 |
Author: David von Oheimb |
|
4 |
Copyright 1999 Technische Universitaet Muenchen |
|
5 |
||
6 |
State for evaluation of Java expressions and statements |
|
7 |
*) |
|
8 |
||
9 |
State = WellType + |
|
10 |
||
8875
ac86b3d44730
Replaced some definitions involving epsilon by more readable primrec
berghofe
parents:
8011
diff
changeset
|
11 |
consts |
8011 | 12 |
the_Bool :: "val \\<Rightarrow> bool" |
9240 | 13 |
the_Intg :: "val \\<Rightarrow> int" |
8011 | 14 |
the_Addr :: "val \\<Rightarrow> loc" |
15 |
||
16 |
defpval :: "prim_ty \\<Rightarrow> val" (* default value for primitive types *) |
|
17 |
default_val :: "ty \\<Rightarrow> val" (* default value for all types *) |
|
18 |
||
19 |
primrec |
|
8875
ac86b3d44730
Replaced some definitions involving epsilon by more readable primrec
berghofe
parents:
8011
diff
changeset
|
20 |
"the_Bool (Bool b) = b" |
ac86b3d44730
Replaced some definitions involving epsilon by more readable primrec
berghofe
parents:
8011
diff
changeset
|
21 |
|
ac86b3d44730
Replaced some definitions involving epsilon by more readable primrec
berghofe
parents:
8011
diff
changeset
|
22 |
primrec |
9240 | 23 |
"the_Intg (Intg i) = i" |
8875
ac86b3d44730
Replaced some definitions involving epsilon by more readable primrec
berghofe
parents:
8011
diff
changeset
|
24 |
|
ac86b3d44730
Replaced some definitions involving epsilon by more readable primrec
berghofe
parents:
8011
diff
changeset
|
25 |
primrec |
ac86b3d44730
Replaced some definitions involving epsilon by more readable primrec
berghofe
parents:
8011
diff
changeset
|
26 |
"the_Addr (Addr a) = a" |
ac86b3d44730
Replaced some definitions involving epsilon by more readable primrec
berghofe
parents:
8011
diff
changeset
|
27 |
|
ac86b3d44730
Replaced some definitions involving epsilon by more readable primrec
berghofe
parents:
8011
diff
changeset
|
28 |
primrec |
8011 | 29 |
"defpval Void = Unit" |
30 |
"defpval Boolean = Bool False" |
|
31 |
"defpval Integer = Intg (#0)" |
|
32 |
||
33 |
primrec |
|
34 |
"default_val (PrimT pt) = defpval pt" |
|
35 |
"default_val (RefT r ) = Null" |
|
36 |
||
37 |
types fields_ |
|
38 |
= "(vname \\<times> cname \\<leadsto> val)" (* field name, defining class, value *) |
|
39 |
||
40 |
types obj = "cname \\<times> fields_" (* class instance with class name and fields *) |
|
41 |
||
42 |
constdefs |
|
43 |
||
44 |
obj_ty :: "obj \\<Rightarrow> ty" |
|
45 |
"obj_ty obj \\<equiv> Class (fst obj)" |
|
46 |
||
47 |
init_vars :: "('a \\<times> ty) list \\<Rightarrow> ('a \\<leadsto> val)" |
|
48 |
"init_vars \\<equiv> map_of o map (\\<lambda>(n,T). (n,default_val T))" |
|
49 |
||
50 |
datatype xcpt (* exceptions *) |
|
51 |
= NullPointer |
|
52 |
| ClassCast |
|
53 |
| OutOfMemory |
|
54 |
||
55 |
types aheap = "loc \\<leadsto> obj" (* "heap" used in a translation below *) |
|
56 |
locals = "vname \\<leadsto> val" |
|
57 |
||
58 |
state (* simple state, i.e. variable contents *) |
|
59 |
= "aheap \\<times> locals" |
|
60 |
(* heap, local parameter including This *) |
|
61 |
||
62 |
xstate (* state including exception information *) |
|
63 |
= "xcpt option \\<times> state" |
|
64 |
||
65 |
syntax |
|
66 |
heap :: "state \\<Rightarrow> aheap" |
|
67 |
locals :: "state \\<Rightarrow> locals" |
|
68 |
Norm :: "state \\<Rightarrow> xstate" |
|
69 |
||
70 |
translations |
|
71 |
"heap" => "fst" |
|
72 |
"locals" => "snd" |
|
73 |
"Norm s" == "(None,s)" |
|
74 |
||
75 |
constdefs |
|
76 |
||
77 |
new_Addr :: "aheap \\<Rightarrow> loc \\<times> xcpt option" |
|
78 |
"new_Addr h \\<equiv> \\<epsilon>(a,x). (h a = None \\<and> x = None) | x = Some OutOfMemory" |
|
79 |
||
80 |
raise_if :: "bool \\<Rightarrow> xcpt \\<Rightarrow> xcpt option \\<Rightarrow> xcpt option" |
|
81 |
"raise_if c x xo \\<equiv> if c \\<and> (xo = None) then Some x else xo" |
|
82 |
||
83 |
np :: "val \\<Rightarrow> xcpt option \\<Rightarrow> xcpt option" |
|
84 |
"np v \\<equiv> raise_if (v = Null) NullPointer" |
|
85 |
||
86 |
c_hupd :: "aheap \\<Rightarrow> xstate \\<Rightarrow> xstate" |
|
87 |
"c_hupd h'\\<equiv> \\<lambda>(xo,(h,l)). if xo = None then (None,(h',l)) else (xo,(h,l))" |
|
88 |
||
89 |
cast_ok :: "'c prog \\<Rightarrow> ty \\<Rightarrow> aheap \\<Rightarrow> val \\<Rightarrow> bool" |
|
90 |
"cast_ok G T h v \\<equiv> ((\\<exists>pt. T = PrimT pt) | (v=Null) | G\\<turnstile>obj_ty (the (h (the_Addr v)))\\<preceq>T)" |
|
91 |
||
92 |
end |