Theory BVConform

Up to index of Isabelle/HOL/Jinja

theory BVConform = BVSpec + JVMExec + Conform:

(*  Title:      HOL/MicroJava/BV/Correct.thy
    ID:         $Id: BVConform.html 1910 2004-05-19 04:46:04Z kleing $
    Author:     Cornelia Pusch, Gerwin Klein
    Copyright   1999 Technische Universitaet Muenchen

The invariant for the type safety proof.
*)

header {* \isaheader{BV Type Safety Invariant} *}

theory BVConform = BVSpec + JVMExec + Conform:


syntax (xsymbols)
  confT :: "'c prog => heap => val => ty err => bool" 
           ("_,_ \<turnstile> _ :≤\<top> _" [51,51,51,51] 50)

constdefs
  confT :: "'c prog => heap => val => ty err => bool" 
           ("_,_ |- _ :<=T _" [51,51,51,51] 50)
  "P,h \<turnstile> v :≤\<top> E ≡ case E of Err => True | OK T => P,h \<turnstile> v :≤ T"

syntax
  confTs :: "'c prog => heap => val list => tyl => bool" 
            ("_,_ |- _ [:<=T] _" [51,51,51,51] 50)

syntax (xsymbols)
  confTs :: "'c prog => heap => val list => tyl => bool"
           ("_,_ \<turnstile> _ [:≤\<top>] _" [51,51,51,51] 50)

translations
  "P,h \<turnstile> vs [:≤\<top>] Ts" == "list_all2 (confT P h) vs Ts"

constdefs
  conf_f  :: "jvm_prog => heap => tyi => bytecode => frame => bool"
  "conf_f P h ≡ λ(ST,LT) is (stk,loc,C,M,pc).
  P,h \<turnstile> stk [:≤] ST ∧ P,h \<turnstile> loc [:≤\<top>] LT ∧ pc < size is"

lemma conf_f_def2:
  "conf_f P h (ST,LT) is (stk,loc,C,M,pc) ≡
  P,h \<turnstile> stk [:≤] ST ∧ P,h \<turnstile> loc [:≤\<top>] LT ∧ pc < size is"
  by (simp add: conf_f_def)


consts
 conf_fs :: "[jvm_prog,heap,tyP,mname,nat,ty,frame list] => bool"
primrec
"conf_fs P h Φ M0 n0 T0 [] = True"

"conf_fs P h Φ M0 n0 T0 (f#frs) =
  (let (stk,loc,C,M,pc) = f in
  (∃ST LT Ts T mxs mxl0 is xt.
    Φ C M ! pc = Some (ST,LT) ∧ 
    (P \<turnstile> C sees M:Ts -> T = (mxs,mxl0,is,xt) in C) ∧
    (∃D Ts' T' m D'.  
       is!pc = (Invoke M0 n0) ∧ ST!n0 = Class D ∧
       P \<turnstile> D sees M0:Ts' -> T' = m in D' ∧ P \<turnstile> T0 ≤ T') ∧
    conf_f P h (ST, LT) is f ∧ conf_fs P h Φ M (size Ts) T frs))"


constdefs
 correct_state :: "[jvm_prog,tyP,jvm_state] => bool"
                  ("_,_ |- _ [ok]"  [61,0,0] 61)
