Theory Value
section ‹Java Values›
theory Value imports Type begin
typedecl loc'
datatype loc
= XcptRef xcpt
| Loc loc'
datatype val
= Unit
| Null
| Bool bool
| Intg int
| Addr loc
primrec the_Bool :: "val => bool" where
"the_Bool (Bool b) = b"
primrec the_Intg :: "val => int" where
"the_Intg (Intg i) = i"
primrec the_Addr :: "val => loc" where
"the_Addr (Addr a) = a"
primrec defpval :: "prim_ty => val" where
"defpval Void = Unit"
| "defpval Boolean = Bool False"
| "defpval Integer = Intg 0"
primrec default_val :: "ty => val" where
"default_val (PrimT pt) = defpval pt"
| "default_val (RefT r ) = Null"
end