Theory State

Up to index of Isabelle/HOL/Bali

theory State
imports DeclConcepts
begin

(*  Title:      HOL/Bali/State.thy
    ID:         $Id: State.thy,v 1.14 2007/11/30 19:13:05 haftmann Exp $
    Author:     David von Oheimb
*)
header {* State for evaluation of Java expressions and statements *}

theory State imports DeclConcepts begin

text {*
design issues:
\begin{itemize}
\item all kinds of objects (class instances, arrays, and class objects)
  are handeled via a general object abstraction
\item the heap and the map for class objects are combined into a single table
  @{text "(recall (loc, obj) table × (qtname, obj) table  ~=  (loc + qtname, obj) table)"}
\end{itemize}
*}

section "objects"

datatype  obj_tag =     --{* tag for generic object   *}
          CInst qtname  --{* class instance           *}
        | Arr  ty int   --{* array with component type and length *}
    --{* | CStat qtname   the tag is irrelevant for a class object,
                           i.e. the static fields of a class,
                           since its type is given already by the reference to 
                           it (see below) *}

types   vn   = "fspec + int"                    --{* variable name      *}
record  obj  = 
          tag :: "obj_tag"                      --{* generalized object *}
          values :: "(vn, val) table"      

translations 
  "fspec" <= (type) "vname × qtname" 
  "vn"    <= (type) "fspec + int"
  "obj"   <= (type) "(|tag::obj_tag, values::vn => val option|)),"
  "obj"   <= (type) "(|tag::obj_tag, values::vn => val option,…::'a|)),"

constdefs
  
  the_Arr :: "obj option => ty × int × (vn, val) table"
 "the_Arr obj ≡ SOME (T,k,t). obj = Some (|tag=Arr T k,values=t|)),"

lemma the_Arr_Arr [simp]: "the_Arr (Some (|tag=Arr T k,values=cs|)),) = (T,k,cs)"
apply (auto simp: the_Arr_def)
done

lemma the_Arr_Arr1 [simp,intro,dest]:
 "[|tag obj = Arr T k|] ==> the_Arr (Some obj) = (T,k,values obj)"
apply (auto simp add: the_Arr_def)
done

constdefs

  upd_obj       :: "vn => val => obj => obj" 
 "upd_obj n v ≡ λ obj . obj (|values:=(values obj)(n\<mapsto>v)|)),"

lemma upd_obj_def2 [simp]: 
  "upd_obj n v obj = obj (|values:=(values obj)(n\<mapsto>v)|))," 
apply (auto simp: upd_obj_def)
done

constdefs
  obj_ty        :: "obj => ty"
 "obj_ty obj    ≡ case tag obj of 
                    CInst C => Class C 
                  | Arr T k => T.[]"

lemma obj_ty_eq [intro!]: "obj_ty (|tag=oi,values=x|)), = obj_ty (|tag=oi,values=y|))," 
by (simp add: obj_ty_def)


lemma obj_ty_eq1 [intro!,dest]: 
  "tag obj = tag obj' ==> obj_ty obj = obj_ty obj'" 
by (simp add: obj_ty_def)

lemma obj_ty_cong [simp]: 
  "obj_ty (obj (|values:=vs|)),) = obj_ty obj" 
by auto

lemma obj_ty_CInst [simp]: 
 "obj_ty (|tag=CInst C,values=vs|)), = Class C" 
by (simp add: obj_ty_def)

lemma obj_ty_CInst1 [simp,intro!,dest]: 
 "[|tag obj = CInst C|] ==> obj_ty obj = Class C" 
by (simp add: obj_ty_def)

lemma obj_ty_Arr [simp]: 
 "obj_ty (|tag=Arr T i,values=vs|)), = T.[]"
by (simp add: obj_ty_def)

lemma obj_ty_Arr1 [simp,intro!,dest]: 
 "[|tag obj = Arr T i|] ==> obj_ty obj = T.[]"
by (simp add: obj_ty_def)

lemma obj_ty_widenD: 
 "G\<turnstile>obj_ty obj\<preceq>RefT t ==> (∃C. tag obj = CInst C) ∨ (∃T k. tag obj = Arr T k)"
apply (unfold obj_ty_def)
apply (auto split add: obj_tag.split_asm)
done

constdefs

  obj_class :: "obj => qtname"
 "obj_class obj ≡ case tag obj of 
                    CInst C => C 
                  | Arr T k => Object"


lemma obj_class_CInst [simp]: "obj_class (|tag=CInst C,values=vs|)), = C" 
by (auto simp: obj_class_def)

lemma obj_class_CInst1 [simp,intro!,dest]: 
  "tag obj = CInst C ==> obj_class obj = C" 
by (auto simp: obj_class_def)

lemma obj_class_Arr [simp]: "obj_class (|tag=Arr T k,values=vs|)), = Object" 
by (auto simp: obj_class_def)

lemma obj_class_Arr1 [simp,intro!,dest]: 
 "tag obj = Arr T k ==> obj_class obj = Object" 
by (auto simp: obj_class_def)

