Theory State

Up to index of Isabelle/HOL/MicroJava

theory State
imports TypeRel Value

(*  Title:      HOL/MicroJava/J/State.thy
Author: David von Oheimb
Copyright 1999 Technische Universitaet Muenchen
*)


header {* \isaheader{Program State} *}

theory State
imports TypeRel Value
begin


types
fields' = "(vname × cname \<rightharpoonup> val)" -- "field name, defining class, value"

obj = "cname × fields'" -- "class instance with class name and fields"


definition obj_ty :: "obj => ty" where
"obj_ty obj == Class (fst obj)"


definition init_vars :: "('a × ty) list => ('a \<rightharpoonup> val)" where
"init_vars == map_of o map (λ(n,T). (n,default_val T))"


types aheap = "loc \<rightharpoonup> obj" -- {* "@{text heap}" used in a translation below *}
locals = "vname \<rightharpoonup> val" -- "simple state, i.e. variable contents"

state = "aheap × locals" -- "heap, local parameter including This"
xstate = "val option × state" -- "state including exception information"


abbreviation (input)
heap :: "state => aheap"
where "heap == fst"


abbreviation (input)
locals :: "state => locals"
where "locals == snd"


abbreviation "Norm s == (None, s)"

abbreviation (input)
abrupt :: "xstate => val option"
where "abrupt == fst"


abbreviation (input)
store :: "xstate => state"
where "store == snd"


abbreviation
lookup_obj :: "state => val => obj"
where "lookup_obj s a' == the (heap s (the_Addr a'))"


definition raise_if :: "bool => xcpt => val option => val option" where
"raise_if b x xo ≡ if b ∧ (xo = None) then Some (Addr (XcptRef x)) else xo"


definition new_Addr :: "aheap => loc × val option" where
"new_Addr h ≡ SOME (a,x). (h a = None ∧ x = None) | x = Some (Addr (XcptRef OutOfMemory))"


definition np :: "val => val option => val option" where
"np v == raise_if (v = Null) NullPointer"


definition c_hupd :: "aheap => xstate => xstate" where
"c_hupd h'== λ(xo,(h,l)). if xo = None then (None,(h',l)) else (xo,(h,l))"


definition cast_ok :: "'c prog => cname => aheap => val => bool" where
"cast_ok G C h v == v = Null ∨ G\<turnstile>obj_ty (the (h (the_Addr v)))\<preceq> Class C"


lemma obj_ty_def2 [simp]: "obj_ty (C,fs) = Class C"
apply (unfold obj_ty_def)
apply (simp (no_asm))
done


lemma new_AddrD: "new_Addr hp = (ref, xcp) ==>
hp ref = None ∧ xcp = None ∨ xcp = Some (Addr (XcptRef OutOfMemory))"

apply (drule sym)
apply (unfold new_Addr_def)
apply(simp add: Pair_fst_snd_eq Eps_split)
apply(rule someI)
apply(rule disjI2)
apply(rule_tac r = "snd (?a,Some (Addr (XcptRef OutOfMemory)))" in trans)
apply auto
done

lemma raise_if_True [simp]: "raise_if True x y ≠ None"
apply (unfold raise_if_def)
apply auto
done

lemma raise_if_False [simp]: "raise_if False x y = y"
apply (unfold raise_if_def)
apply auto
done

lemma raise_if_Some [simp]: "raise_if c x (Some y) ≠ None"
apply (unfold raise_if_def)
apply auto
done

lemma raise_if_Some2 [simp]:
"raise_if c z (if x = None then Some y else x) ≠ None"

apply (unfold raise_if_def)
apply(induct_tac "x")
apply auto
done

lemma raise_if_SomeD [rule_format (no_asm)]:
"raise_if c x y = Some z --> c ∧ Some z = Some (Addr (XcptRef x)) | y = Some z"

apply (unfold raise_if_def)
apply auto
done

lemma raise_if_NoneD [rule_format (no_asm)]:
"raise_if c x y = None --> ¬ c ∧ y = None"

apply (unfold raise_if_def)
apply auto
done

lemma np_NoneD [rule_format (no_asm)]:
"np a' x' = None --> x' = None ∧ a' ≠ Null"

apply (unfold np_def raise_if_def)
apply auto
done

lemma np_None [rule_format (no_asm), simp]: "a' ≠ Null --> np a' x' = x'"
apply (unfold np_def raise_if_def)
apply auto
done

lemma np_Some [simp]: "np a' (Some xc) = Some xc"
apply (unfold np_def raise_if_def)
apply auto
done

lemma np_Null [simp]: "np Null None = Some (Addr (XcptRef NullPointer))"
apply (unfold np_def raise_if_def)
apply auto
done

lemma np_Addr [simp]: "np (Addr a) None = None"
apply (unfold np_def raise_if_def)
apply auto
done

lemma np_raise_if [simp]: "(np Null (raise_if c xc None)) =
Some (Addr (XcptRef (if c then xc else NullPointer)))"

apply (unfold raise_if_def)
apply (simp (no_asm))
done

lemma c_hupd_fst [simp]: "fst (c_hupd h (x, s)) = x"
by (simp add: c_hupd_def split_beta)

end