"correct_state P Φ ≡ λ(xp,h,frs).
  case xp of
     None => (case frs of
             [] => True
             | (f#fs) => P\<turnstile> h\<surd> ∧ 
             (let (stk,loc,C,M,pc) = f
              in ∃Ts T mxs mxl0 is xt τ.
                    (P \<turnstile> C sees M:Ts->T = (mxs,mxl0,is,xt) in C) ∧
                    Φ C M ! pc = Some τ ∧
                    conf_f P h τ is f ∧ conf_fs P h Φ M (size Ts) T fs))
  | Some x => frs = []" 

syntax (xsymbols)
 correct_state :: "[jvm_prog,tyP,jvm_state] => bool"
                  ("_,_ \<turnstile> _ \<surd>"  [61,0,0] 61)



section {* Values and @{text "\<top>"} *}

lemma confT_Err [iff]: "P,h \<turnstile> x :≤\<top> Err" 
  by (simp add: confT_def)

lemma confT_OK [iff]:  "P,h \<turnstile> x :≤\<top> OK T = (P,h \<turnstile> x :≤ T)"
  by (simp add: confT_def)

lemma confT_cases:
  "P,h \<turnstile> x :≤\<top> X = (X = Err ∨ (∃T. X = OK T ∧ P,h \<turnstile> x :≤ T))"
  by (cases X) auto

lemma confT_hext [intro?, trans]:
  "[| P,h \<turnstile> x :≤\<top> T; h \<unlhd> h' |] ==> P,h' \<turnstile> x :≤\<top> T"
  by (cases T) (blast intro: conf_hext)+

lemma confT_widen [intro?, trans]:
  "[| P,h \<turnstile> x :≤\<top> T; P \<turnstile> T ≤\<top> T' |] ==> P,h \<turnstile> x :≤\<top> T'"
  by (cases T', auto intro: conf_widen)


section {* Stack and Registers *}

lemmas confTs_Cons1 [iff] = list_all2_Cons1 [of "confT P h", standard]

lemma confTs_confT_sup:
  "[| P,h \<turnstile> loc [:≤\<top>] LT; n < size LT; LT!n = OK T; P \<turnstile> T ≤ T' |] 
  ==> P,h \<turnstile> (loc!n) :≤ T'"
(*<*)
  apply (frule list_all2_lengthD)
  apply (drule list_all2_nthD, simp)
  apply simp
  apply (erule conf_widen, assumption+)
  done
(*>*)

lemma confTs_hext [intro?]:
  "P,h \<turnstile> loc [:≤\<top>] LT ==> h \<unlhd> h' ==> P,h' \<turnstile> loc [:≤\<top>] LT"
  by (fast elim: list_all2_mono confT_hext)    

lemma confTs_widen [intro?, trans]:
  "P,h \<turnstile> loc [:≤\<top>] LT ==> P \<turnstile> LT [≤\<top>] LT' ==> P,h \<turnstile> loc [:≤\<top>] LT'"
  by (rule list_all2_trans, rule confT_widen)

lemma confTs_map [iff]:
  "!!vs. (P,h \<turnstile> vs [:≤\<top>] map OK Ts) = (P,h \<turnstile> vs [:≤] Ts)"
  by (induct Ts) (auto simp add: list_all2_Cons2)

lemma reg_widen_Err [iff]:
  "!!LT. (P \<turnstile> replicate n Err [≤\<top>] LT) = (LT = replicate n Err)"
  by (induct n) (auto simp add: list_all2_Cons1)
    
lemma confTs_Err [iff]:
  "P,h \<turnstile> replicate n v [:≤\<top>] replicate n Err"
  by (induct n) auto

  
section {* correct-frames *}

lemmas [simp del] = fun_upd_apply

lemma conf_fs_hext:
  "!!M n Tr. 
  [| conf_fs P h Φ M n Tr frs; h \<unlhd> h' |] ==> conf_fs P h' Φ M n Tr frs"
(*<*)
apply (induct frs)
 apply simp
apply clarify
apply (simp (no_asm_use))
apply clarify
apply (unfold conf_f_def)
apply (simp (no_asm_use))
apply clarify
apply (fast elim!: confs_hext confTs_hext)
done
(*>*)

end

lemma conf_f_def2:

  conf_f P h (ST, LT) is (stk, loc, C, M, pc) ==
  P,h \<turnstile> stk [:≤] ST & P,h |- loc [:<=T] LT & pc < length is

Values and @{text "\<top>"}

lemma confT_Err:

  P,h |- x :<=T Err

lemma confT_OK:

  (P,h |- x :<=T OK T) = (P,h \<turnstile> x :≤ T)

lemma confT_cases:

  (P,h |- x :<=T X) = (X = Err | (EX T. X = OK T & P,h \<turnstile> x :≤ T))

lemma confT_hext:

  [| P,h |- x :<=T T; h \<unlhd> h' |] ==> P,h' |- x :<=T T

lemma confT_widen:

  [| P,h |- x :<=T T; P |- T <=T T' |] ==> P,h |- x :<=T T'

Stack and Registers

lemmas confTs_Cons1:

  (P,h |- x # xs [:<=T] ys) =
  (EX z zs. ys = z # zs & P,h |- x :<=T z & P,h |- xs [:<=T] zs)

lemma confTs_confT_sup:

  [| P,h |- loc [:<=T] LT; n < length LT; LT ! n = OK T; widen P T T' |]
  ==> P,h \<turnstile> loc ! n :≤ T'

lemma confTs_hext:

  [| P,h |- loc [:<=T] LT; h \<unlhd> h' |] ==> P,h' |- loc [:<=T] LT

lemma confTs_widen:

  [| P,h |- loc [:<=T] LT; P |- LT [<=T] LT' |] ==> P,h |- loc [:<=T] LT'

lemma confTs_map:

  (P,h |- vs [:<=T] map OK Ts) = (P,h \<turnstile> vs [:≤] Ts)

lemma reg_widen_Err:

  P |- replicate n Err [<=T] LT = (LT = replicate n Err)

lemma confTs_Err:

  P,h |- replicate n v [:<=T] replicate n Err

correct-frames

lemmas

  (f(x := y)) z = (if z = x then y else f z)

lemma conf_fs_hext:

  [| conf_fs P h Φ M n Tr frs; h \<unlhd> h' |] ==> conf_fs P h' Φ M n Tr frs