lemma obj_ty_obj_class: "G\<turnstile>obj_ty obj\<preceq> Class statC = G\<turnstile>obj_class obj \<preceq>C statC"
apply (case_tac "tag obj")
apply (auto simp add: obj_ty_def obj_class_def)
apply (case_tac "statC = Object")
apply (auto dest: widen_Array_Class)
done

section "object references"

types oref = "loc + qtname"         --{* generalized object reference *}
syntax
  Heap  :: "loc   => oref"
  Stat  :: "qtname => oref"

translations
  "Heap" => "Inl"
  "Stat" => "Inr"
  "oref" <= (type) "loc + qtname"

constdefs
  fields_table::
    "prog => qtname => (fspec => field => bool)  => (fspec, ty) table"
 "fields_table G C P 
    ≡ option_map type o table_of (filter (split P) (DeclConcepts.fields G C))"

lemma fields_table_SomeI: 
"[|table_of (DeclConcepts.fields G C) n = Some f; P n f|] 
 ==> fields_table G C P n = Some (type f)"
apply (unfold fields_table_def)
apply clarsimp
apply (rule exI)
apply (rule conjI)
apply (erule map_of_filter_in)
apply assumption
apply simp
done

(* unused *)
lemma fields_table_SomeD': "fields_table G C P fn = Some T ==>  
  ∃f. (fn,f)∈set(DeclConcepts.fields G C) ∧ type f = T"
apply (unfold fields_table_def)
apply clarsimp
apply (drule map_of_SomeD)
apply auto
done

lemma fields_table_SomeD: 
"[|fields_table G C P fn = Some T; unique (DeclConcepts.fields G C)|] ==>  
  ∃f. table_of (DeclConcepts.fields G C) fn = Some f ∧ type f = T"
apply (unfold fields_table_def)
apply clarsimp
apply (rule exI)
apply (rule conjI)
apply (erule table_of_filter_unique_SomeD)
apply assumption
apply simp
done

constdefs
  in_bounds :: "int => int => bool"            ("(_/ in'_bounds _)" [50, 51] 50)
 "i in_bounds k ≡ 0 ≤ i ∧ i < k"

  arr_comps :: "'a => int => int => 'a option"
 "arr_comps T k ≡ λi. if i in_bounds k then Some T else None"
  
  var_tys       :: "prog => obj_tag => oref => (vn, ty) table"
"var_tys G oi r 
  ≡ case r of 
      Heap a => (case oi of 
                   CInst C => fields_table G C (λn f. ¬static f) (+) empty
                 | Arr T k => empty (+) arr_comps T k)
    | Stat C => fields_table G C (λfn f. declclassf fn = C ∧ static f) 
                (+) empty"

lemma var_tys_Some_eq: 
 "var_tys G oi r n = Some T 
  = (case r of 
       Inl a => (case oi of  
                   CInst C => (∃nt. n = Inl nt ∧ fields_table G C (λn f. 
                               ¬static f) nt = Some T)  
                 | Arr t k => (∃ i. n = Inr i  ∧ i in_bounds k ∧ t = T))  
     | Inr C => (∃nt. n = Inl nt ∧ 
                 fields_table G C (λfn f. declclassf fn = C ∧ static f) nt 
                  = Some T))"
apply (unfold var_tys_def arr_comps_def)
apply (force split add: sum.split_asm sum.split obj_tag.split)
done


section "stores"

types   globs               --{* global variables: heap and static variables *}
        = "(oref , obj) table"
        heap
        = "(loc  , obj) table"
(*      locals                   
        = "(lname, val) table" *) (* defined in Value.thy local variables *)

translations
 "globs"  <= (type) "(oref , obj) table"
 "heap"   <= (type) "(loc  , obj) table"
(*  "locals" <= (type) "(lname, val) table" *)

datatype st = (* pure state, i.e. contents of all variables *)
         st globs locals

subsection "access"

constdefs

  globs  :: "st => globs"
 "globs  ≡ st_case (λg l. g)"
  
  locals :: "st => locals"
 "locals ≡ st_case (λg l. l)"

  heap   :: "st => heap"
 "heap s ≡ globs s o Heap"


lemma globs_def2 [simp]: " globs (st g l) = g"
by (simp add: globs_def)

lemma locals_def2 [simp]: "locals (st g l) = l"
by (simp add: locals_def)

lemma heap_def2 [simp]:  "heap s a=globs s (Heap a)"
by (simp add: heap_def)


syntax
  val_this     :: "st => val"
  lookup_obj   :: "st => val => obj"

translations
 "val_this s"       == "the (locals s This)" 
 "lookup_obj s a'"  == "the (heap s (the_Addr a'))"

subsection "memory allocation"

constdefs
  new_Addr     :: "heap => loc option"
 "new_Addr h   ≡ if (∀a. h a ≠ None) then None else Some (SOME a. h a = None)"

lemma new_AddrD: "new_Addr h = Some a ==> h a = None"
apply (auto simp add: new_Addr_def)
apply (erule someI) 
done

lemma new_AddrD2: "new_Addr h = Some a ==> ∀b. h b ≠ None --> b ≠ a"
apply (drule new_AddrD)
apply auto
done

lemma new_Addr_SomeI: "h a = None ==> ∃b. new_Addr h = Some b ∧ h b = None"
apply (simp add: new_Addr_def)
apply (fast intro: someI2)
done


