diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/Conform.thy --- a/src/HOL/MicroJava/J/Conform.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/Conform.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/J/Conform.thy - ID: $Id$ Author: David von Oheimb Copyright 1999 Technische Universitaet Muenchen *) @@ -10,29 +9,27 @@ types 'c env' = "'c prog \ (vname \ ty)" -- "same as @{text env} of @{text WellType.thy}" -constdefs - - hext :: "aheap => aheap => bool" ("_ <=| _" [51,51] 50) +definition hext :: "aheap => aheap => bool" ("_ <=| _" [51,51] 50) where "h<=|h' == \a C fs. h a = Some(C,fs) --> (\fs'. h' a = Some(C,fs'))" - conf :: "'c prog => aheap => val => ty => bool" - ("_,_ |- _ ::<= _" [51,51,51,51] 50) +definition conf :: "'c prog => aheap => val => ty => bool" + ("_,_ |- _ ::<= _" [51,51,51,51] 50) where "G,h|-v::<=T == \T'. typeof (Option.map obj_ty o h) v = Some T' \ G\T'\T" - lconf :: "'c prog => aheap => ('a \ val) => ('a \ ty) => bool" - ("_,_ |- _ [::<=] _" [51,51,51,51] 50) +definition lconf :: "'c prog => aheap => ('a \ val) => ('a \ ty) => bool" + ("_,_ |- _ [::<=] _" [51,51,51,51] 50) where "G,h|-vs[::<=]Ts == \n T. Ts n = Some T --> (\v. vs n = Some v \ G,h|-v::<=T)" - oconf :: "'c prog => aheap => obj => bool" ("_,_ |- _ [ok]" [51,51,51] 50) +definition oconf :: "'c prog => aheap => obj => bool" ("_,_ |- _ [ok]" [51,51,51] 50) where "G,h|-obj [ok] == G,h|-snd obj[::<=]map_of (fields (G,fst obj))" - hconf :: "'c prog => aheap => bool" ("_ |-h _ [ok]" [51,51] 50) +definition hconf :: "'c prog => aheap => bool" ("_ |-h _ [ok]" [51,51] 50) where "G|-h h [ok] == \a obj. h a = Some obj --> G,h|-obj [ok]" - xconf :: "aheap \ val option \ bool" +definition xconf :: "aheap \ val option \ bool" where "xconf hp vo == preallocated hp \ (\ v. (vo = Some v) \ (\ xc. v = (Addr (XcptRef xc))))" - conforms :: "xstate => java_mb env' => bool" ("_ ::<= _" [51,51] 50) +definition conforms :: "xstate => java_mb env' => bool" ("_ ::<= _" [51,51] 50) where "s::<=E == prg E|-h heap (store s) [ok] \ prg E,heap (store s)|-locals (store s)[::<=]localT E \ xconf (heap (store s)) (abrupt s)"