author | wenzelm |
Wed, 04 Oct 2017 12:00:53 +0200 | |
changeset 66787 | 64b47495676d |
parent 62042 | 6c6ccf573479 |
child 67443 | 3abf6a722518 |
permissions | -rw-r--r-- |
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10061
diff
changeset
|
1 |
(* Title: HOL/MicroJava/J/Value.thy |
41589 | 2 |
Author: David von Oheimb, Technische Universitaet Muenchen |
11070 | 3 |
*) |
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
4 |
|
61361 | 5 |
section \<open>Java Values\<close> |
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
6 |
|
16417 | 7 |
theory Value imports Type begin |
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
8 |
|
62042 | 9 |
typedecl loc' \<comment> "locations, i.e. abstract references on objects" |
12517 | 10 |
|
58310 | 11 |
datatype loc |
62042 | 12 |
= XcptRef xcpt \<comment> "special locations for pre-allocated system exceptions" |
13 |
| Loc loc' \<comment> "usual locations (references on objects)" |
|
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
14 |
|
58310 | 15 |
datatype val |
62042 | 16 |
= Unit \<comment> "dummy result value of void methods" |
17 |
| Null \<comment> "null reference" |
|
18 |
| Bool bool \<comment> "Boolean value" |
|
19 |
| Intg int \<comment> "integer value, name Intg instead of Int because |
|
12517 | 20 |
of clash with HOL/Set.thy" |
62042 | 21 |
| Addr loc \<comment> "addresses, i.e. locations of objects " |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10042
diff
changeset
|
22 |
|
39758 | 23 |
primrec the_Bool :: "val => bool" where |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10042
diff
changeset
|
24 |
"the_Bool (Bool b) = b" |
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
25 |
|
39758 | 26 |
primrec the_Intg :: "val => int" where |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10042
diff
changeset
|
27 |
"the_Intg (Intg i) = i" |
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
28 |
|
39758 | 29 |
primrec the_Addr :: "val => loc" where |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10042
diff
changeset
|
30 |
"the_Addr (Addr a) = a" |
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
31 |
|
62042 | 32 |
primrec defpval :: "prim_ty => val" \<comment> "default value for primitive types" where |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10042
diff
changeset
|
33 |
"defpval Void = Unit" |
39758 | 34 |
| "defpval Boolean = Bool False" |
35 |
| "defpval Integer = Intg 0" |
|
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
36 |
|
62042 | 37 |
primrec default_val :: "ty => val" \<comment> "default value for all types" where |
10061
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
kleing
parents:
10042
diff
changeset
|
38 |
"default_val (PrimT pt) = defpval pt" |
39758 | 39 |
| "default_val (RefT r ) = Null" |
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
40 |
|
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
diff
changeset
|
41 |
end |