diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/NanoJava/State.thy --- a/src/HOL/NanoJava/State.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/NanoJava/State.thy Mon Mar 01 13:42:31 2010 +0100 @@ -7,9 +7,7 @@ theory State imports TypeRel begin -constdefs - - body :: "cname \ mname => stmt" +definition body :: "cname \ mname => stmt" where "body \ \(C,m). bdy (the (method C m))" text {* Locations, i.e.\ abstract references to objects *} @@ -29,9 +27,7 @@ "fields" \ (type)"fname => val option" "obj" \ (type)"cname \ fields" -constdefs - - init_vars:: "('a \ 'b) => ('a \ val)" +definition init_vars :: "('a \ 'b) => ('a \ val)" where "init_vars m == Option.map (\T. Null) o m" text {* private: *} @@ -49,54 +45,49 @@ "locals" \ (type)"vname => val option" "state" \ (type)"(|heap :: heap, locals :: locals|)" -constdefs - - del_locs :: "state => state" +definition del_locs :: "state => state" where "del_locs s \ s (| locals := empty |)" - init_locs :: "cname => mname => state => state" +definition init_locs :: "cname => mname => state => state" where "init_locs C m s \ s (| locals := locals s ++ init_vars (map_of (lcl (the (method C m)))) |)" text {* The first parameter of @{term set_locs} is of type @{typ state} rather than @{typ locals} in order to keep @{typ locals} private.*} -constdefs - set_locs :: "state => state => state" +definition set_locs :: "state => state => state" where "set_locs s s' \ s' (| locals := locals s |)" - get_local :: "state => vname => val" ("_<_>" [99,0] 99) +definition get_local :: "state => vname => val" ("_<_>" [99,0] 99) where "get_local s x \ the (locals s x)" --{* local function: *} - get_obj :: "state => loc => obj" +definition get_obj :: "state => loc => obj" where "get_obj s a \ the (heap s a)" - obj_class :: "state => loc => cname" +definition obj_class :: "state => loc => cname" where "obj_class s a \ fst (get_obj s a)" - get_field :: "state => loc => fname => val" +definition get_field :: "state => loc => fname => val" where "get_field s a f \ the (snd (get_obj s a) f)" --{* local function: *} - hupd :: "loc => obj => state => state" ("hupd'(_|->_')" [10,10] 1000) +definition hupd :: "loc => obj => state => state" ("hupd'(_|->_')" [10,10] 1000) where "hupd a obj s \ s (| heap := ((heap s)(a\obj))|)" - lupd :: "vname => val => state => state" ("lupd'(_|->_')" [10,10] 1000) +definition lupd :: "vname => val => state => state" ("lupd'(_|->_')" [10,10] 1000) where "lupd x v s \ s (| locals := ((locals s)(x\v ))|)" notation (xsymbols) hupd ("hupd'(_\_')" [10,10] 1000) and lupd ("lupd'(_\_')" [10,10] 1000) -constdefs - - new_obj :: "loc => cname => state => state" +definition new_obj :: "loc => cname => state => state" where "new_obj a C \ hupd(a\(C,init_vars (field C)))" - upd_obj :: "loc => fname => val => state => state" +definition upd_obj :: "loc => fname => val => state => state" where "upd_obj a f v s \ let (C,fs) = the (heap s a) in hupd(a\(C,fs(f\v))) s" - new_Addr :: "state => val" +definition new_Addr :: "state => val" where "new_Addr s == SOME v. (\a. v = Addr a \ (heap s) a = None) | v = Null"