subsection "initialization"

syntax

  init_vals     :: "('a, ty) table => ('a, val) table"

translations
 "init_vals vs"    == "CONST option_map default_val o vs"

lemma init_arr_comps_base [simp]: "init_vals (arr_comps T 0) = empty"
apply (unfold arr_comps_def in_bounds_def)
apply (rule ext)
apply auto
done

lemma init_arr_comps_step [simp]: 
"0 < j ==> init_vals (arr_comps T  j    ) =  
           init_vals (arr_comps T (j - 1))(j - 1\<mapsto>default_val T)"
apply (unfold arr_comps_def in_bounds_def)
apply (rule ext)
apply auto
done

subsection "update"

constdefs
  gupd       :: "oref  => obj => st => st"        ("gupd'(_\<mapsto>_')"[10,10]1000)
 "gupd r obj  ≡ st_case (λg l. st (g(r\<mapsto>obj)) l)"

  lupd       :: "lname => val => st => st"        ("lupd'(_\<mapsto>_')"[10,10]1000)
 "lupd vn v   ≡ st_case (λg l. st g (l(vn\<mapsto>v)))"

  upd_gobj   :: "oref => vn => val => st => st" 
 "upd_gobj r n v ≡ st_case (λg l. st (chg_map (upd_obj n v) r g) l)"

  set_locals  :: "locals => st => st"
 "set_locals l ≡ st_case (λg l'. st g l)"
  
  init_obj    :: "prog => obj_tag => oref => st => st"
 "init_obj G oi r ≡ gupd(r\<mapsto>(|tag=oi, values=init_vals (var_tys G oi r)|)),)"

syntax
  init_class_obj :: "prog => qtname => st => st"

translations
 "init_class_obj G C" == "init_obj G arbitrary (Inr C)"

lemma gupd_def2 [simp]: "gupd(r\<mapsto>obj) (st g l) = st (g(r\<mapsto>obj)) l"
apply (unfold gupd_def)
apply (simp (no_asm))
done

lemma lupd_def2 [simp]: "lupd(vn\<mapsto>v) (st g l) = st g (l(vn\<mapsto>v))"
apply (unfold lupd_def)
apply (simp (no_asm))
done

lemma globs_gupd [simp]: "globs  (gupd(r\<mapsto>obj) s) = globs s(r\<mapsto>obj)"
apply (induct "s")
by (simp add: gupd_def)

lemma globs_lupd [simp]: "globs  (lupd(vn\<mapsto>v ) s) = globs  s"
apply (induct "s")
by (simp add: lupd_def)

lemma locals_gupd [simp]: "locals (gupd(r\<mapsto>obj) s) = locals s"
apply (induct "s")
by (simp add: gupd_def)

lemma locals_lupd [simp]: "locals (lupd(vn\<mapsto>v ) s) = locals s(vn\<mapsto>v )"
apply (induct "s")
by (simp add: lupd_def)

lemma globs_upd_gobj_new [rule_format (no_asm), simp]: 
  "globs s r = None --> globs (upd_gobj r n v s) = globs s"
apply (unfold upd_gobj_def)
apply (induct "s")
apply auto
done

lemma globs_upd_gobj_upd [rule_format (no_asm), simp]: 
"globs s r=Some obj--> globs (upd_gobj r n v s) = globs s(r\<mapsto>upd_obj n v obj)"
apply (unfold upd_gobj_def)
apply (induct "s")
apply auto
done

lemma locals_upd_gobj [simp]: "locals (upd_gobj r n v s) = locals s"
apply (induct "s")
by (simp add: upd_gobj_def) 


lemma globs_init_obj [simp]: "globs (init_obj G oi r s) t =  
  (if t=r then Some (|tag=oi,values=init_vals (var_tys G oi r)|)), else globs s t)"
apply (unfold init_obj_def)
apply (simp (no_asm))
done

lemma locals_init_obj [simp]: "locals (init_obj G oi r s) = locals s"
by (simp add: init_obj_def)
  
lemma surjective_st [simp]: "st (globs s) (locals s) = s"
apply (induct "s")
by auto

lemma surjective_st_init_obj: 
 "st (globs (init_obj G oi r s)) (locals s) = init_obj G oi r s"
apply (subst locals_init_obj [THEN sym])
apply (rule surjective_st)
done

lemma heap_heap_upd [simp]: 
  "heap (st (g(Inl a\<mapsto>obj)) l) = heap (st g l)(a\<mapsto>obj)"
apply (rule ext)
apply (simp (no_asm))
done
lemma heap_stat_upd [simp]: "heap (st (g(Inr C\<mapsto>obj)) l) = heap (st g l)"
apply (rule ext)
apply (simp (no_asm))
done
lemma heap_local_upd [simp]: "heap (st g (l(vn\<mapsto>v))) = heap (st g l)"
apply (rule ext)
apply (simp (no_asm))
done

lemma heap_gupd_Heap [simp]: "heap (gupd(Heap a\<mapsto>obj) s) = heap s(a\<mapsto>obj)"
apply (rule ext)
apply (simp (no_asm))
done
lemma heap_gupd_Stat [simp]: "heap (gupd(Stat C\<mapsto>obj) s) = heap s"
apply (rule ext)
apply (simp (no_asm))
done
lemma heap_lupd [simp]: "heap (lupd(vn\<mapsto>v) s) = heap s"
apply (rule ext)
apply (simp (no_asm))
done

