| 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 | 
 | 
|  |     11 | constdefs
 | 
|  |     12 | 
 | 
|  |     13 |   the_Bool	:: "val \\<Rightarrow> bool"
 | 
|  |     14 |  "the_Bool v  \\<equiv> \\<epsilon>b. v = Bool b"
 | 
|  |     15 | 
 | 
|  |     16 |   the_Int	:: "val \\<Rightarrow> int"
 | 
|  |     17 |  "the_Int  v  \\<equiv> \\<epsilon>i. v = Intg i"
 | 
|  |     18 | 
 | 
|  |     19 |   the_Addr	:: "val \\<Rightarrow> loc"
 | 
|  |     20 |  "the_Addr  v  \\<equiv> \\<epsilon>a. v = Addr a"
 | 
|  |     21 | 
 | 
|  |     22 | consts
 | 
|  |     23 | 
 | 
|  |     24 |   defpval	:: "prim_ty \\<Rightarrow> val"	(* default value for primitive types *)
 | 
|  |     25 |   default_val	:: "ty \\<Rightarrow> val"		(* default value for all types *)
 | 
|  |     26 | 
 | 
|  |     27 | primrec
 | 
|  |     28 | 	"defpval Void    = Unit"
 | 
|  |     29 | 	"defpval Boolean = Bool False"
 | 
|  |     30 | 	"defpval Integer = Intg (#0)"
 | 
|  |     31 | 
 | 
|  |     32 | primrec
 | 
|  |     33 | 	"default_val (PrimT pt) = defpval pt"
 | 
|  |     34 | 	"default_val (RefT  r ) = Null"
 | 
|  |     35 | 
 | 
|  |     36 | types	fields_
 | 
|  |     37 | 	= "(vname \\<times> cname \\<leadsto> val)" (* field name, defining class, value *)
 | 
|  |     38 | 
 | 
|  |     39 | types obj = "cname \\<times> fields_"	(* class instance with class name and fields *)
 | 
|  |     40 | 
 | 
|  |     41 | constdefs
 | 
|  |     42 | 
 | 
|  |     43 |   obj_ty	:: "obj \\<Rightarrow> ty"
 | 
|  |     44 |  "obj_ty obj  \\<equiv> Class (fst obj)"
 | 
|  |     45 | 
 | 
|  |     46 |   init_vars	:: "('a \\<times> ty) list \\<Rightarrow> ('a \\<leadsto> val)"
 | 
|  |     47 |  "init_vars	\\<equiv> map_of o map (\\<lambda>(n,T). (n,default_val T))"
 | 
|  |     48 |   
 | 
|  |     49 | datatype xcpt		(* exceptions *)
 | 
|  |     50 | 	= NullPointer
 | 
|  |     51 | 	| ClassCast
 | 
|  |     52 | 	| OutOfMemory
 | 
|  |     53 | 
 | 
|  |     54 | types	aheap  = "loc \\<leadsto> obj" (* "heap" used in a translation below *)
 | 
|  |     55 |         locals = "vname \\<leadsto> val"	
 | 
|  |     56 | 
 | 
|  |     57 |         state		(* simple state, i.e. variable contents *)
 | 
|  |     58 | 	= "aheap \\<times> locals"
 | 
|  |     59 | 	(* heap, local parameter including This *)
 | 
|  |     60 | 
 | 
|  |     61 | 	xstate		(* state including exception information *)
 | 
|  |     62 | 	 = "xcpt option \\<times> state"
 | 
|  |     63 | 
 | 
|  |     64 | syntax
 | 
|  |     65 |   heap		:: "state \\<Rightarrow> aheap"
 | 
|  |     66 |   locals	:: "state \\<Rightarrow> locals"
 | 
|  |     67 |   Norm		:: "state \\<Rightarrow> xstate"
 | 
|  |     68 | 
 | 
|  |     69 | translations
 | 
|  |     70 |   "heap"	=> "fst"
 | 
|  |     71 |   "locals"	=> "snd"
 | 
|  |     72 |   "Norm s"      == "(None,s)"  
 | 
|  |     73 | 
 | 
|  |     74 | constdefs
 | 
|  |     75 | 
 | 
|  |     76 |   new_Addr	:: "aheap \\<Rightarrow> loc \\<times> xcpt option"
 | 
|  |     77 |  "new_Addr h \\<equiv> \\<epsilon>(a,x). (h a = None \\<and>  x = None) |  x = Some OutOfMemory"
 | 
|  |     78 | 
 | 
|  |     79 |   raise_if	:: "bool \\<Rightarrow> xcpt \\<Rightarrow> xcpt option \\<Rightarrow> xcpt option"
 | 
|  |     80 |  "raise_if c x xo \\<equiv> if c \\<and>  (xo = None) then Some x else xo"
 | 
|  |     81 | 
 | 
|  |     82 |   np		:: "val \\<Rightarrow> xcpt option \\<Rightarrow> xcpt option"
 | 
|  |     83 |  "np v \\<equiv> raise_if (v = Null) NullPointer"
 | 
|  |     84 | 
 | 
|  |     85 |   c_hupd	:: "aheap \\<Rightarrow> xstate \\<Rightarrow> xstate"
 | 
|  |     86 |  "c_hupd h'\\<equiv> \\<lambda>(xo,(h,l)). if xo = None then (None,(h',l)) else (xo,(h,l))"
 | 
|  |     87 | 
 | 
|  |     88 |   cast_ok	:: "'c prog \\<Rightarrow> ty \\<Rightarrow> aheap \\<Rightarrow> val \\<Rightarrow> bool"
 | 
|  |     89 |  "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)"
 | 
|  |     90 | 
 | 
|  |     91 | end
 |