lemma heap_upd_gobj_Stat [simp]: "heap (upd_gobj (Stat C) n v s) = heap s"
apply (rule ext)
apply (simp (no_asm))
apply (case_tac "globs s (Stat C)")
apply  auto
done

lemma set_locals_def2 [simp]: "set_locals l (st g l') = st g l"
apply (unfold set_locals_def)
apply (simp (no_asm))
done

lemma set_locals_id [simp]: "set_locals (locals s) s = s"
apply (unfold set_locals_def)
apply (induct_tac "s")
apply (simp (no_asm))
done

lemma set_set_locals [simp]: "set_locals l (set_locals l' s) = set_locals l s"
apply (unfold set_locals_def)
apply (induct_tac "s")
apply (simp (no_asm))
done

lemma locals_set_locals [simp]: "locals (set_locals l s) = l"
apply (unfold set_locals_def)
apply (induct_tac "s")
apply (simp (no_asm))
done

lemma globs_set_locals [simp]: "globs (set_locals l s) = globs s"
apply (unfold set_locals_def)
apply (induct_tac "s")
apply (simp (no_asm))
done

lemma heap_set_locals [simp]: "heap (set_locals l s) = heap s"
apply (unfold heap_def)
apply (induct_tac "s")
apply (simp (no_asm))
done


section "abrupt completion"



consts

  the_Xcpt :: "abrupt => xcpt"
  the_Jump :: "abrupt => jump"
  the_Loc  :: "xcpt => loc"
  the_Std  :: "xcpt => xname"

primrec "the_Xcpt (Xcpt x) = x"
primrec "the_Jump (Jump j) = j"
primrec "the_Loc (Loc a) = a"
primrec "the_Std (Std x) = x"


        

constdefs
  abrupt_if    :: "bool => abopt => abopt => abopt"
 "abrupt_if c x' x ≡ if c ∧ (x = None) then x' else x"

lemma abrupt_if_True_None [simp]: "abrupt_if True x None = x"
by (simp add: abrupt_if_def)

lemma abrupt_if_True_not_None [simp]: "x ≠ None ==> abrupt_if True x y ≠ None"
by (simp add: abrupt_if_def)

lemma abrupt_if_False [simp]: "abrupt_if False x y = y"
by (simp add: abrupt_if_def)

lemma abrupt_if_Some [simp]: "abrupt_if c x (Some y) = Some y"
by (simp add: abrupt_if_def)

lemma abrupt_if_not_None [simp]: "y ≠ None ==> abrupt_if c x y = y"
apply (simp add: abrupt_if_def)
by auto


lemma split_abrupt_if: 
"P (abrupt_if c x' x) = 
      ((c ∧ x = None --> P x') ∧ (¬ (c ∧ x = None) --> P x))"
apply (unfold abrupt_if_def)
apply (split split_if)
apply auto
done

syntax

  raise_if :: "bool => xname => abopt => abopt"
  np       :: "val            => abopt => abopt"
  check_neg:: "val            => abopt => abopt"
  error_if :: "bool => error => abopt => abopt"
  
translations

 "raise_if c xn" == "abrupt_if c (Some (Xcpt (Std xn)))"
 "np v"          == "raise_if (v = Null)      NullPointer"
 "check_neg i'"  == "raise_if (the_Intg i'<0) NegArrSize"
 "error_if c e"  == "abrupt_if c (Some (Error e))"

lemma raise_if_None [simp]: "(raise_if c x y = None) = (¬c ∧ y = None)"
apply (simp add: abrupt_if_def)
by auto
declare raise_if_None [THEN iffD1, dest!]

lemma if_raise_if_None [simp]: 
  "((if b then y else raise_if c x y) = None) = ((c --> b) ∧ y = None)"
apply (simp add: abrupt_if_def)
apply auto
done

lemma raise_if_SomeD [dest!]:
  "raise_if c x y = Some z ==> c ∧ z=(Xcpt (Std x)) ∧ y=None ∨ (y=Some z)"
apply (case_tac y)
apply (case_tac c)
apply (simp add: abrupt_if_def)
apply (simp add: abrupt_if_def)
apply auto
done

lemma error_if_None [simp]: "(error_if c e y = None) = (¬c ∧ y = None)"
apply (simp add: abrupt_if_def)
by auto
declare error_if_None [THEN iffD1, dest!]

lemma if_error_if_None [simp]: 
  "((if b then y else error_if c e y) = None) = ((c --> b) ∧ y = None)"
apply (simp add: abrupt_if_def)
apply auto
done

lemma error_if_SomeD [dest!]:
  "error_if c e y = Some z ==> c ∧ z=(Error e) ∧ y=None ∨ (y=Some z)"
apply (case_tac y)
apply (case_tac c)
apply (simp add: abrupt_if_def)
apply (simp add: abrupt_if_def)
apply auto
done

constdefs
   absorb :: "jump => abopt => abopt"
  "absorb j a ≡ if a=Some (Jump j) then None else a"

lemma absorb_SomeD [dest!]: "absorb j a = Some x ==> a = Some x"
by (auto simp add: absorb_def)

lemma absorb_same [simp]: "absorb j (Some (Jump j)) = None"
by (auto simp add: absorb_def)

lemma absorb_other [simp]: "a ≠ Some (Jump j) ==> absorb j a = a"
by (auto simp add: absorb_def)

lemma absorb_Some_NoneD: "absorb j (Some abr) = None ==> abr = Jump j"
  by (simp add: absorb_def)

lemma absorb_Some_JumpD: "absorb j s = Some (Jump j') ==> j'≠j"
  by (simp add: absorb_def)


section "full program state"

types
  state = "abopt × st"          --{* state including abruption information *}

syntax 
  Norm   :: "st => state"
  abrupt :: "state => abopt"
  store  :: "state => st"

translations
   
  "Norm s"     == "(None,s)" 
  "abrupt"     => "fst"
  "store"      => "snd"
  "abopt"       <= (type) "State.abrupt option"
  "abopt"       <= (type) "abrupt option"
  "state"      <= (type) "abopt × State.st"
  "state"      <= (type) "abopt × st"



lemma single_stateE: "∀Z. Z = (s::state) ==> False"
apply (erule_tac x = "(Some k,y)" in all_dupE)
apply (erule_tac x = "(None,y)" in allE)
apply clarify
done

lemma state_not_single: "All (op = (x::state)) ==> R"
apply (drule_tac x = "(if abrupt x = None then Some ?x else None,?y)" in spec)
apply clarsimp
done

constdefs

  normal     :: "state => bool"
 "normal ≡ λs. abrupt s = None"

lemma normal_def2 [simp]: "normal s = (abrupt s = None)"
apply (unfold normal_def)
apply (simp (no_asm))
done

constdefs
  heap_free :: "nat => state => bool"
 "heap_free n ≡ λs. atleast_free (heap (store s)) n"

lemma heap_free_def2 [simp]: "heap_free n s = atleast_free (heap (store s)) n"
apply (unfold heap_free_def)
apply simp
done

subsection "update"

constdefs
 
  abupd     :: "(abopt => abopt) => state => state"
 "abupd f ≡ prod_fun f id"

  supd     :: "(st => st) => state => state" 
 "supd ≡ prod_fun id"
  
lemma abupd_def2 [simp]: "abupd f (x,s) = (f x,s)"
by (simp add: abupd_def)

lemma abupd_abrupt_if_False [simp]: "!! s. abupd (abrupt_if False xo) s = s"
by simp

lemma supd_def2 [simp]: "supd f (x,s) = (x,f s)"
by (simp add: supd_def)

lemma supd_lupd [simp]: 
 "!! s. supd (lupd vn v ) s = (abrupt s,lupd vn v (store s))"
apply (simp (no_asm_simp) only: split_tupled_all)
apply (simp (no_asm))
done


lemma supd_gupd [simp]: 
 "!! s. supd (gupd r obj) s = (abrupt s,gupd r obj (store s))"
apply (simp (no_asm_simp) only: split_tupled_all)
apply (simp (no_asm))
done

lemma supd_init_obj [simp]: 
 "supd (init_obj G oi r) s = (abrupt s,init_obj G oi r (store s))"
apply (unfold init_obj_def)
apply (simp (no_asm))
done

lemma abupd_store_invariant [simp]: "store (abupd f s) = store s"
  by (cases s) simp

lemma supd_abrupt_invariant [simp]: "abrupt (supd f s) = abrupt s"
  by (cases s) simp

syntax

  set_lvars     :: "locals => state => state"
  restore_lvars :: "state  => state => state"
  
translations

 "set_lvars l" == "supd (set_locals l)"
 "restore_lvars s' s" == "set_lvars (locals (store s')) s"

lemma set_set_lvars [simp]: "!! s. set_lvars l (set_lvars l' s) = set_lvars l s"
apply (simp (no_asm_simp) only: split_tupled_all)
apply (simp (no_asm))
done

lemma set_lvars_id [simp]: "!! s. set_lvars (locals (store s)) s = s"
apply (simp (no_asm_simp) only: split_tupled_all)
apply (simp (no_asm))
done

section "initialisation test"

constdefs

  inited   :: "qtname => globs => bool"
 "inited C g ≡ g (Stat C) ≠ None"

  initd    :: "qtname => state => bool"
 "initd C ≡ inited C o globs o store"

lemma not_inited_empty [simp]: "¬inited C empty"
apply (unfold inited_def)
apply (simp (no_asm))
done

lemma inited_gupdate [simp]: "inited C (g(r\<mapsto>obj)) = (inited C g ∨ r = Stat C)"
apply (unfold inited_def)
apply (auto split add: st.split)
done

lemma inited_init_class_obj [intro!]: "inited C (globs (init_class_obj G C s))"
apply (unfold inited_def)
apply (simp (no_asm))
done

lemma not_initedD: "¬ inited C g ==> g (Stat C) = None"
apply (unfold inited_def)
apply (erule notnotD)
done

lemma initedD: "inited C g ==> ∃ obj. g (Stat C) = Some obj"
apply (unfold inited_def)
apply auto
done

lemma initd_def2 [simp]: "initd C s = inited C (globs (store s))"
apply (unfold initd_def)
apply (simp (no_asm))
done

section {* @{text error_free} *}
constdefs error_free:: "state => bool"
"error_free s ≡ ¬ (∃ err. abrupt s = Some (Error err))"

lemma error_free_Norm [simp,intro]: "error_free (Norm s)"
by (simp add: error_free_def)

lemma error_free_normal [simp,intro]: "normal s ==> error_free s"
by (simp add: error_free_def)

lemma error_free_Xcpt [simp]: "error_free (Some (Xcpt x),s)"
by (simp add: error_free_def)

lemma error_free_Jump [simp,intro]: "error_free (Some (Jump j),s)"
by (simp add: error_free_def)

lemma error_free_Error [simp]: "error_free (Some (Error e),s) = False"
by (simp add: error_free_def)  

lemma error_free_Some [simp,intro]: 
 "¬ (∃ err. x=Error err) ==> error_free ((Some x),s)"
by (auto simp add: error_free_def)

lemma error_free_abupd_absorb [simp,intro]: 
 "error_free s ==> error_free (abupd (absorb j) s)"
by (cases s) 
   (auto simp add: error_free_def absorb_def
         split: split_if_asm)

lemma error_free_absorb [simp,intro]: 
 "error_free (a,s) ==> error_free (absorb j a, s)"
by (auto simp add: error_free_def absorb_def
            split: split_if_asm)

lemma error_free_abrupt_if [simp,intro]:
"[|error_free s; ¬ (∃ err. x=Error err)|]
 ==> error_free (abupd (abrupt_if p (Some x)) s)"
by (cases s)
   (auto simp add: abrupt_if_def
            split: split_if)

lemma error_free_abrupt_if1 [simp,intro]:
"[|error_free (a,s); ¬ (∃ err. x=Error err)|]
 ==> error_free (abrupt_if p (Some x) a, s)"
by  (auto simp add: abrupt_if_def
            split: split_if)

lemma error_free_abrupt_if_Xcpt [simp,intro]:
 "error_free s 
  ==> error_free (abupd (abrupt_if p (Some (Xcpt x))) s)"
by simp 

lemma error_free_abrupt_if_Xcpt1 [simp,intro]:
 "error_free (a,s) 
  ==> error_free (abrupt_if p (Some (Xcpt x)) a, s)" 
by simp 

lemma error_free_abrupt_if_Jump [simp,intro]:
 "error_free s 
  ==> error_free (abupd (abrupt_if p (Some (Jump j))) s)" 
by simp

lemma error_free_abrupt_if_Jump1 [simp,intro]:
 "error_free (a,s) 
  ==> error_free (abrupt_if p (Some (Jump j)) a, s)" 
by simp

lemma error_free_raise_if [simp,intro]:
 "error_free s ==> error_free (abupd (raise_if p x) s)"
by simp 

lemma error_free_raise_if1 [simp,intro]:
 "error_free (a,s) ==> error_free ((raise_if p x a), s)"
by simp 

lemma error_free_supd [simp,intro]:
 "error_free s ==> error_free (supd f s)"
by (cases s) (simp add: error_free_def)

lemma error_free_supd1 [simp,intro]:
 "error_free (a,s) ==> error_free (a,f s)"
by (simp add: error_free_def)

lemma error_free_set_lvars [simp,intro]:
"error_free s ==> error_free ((set_lvars l) s)"
by (cases s) simp

lemma error_free_set_locals [simp,intro]: 
"error_free (x, s)
       ==> error_free (x, set_locals l s')"
by (simp add: error_free_def)


end


objects

lemma the_Arr_Arr:

  the_Arr (Some (| tag = Arr T k, values = cs |)) = (T, k, cs)

lemma the_Arr_Arr1:

  tag obj = Arr T k ==> the_Arr (Some obj) = (T, k, values obj)

lemma upd_obj_def2:

  upd_obj n v obj = obj(| values := values obj(n |-> v) |)

lemma obj_ty_eq:

  obj_ty (| tag = oi, values = x |) = obj_ty (| tag = oi, values = y |)

lemma obj_ty_eq1:

  tag obj = tag obj' ==> obj_ty obj = obj_ty obj'

lemma obj_ty_cong:

  obj_ty (obj(| values := vs |)) = obj_ty obj

lemma obj_ty_CInst:

  obj_ty (| tag = CInst C, values = vs |) = Class C

lemma obj_ty_CInst1:

  tag obj = CInst C ==> obj_ty obj = Class C

lemma obj_ty_Arr:

  obj_ty (| tag = Arr T i, values = vs |) = T.[]

lemma obj_ty_Arr1:

  tag obj = Arr T i ==> obj_ty obj = T.[]

lemma obj_ty_widenD:

  G\<turnstile>obj_ty obj\<preceq>RefT t
  ==> (∃C. tag obj = CInst C) ∨ (∃T k. tag obj = Arr T k)

lemma obj_class_CInst:

  obj_class (| tag = CInst C, values = vs |) = C

lemma obj_class_CInst1:

  tag obj = CInst C ==> obj_class obj = C

lemma obj_class_Arr:

  obj_class (| tag = Arr T k, values = vs |) = Object

lemma obj_class_Arr1:

  tag obj = Arr T k ==> obj_class obj = Object

lemma obj_ty_obj_class:

  G\<turnstile>obj_ty obj\<preceq>Class statC = G|-obj_class obj<=:C statC

object references

lemma fields_table_SomeI:

  [| table_of (DeclConcepts.fields G C) n = Some f; P n f |]
  ==> fields_table G C P n = Some (type f)

lemma fields_table_SomeD':

  fields_table G C P fn = Some T
  ==> ∃f. (fn, f) ∈ set (DeclConcepts.fields G C) ∧ type f = T

lemma fields_table_SomeD:

  [| fields_table G C P fn = Some T; unique (DeclConcepts.fields G C) |]
  ==> ∃f. table_of (DeclConcepts.fields G C) fn = Some f ∧ type f = T

lemma var_tys_Some_eq:

  (var_tys G oi r n = Some T) =
  (case r of
   Inl a =>
     case oi of
     CInst C => ∃nt. n = Inl nt ∧ fields_table G Cn f. ¬ static f) nt = Some T
     | Arr t k => ∃i. n = Inr ii in_bounds kt = T
   | Inr C =>
       ∃nt. n = Inl nt ∧
            fields_table G Cfn f. declclassf fn = C ∧ static f) nt = Some T)

stores

access

lemma globs_def2:

  globs (st g l) = g

lemma locals_def2:

  locals (st g l) = l

lemma heap_def2:

  heap s a = globs s (Inl a)

memory allocation

lemma new_AddrD:

  new_Addr h = Some a ==> h a = None

lemma new_AddrD2:

  new_Addr h = Some a ==> ∀b. h b  None --> b  a

lemma new_Addr_SomeI:

  h a = None ==> ∃b. new_Addr h = Some bh b = None

initialization

lemma init_arr_comps_base:

  init_vals (arr_comps T 0) = empty

lemma init_arr_comps_step:

  0 < j
  ==> init_vals (arr_comps T j) = init_vals (arr_comps T (j - 1))(j - 1 |->
      default_val T)

update

lemma gupd_def2:

  gupd(r\<mapsto>obj) (st g l) = st (g(r |-> obj)) l

lemma lupd_def2:

  lupd(vn\<mapsto>v) (st g l) = st g (l(vn |-> v))

lemma globs_gupd:

  globs (gupd(r\<mapsto>obj) s) = globs s(r |-> obj)

lemma globs_lupd:

  globs (lupd(vn\<mapsto>v) s) = globs s

lemma locals_gupd:

  locals (gupd(r\<mapsto>obj) s) = locals s

lemma locals_lupd:

  locals (lupd(vn\<mapsto>v) s) = locals s(vn |-> v)

lemma globs_upd_gobj_new:

  globs s r = None ==> globs (upd_gobj r n v s) = globs s

lemma globs_upd_gobj_upd:

  globs s r = Some obj
  ==> globs (upd_gobj r n v s) = globs s(r |-> upd_obj n v obj)

lemma locals_upd_gobj:

  locals (upd_gobj r n v s) = locals s

lemma globs_init_obj:

  globs (init_obj G oi r s) t =
  (if t = r then Some (| tag = oi, values = init_vals (var_tys G oi r) |)
   else globs s t)

lemma locals_init_obj:

  locals (init_obj G oi r s) = locals s

lemma surjective_st:

  st (globs s) (locals s) = s

lemma surjective_st_init_obj:

  st (globs (init_obj G oi r s)) (locals s) = init_obj G oi r s

lemma heap_heap_upd:

  heap (st (g(Inl a |-> obj)) l) = heap (st g l)(a |-> obj)

lemma heap_stat_upd:

  heap (st (g(Inr C |-> obj)) l) = heap (st g l)

lemma heap_local_upd:

  heap (st g (l(vn |-> v))) = heap (st g l)

lemma heap_gupd_Heap:

  heap (gupd(Inl a\<mapsto>obj) s) = heap s(a |-> obj)

lemma heap_gupd_Stat:

  heap (gupd(Inr C\<mapsto>obj) s) = heap s

lemma heap_lupd:

  heap (lupd(vn\<mapsto>v) s) = heap s

lemma heap_upd_gobj_Stat:

  heap (upd_gobj (Inr C) n v s) = heap s

lemma set_locals_def2:

  set_locals l (st g l') = st g l

lemma set_locals_id:

  set_locals (locals s) s = s

lemma set_set_locals:

  set_locals l (set_locals l' s) = set_locals l s

lemma locals_set_locals:

  locals (set_locals l s) = l

lemma globs_set_locals:

  globs (set_locals l s) = globs s

lemma heap_set_locals:

  heap (set_locals l s) = heap s

abrupt completion

lemma abrupt_if_True_None:

  abrupt_if True x None = x

lemma abrupt_if_True_not_None:

  x  None ==> abrupt_if True x y  None

lemma abrupt_if_False:

  abrupt_if False x y = y

lemma abrupt_if_Some:

  abrupt_if c x (Some y) = Some y

lemma abrupt_if_not_None:

  y  None ==> abrupt_if c x y = y

lemma split_abrupt_if:

  P (abrupt_if c x' x) = ((cx = None --> P x') ∧ (¬ (cx = None) --> P x))

lemma raise_if_None:

  ((raise_if c x) y = None) = (¬ cy = None)

lemma if_raise_if_None:

  ((if b then y else (raise_if c x) y) = None) = ((c --> b) ∧ y = None)

lemma raise_if_SomeD:

  (raise_if c x) y = Some z ==> cz = Xcpt (Std x) ∧ y = None ∨ y = Some z

lemma error_if_None:

  ((error_if c e) y = None) = (¬ cy = None)

lemma if_error_if_None:

  ((if b then y else (error_if c e) y) = None) = ((c --> b) ∧ y = None)

lemma error_if_SomeD:

  (error_if c e) y = Some z ==> cz = Error ey = None ∨ y = Some z

lemma absorb_SomeD:

  absorb j a = Some x ==> a = Some x

lemma absorb_same:

  absorb j (Some (Jump j)) = None

lemma absorb_other:

  a  Some (Jump j) ==> absorb j a = a

lemma absorb_Some_NoneD:

  absorb j (Some abr) = None ==> abr = Jump j

lemma absorb_Some_JumpD:

  absorb j s = Some (Jump j') ==> j'  j

full program state

lemma single_stateE:

  Z. Z = s ==> False

lemma state_not_single:

  All (op = x) ==> R

lemma normal_def2:

  normal s = (fst s = None)

lemma heap_free_def2:

  heap_free n s = atleast_free (heap (snd s)) n

update

lemma abupd_def2:

  abupd f (x, s) = (f x, s)

lemma abupd_abrupt_if_False:

  abupd (abrupt_if False xo) s = s

lemma supd_def2:

  supd f (x, s) = (x, f s)

lemma supd_lupd:

  supd lupd(vn\<mapsto>v) s = (fst s, lupd(vn\<mapsto>v) (snd s))

lemma supd_gupd:

  supd gupd(r\<mapsto>obj) s = (fst s, gupd(r\<mapsto>obj) (snd s))

lemma supd_init_obj:

  supd (init_obj G oi r) s = (fst s, init_obj G oi r (snd s))

lemma abupd_store_invariant:

  snd (abupd f s) = snd s

lemma supd_abrupt_invariant:

  fst (supd f s) = fst s

lemma set_set_lvars:

  (set_lvars l) ((set_lvars l') s) = (set_lvars l) s

lemma set_lvars_id:

  (set_lvars (locals (snd s))) s = s

initialisation test

lemma not_inited_empty:

  ¬ inited C empty

lemma inited_gupdate:

  inited C (g(r |-> obj)) = (inited C gr = Inr C)

lemma inited_init_class_obj:

  inited C (globs ((init_class_obj G C) s))

lemma not_initedD:

  ¬ inited C g ==> g (Inr C) = None

lemma initedD:

  inited C g ==> ∃obj. g (Inr C) = Some obj

lemma initd_def2:

  initd C s = inited C (globs (snd s))

@{text error_free}

lemma error_free_Norm:

  error_free (Norm s)

lemma error_free_normal:

  normal s ==> error_free s

lemma error_free_Xcpt:

  error_free (Some (Xcpt x), s)

lemma error_free_Jump:

  error_free (Some (Jump j), s)

lemma error_free_Error:

  error_free (Some (Error e), s) = False

lemma error_free_Some:

  ¬ (∃err. x = Error err) ==> error_free (Some x, s)

lemma error_free_abupd_absorb:

  error_free s ==> error_free (abupd (absorb j) s)

lemma error_free_absorb:

  error_free (a, s) ==> error_free (absorb j a, s)

lemma error_free_abrupt_if:

  [| error_free s; ¬ (∃err. x = Error err) |]
  ==> error_free (abupd (abrupt_if p (Some x)) s)

lemma error_free_abrupt_if1:

  [| error_free (a, s); ¬ (∃err. x = Error err) |]
  ==> error_free (abrupt_if p (Some x) a, s)

lemma error_free_abrupt_if_Xcpt:

  error_free s ==> error_free (abupd (abrupt_if p (Some (Xcpt x))) s)

lemma error_free_abrupt_if_Xcpt1:

  error_free (a, s) ==> error_free (abrupt_if p (Some (Xcpt x)) a, s)

lemma error_free_abrupt_if_Jump:

  error_free s ==> error_free (abupd (abrupt_if p (Some (Jump j))) s)

lemma error_free_abrupt_if_Jump1:

  error_free (a, s) ==> error_free (abrupt_if p (Some (Jump j)) a, s)

lemma error_free_raise_if:

  error_free s ==> error_free (abupd (raise_if p x) s)

lemma error_free_raise_if1:

  error_free (a, s) ==> error_free ((raise_if p x) a, s)

lemma error_free_supd:

  error_free s ==> error_free (supd f s)

lemma error_free_supd1:

  error_free (a, s) ==> error_free (a, f s)

lemma error_free_set_lvars:

  error_free s ==> error_free ((set_lvars l) s)

lemma error_free_set_locals:

  error_free (x, s) ==> error_free (x, set_locals l s')