(* Title: HOL/Bali/AxSem.thy ID: $Id: AxSem.thy,v 1.20 2007/09/30 19:55:15 wenzelm Exp $ Author: David von Oheimb *) header {* Axiomatic semantics of Java expressions and statements (see also Eval.thy) *} theory AxSem imports Evaln TypeSafe begin text {* design issues: \begin{itemize} \item a strong version of validity for triples with premises, namely one that takes the recursive depth needed to complete execution, enables correctness proof \item auxiliary variables are handled first-class (-> Thomas Kleymann) \item expressions not flattened to elementary assignments (as usual for axiomatic semantics) but treated first-class => explicit result value handling \item intermediate values not on triple, but on assertion level (with result entry) \item multiple results with semantical substitution mechnism not requiring a stack \item because of dynamic method binding, terms need to be dependent on state. this is also useful for conditional expressions and statements \item result values in triples exactly as in eval relation (also for xcpt states) \item validity: additional assumption of state conformance and well-typedness, which is required for soundness and thus rule hazard required of completeness \end{itemize} restrictions: \begin{itemize} \item all triples in a derivation are of the same type (due to weak polymorphism) \end{itemize} *} types res = vals --{* result entry *} syntax Val :: "val => res" Var :: "var => res" Vals :: "val list => res" translations "Val x" => "(In1 x)" "Var x" => "(In2 x)" "Vals x" => "(In3 x)" syntax "_Val" :: "[pttrn] => pttrn" ("Val:_" [951] 950) "_Var" :: "[pttrn] => pttrn" ("Var:_" [951] 950) "_Vals" :: "[pttrn] => pttrn" ("Vals:_" [951] 950) translations "λVal:v . b" == "(λv. b) o the_In1" "λVar:v . b" == "(λv. b) o the_In2" "λVals:v. b" == "(λv. b) o the_In3" --{* relation on result values, state and auxiliary variables *} types 'a assn = "res => state => 'a => bool" translations "res" <= (type) "AxSem.res" "a assn" <= (type) "vals => state => a => bool" constdefs assn_imp :: "'a assn => 'a assn => bool" (infixr "=>" 25) "P => Q ≡ ∀Y s Z. P Y s Z --> Q Y s Z" lemma assn_imp_def2 [iff]: "(P => Q) = (∀Y s Z. P Y s Z --> Q Y s Z)" apply (unfold assn_imp_def) apply (rule HOL.refl) done section "assertion transformers" subsection "peek-and" constdefs peek_and :: "'a assn => (state => bool) => 'a assn" (infixl "∧." 13) "P ∧. p ≡ λY s Z. P Y s Z ∧ p s" lemma peek_and_def2 [simp]: "peek_and P p Y s = (λZ. (P Y s Z ∧ p s))" apply (unfold peek_and_def) apply (simp (no_asm)) done lemma peek_and_Not [simp]: "(P ∧. (λs. ¬ f s)) = (P ∧. Not o f)" apply (rule ext) apply (rule ext) apply (simp (no_asm)) done lemma peek_and_and [simp]: "peek_and (peek_and P p) p = peek_and P p" apply (unfold peek_and_def) apply (simp (no_asm)) done lemma peek_and_commut: "(P ∧. p ∧. q) = (P ∧. q ∧. p)" apply (rule ext) apply (rule ext) apply (rule ext) apply auto done syntax Normal :: "'a assn => 'a assn" translations "Normal P" == "P ∧. normal" lemma peek_and_Normal [simp]: "peek_and (Normal P) p = Normal (peek_and P p)" apply (rule ext) apply (rule ext) apply (rule ext) apply auto done subsection "assn-supd" constdefs assn_supd :: "'a assn => (state => state) => 'a assn" (infixl ";." 13) "P ;. f ≡ λY s' Z. ∃s. P Y s Z ∧ s' = f s" lemma assn_supd_def2 [simp]: "assn_supd P f Y s' Z = (∃s. P Y s Z ∧ s' = f s)" apply (unfold assn_supd_def) apply (simp (no_asm)) done subsection "supd-assn" constdefs supd_assn :: "(state => state) => 'a assn => 'a assn" (infixr ".;" 13) "f .; P ≡ λY s. P Y (f s)" lemma supd_assn_def2 [simp]: "(f .; P) Y s = P Y (f s)" apply (unfold supd_assn_def) apply (simp (no_asm)) done lemma supd_assn_supdD [elim]: "((f .; Q) ;. f) Y s Z ==> Q Y s Z" apply auto done lemma supd_assn_supdI [elim]: "Q Y s Z ==> (f .; (Q ;. f)) Y s Z" apply (auto simp del: split_paired_Ex) done subsection "subst-res" constdefs subst_res :: "'a assn => res => 'a assn" ("_\<leftarrow>_" [60,61] 60) "P\<leftarrow>w ≡ λY. P w" lemma subst_res_def2 [simp]: "(P\<leftarrow>w) Y = P w" apply (unfold subst_res_def) apply (simp (no_asm)) done lemma subst_subst_res [simp]: "P\<leftarrow>w\<leftarrow>v = P\<leftarrow>w" apply (rule ext) apply (simp (no_asm)) done lemma peek_and_subst_res [simp]: "(P ∧. p)\<leftarrow>w = (P\<leftarrow>w ∧. p)" apply (rule ext) apply (rule ext) apply (simp (no_asm)) done (*###Do not work for some strange (unification?) reason lemma subst_res_Val_beta [simp]: "(λY. P (the_In1 Y))\<leftarrow>Val v = (λY. P v)" apply (rule ext) by simp lemma subst_res_Var_beta [simp]: "(λY. P (the_In2 Y))\<leftarrow>Var vf = (λY. P vf)"; apply (rule ext) by simp lemma subst_res_Vals_beta [simp]: "(λY. P (the_In3 Y))\<leftarrow>Vals vs = (λY. P vs)"; apply (rule ext) by simp *) subsection "subst-Bool" constdefs subst_Bool :: "'a assn => bool => 'a assn" ("_\<leftarrow>=_" [60,61] 60) "P\<leftarrow>=b ≡ λY s Z. ∃v. P (Val v) s Z ∧ (normal s --> the_Bool v=b)" lemma subst_Bool_def2 [simp]: "(P\<leftarrow>=b) Y s Z = (∃v. P (Val v) s Z ∧ (normal s --> the_Bool v=b))" apply (unfold subst_Bool_def) apply (simp (no_asm)) done lemma subst_Bool_the_BoolI: "P (Val b) s Z ==> (P\<leftarrow>=the_Bool b) Y s Z" apply auto done subsection "peek-res" constdefs peek_res :: "(res => 'a assn) => 'a assn" "peek_res Pf ≡ λY. Pf Y Y" syntax "@peek_res" :: "pttrn => 'a assn => 'a assn" ("λ_:. _" [0,3] 3) translations "λw:. P" == "peek_res (λw. P)" lemma peek_res_def2 [simp]: "peek_res P Y = P Y Y" apply (unfold peek_res_def) apply (simp (no_asm)) done lemma peek_res_subst_res [simp]: "peek_res P\<leftarrow>w = P w\<leftarrow>w" apply (rule ext) apply (simp (no_asm)) done (* unused *) lemma peek_subst_res_allI: "(!!a. T a (P (f a)\<leftarrow>f a)) ==> ∀a. T a (peek_res P\<leftarrow>f a)" apply (rule allI) apply (simp (no_asm)) apply fast done subsection "ign-res" constdefs ign_res :: " 'a assn => 'a assn" ("_\<down>" [1000] 1000) "P\<down> ≡ λY s Z. ∃Y. P Y s Z" lemma ign_res_def2 [simp]: "P\<down> Y s Z = (∃Y. P Y s Z)" apply (unfold ign_res_def) apply (simp (no_asm)) done lemma ign_ign_res [simp]: "P\<down>\<down> = P\<down>" apply (rule ext) apply (rule ext) apply (rule ext) apply (simp (no_asm)) done lemma ign_subst_res [simp]: "P\<down>\<leftarrow>w = P\<down>" apply (rule ext) apply (rule ext) apply (rule ext) apply (simp (no_asm)) done lemma peek_and_ign_res [simp]: "(P ∧. p)\<down> = (P\<down> ∧. p)" apply (rule ext) apply (rule ext) apply (rule ext) apply (simp (no_asm)) done subsection "peek-st" constdefs peek_st :: "(st => 'a assn) => 'a assn" "peek_st P ≡ λY s. P (store s) Y s" syntax "@peek_st" :: "pttrn => 'a assn => 'a assn" ("λ_.. _" [0,3] 3) translations "λs.. P" == "peek_st (λs. P)" lemma peek_st_def2 [simp]: "(λs.. Pf s) Y s = Pf (store s) Y s" apply (unfold peek_st_def) apply (simp (no_asm)) done lemma peek_st_triv [simp]: "(λs.. P) = P" apply (rule ext) apply (rule ext) apply (simp (no_asm)) done lemma peek_st_st [simp]: "(λs.. λs'.. P s s') = (λs.. P s s)" apply (rule ext) apply (rule ext) apply (simp (no_asm)) done lemma peek_st_split [simp]: "(λs.. λY s'. P s Y s') = (λY s. P (store s) Y s)" apply (rule ext) apply (rule ext) apply (simp (no_asm)) done lemma peek_st_subst_res [simp]: "(λs.. P s)\<leftarrow>w = (λs.. P s\<leftarrow>w)" apply (rule ext) apply (simp (no_asm)) done lemma peek_st_Normal [simp]: "(λs..(Normal (P s))) = Normal (λs.. P s)" apply (rule ext) apply (rule ext) apply (simp (no_asm)) done subsection "ign-res-eq" constdefs ign_res_eq :: "'a assn => res => 'a assn" ("_\<down>=_" [60,61] 60) "P\<down>=w ≡ λY:. P\<down> ∧. (λs. Y=w)" lemma ign_res_eq_def2 [simp]: "(P\<down>=w) Y s Z = ((∃Y. P Y s Z) ∧ Y=w)" apply (unfold ign_res_eq_def) apply auto done lemma ign_ign_res_eq [simp]: "(P\<down>=w)\<down> = P\<down>" apply (rule ext) apply (rule ext) apply (rule ext) apply (simp (no_asm)) done (* unused *) lemma ign_res_eq_subst_res: "P\<down>=w\<leftarrow>w = P\<down>" apply (rule ext) apply (rule ext) apply (rule ext) apply (simp (no_asm)) done (* unused *) lemma subst_Bool_ign_res_eq: "((P\<leftarrow>=b)\<down>=x) Y s Z = ((P\<leftarrow>=b) Y s Z ∧ Y=x)" apply (simp (no_asm)) done subsection "RefVar" constdefs RefVar :: "(state => vvar × state) => 'a assn => 'a assn"(infixr "..;" 13) "vf ..; P ≡ λY s. let (v,s') = vf s in P (Var v) s'" lemma RefVar_def2 [simp]: "(vf ..; P) Y s = P (Var (fst (vf s))) (snd (vf s))" apply (unfold RefVar_def Let_def) apply (simp (no_asm) add: split_beta) done subsection "allocation" constdefs Alloc :: "prog => obj_tag => 'a assn => 'a assn" "Alloc G otag P ≡ λY s Z. ∀s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a-> s'--> P (Val (Addr a)) s' Z" SXAlloc :: "prog => 'a assn => 'a assn" "SXAlloc G P ≡ λY s Z. ∀s'. G\<turnstile>s \<midarrow>sxalloc-> s' --> P Y s' Z" lemma Alloc_def2 [simp]: "Alloc G otag P Y s Z = (∀s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a-> s'--> P (Val (Addr a)) s' Z)" apply (unfold Alloc_def) apply (simp (no_asm)) done lemma SXAlloc_def2 [simp]: "SXAlloc G P Y s Z = (∀s'. G\<turnstile>s \<midarrow>sxalloc-> s' --> P Y s' Z)" apply (unfold SXAlloc_def) apply (simp (no_asm)) done section "validity" constdefs type_ok :: "prog => term => state => bool" "type_ok G t s ≡ ∃L T C A. (normal s --> (|prg=G,cls=C,lcl=L|)),\<turnstile>t::T ∧ (|prg=G,cls=C,lcl=L|)),\<turnstile>dom (locals (store s))»t»A ) ∧ s::\<preceq>(G,L)" datatype 'a triple = triple "('a assn)" "term" "('a assn)" (** should be something like triple = ∀'a. triple ('a assn) term ('a assn) **) ("{(1_)}/ _>/ {(1_)}" [3,65,3]75) types 'a triples = "'a triple set" syntax var_triple :: "['a assn, var ,'a assn] => 'a triple" ("{(1_)}/ _=>/ {(1_)}" [3,80,3] 75) expr_triple :: "['a assn, expr ,'a assn] => 'a triple" ("{(1_)}/ _->/ {(1_)}" [3,80,3] 75) exprs_triple :: "['a assn, expr list ,'a assn] => 'a triple" ("{(1_)}/ _#>/ {(1_)}" [3,65,3] 75) stmt_triple :: "['a assn, stmt, 'a assn] => 'a triple" ("{(1_)}/ ._./ {(1_)}" [3,65,3] 75) syntax (xsymbols) triple :: "['a assn, term ,'a assn] => 'a triple" ("{(1_)}/ _\<succ>/ {(1_)}" [3,65,3] 75) var_triple :: "['a assn, var ,'a assn] => 'a triple" ("{(1_)}/ _=\<succ>/ {(1_)}" [3,80,3] 75) expr_triple :: "['a assn, expr ,'a assn] => 'a triple" ("{(1_)}/ _-\<succ>/ {(1_)}" [3,80,3] 75) exprs_triple :: "['a assn, expr list ,'a assn] => 'a triple" ("{(1_)}/ _\<doteq>\<succ>/ {(1_)}" [3,65,3] 75) translations "{P} e-\<succ> {Q}" == "{P} In1l e\<succ> {Q}" "{P} e=\<succ> {Q}" == "{P} In2 e\<succ> {Q}" "{P} e\<doteq>\<succ> {Q}" == "{P} In3 e\<succ> {Q}" "{P} .c. {Q}" == "{P} In1r c\<succ> {Q}" lemma inj_triple: "inj (λ(P,t,Q). {P} t\<succ> {Q})" apply (rule inj_onI) apply auto done lemma triple_inj_eq: "({P} t\<succ> {Q} = {P'} t'\<succ> {Q'} ) = (P=P' ∧ t=t' ∧ Q=Q')" apply auto done constdefs mtriples :: "('c => 'sig => 'a assn) => ('c => 'sig => expr) => ('c => 'sig => 'a assn) => ('c × 'sig) set => 'a triples" ("{{(1_)}/ _-\<succ>/ {(1_)} | _}"[3,65,3,65]75) "{{P} tf-\<succ> {Q} | ms} ≡ (λ(C,sig). {Normal(P C sig)} tf C sig-\<succ> {Q C sig})`ms" consts triple_valid :: "prog => nat => 'a triple => bool" ( "_\<Turnstile>_:_" [61,0, 58] 57) ax_valids :: "prog => 'b triples => 'a triples => bool" ("_,_|\<Turnstile>_" [61,58,58] 57) syntax triples_valid:: "prog => nat => 'a triples => bool" ( "_||=_:_" [61,0, 58] 57) ax_valid :: "prog => 'b triples => 'a triple => bool" ( "_,_|=_" [61,58,58] 57) syntax (xsymbols) triples_valid:: "prog => nat => 'a triples => bool" ( "_|\<Turnstile>_:_" [61,0, 58] 57) ax_valid :: "prog => 'b triples => 'a triple => bool" ( "_,_\<Turnstile>_" [61,58,58] 57) defs triple_valid_def: "G\<Turnstile>n:t ≡ case t of {P} t\<succ> {Q} => ∀Y s Z. P Y s Z --> type_ok G t s --> (∀Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n-> (Y',s') --> Q Y' s' Z)" translations "G|\<Turnstile>n:ts" == "Ball ts (triple_valid G n)" defs ax_valids_def:"G,A|\<Turnstile>ts ≡ ∀n. G|\<Turnstile>n:A --> G|\<Turnstile>n:ts" translations "G,A \<Turnstile>t" == "G,A|\<Turnstile>{t}" lemma triple_valid_def2: "G\<Turnstile>n:{P} t\<succ> {Q} = (∀Y s Z. P Y s Z --> (∃L. (normal s --> (∃ C T A. (|prg=G,cls=C,lcl=L|)),\<turnstile>t::T ∧ (|prg=G,cls=C,lcl=L|)),\<turnstile>dom (locals (store s))»t»A)) ∧ s::\<preceq>(G,L)) --> (∀Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n-> (Y',s')--> Q Y' s' Z))" apply (unfold triple_valid_def type_ok_def) apply (simp (no_asm)) done declare split_paired_All [simp del] split_paired_Ex [simp del] declare split_if [split del] split_if_asm [split del] option.split [split del] option.split_asm [split del] declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} declaration {* K (Classical.map_cs (fn cs => cs delSWrapper "split_all_tac")) *} inductive ax_derivs :: "prog => 'a triples => 'a triples => bool" ("_,_|\<turnstile>_" [61,58,58] 57) and ax_deriv :: "prog => 'a triples => 'a triple => bool" ("_,_\<turnstile>_" [61,58,58] 57) for G :: prog where "G,A \<turnstile>t ≡ G,A|\<turnstile>{t}" | empty: " G,A|\<turnstile>{}" | insert:"[|G,A\<turnstile>t; G,A|\<turnstile>ts|] ==> G,A|\<turnstile>insert t ts" | asm: "ts⊆A ==> G,A|\<turnstile>ts" (* could be added for convenience and efficiency, but is not necessary cut: "[|G,A'|\<turnstile>ts; G,A|\<turnstile>A'|] ==> G,A |\<turnstile>ts" *) | weaken:"[|G,A|\<turnstile>ts'; ts ⊆ ts'|] ==> G,A|\<turnstile>ts" | conseq:"∀Y s Z . P Y s Z --> (∃P' Q'. G,A\<turnstile>{P'} t\<succ> {Q'} ∧ (∀Y' s'. (∀Y Z'. P' Y s Z' --> Q' Y' s' Z') --> Q Y' s' Z )) ==> G,A\<turnstile>{P } t\<succ> {Q }" | hazard:"G,A\<turnstile>{P ∧. Not o type_ok G t} t\<succ> {Q}" | Abrupt: "G,A\<turnstile>{P\<leftarrow>(arbitrary3 t) ∧. Not o normal} t\<succ> {P}" --{* variables *} | LVar: " G,A\<turnstile>{Normal (λs.. P\<leftarrow>Var (lvar vn s))} LVar vn=\<succ> {P}" | FVar: "[|G,A\<turnstile>{Normal P} .Init C. {Q}; G,A\<turnstile>{Q} e-\<succ> {λVal:a:. fvar C stat fn a ..; R}|] ==> G,A\<turnstile>{Normal P} {accC,C,stat}e..fn=\<succ> {R}" | AVar: "[|G,A\<turnstile>{Normal P} e1-\<succ> {Q}; ∀a. G,A\<turnstile>{Q\<leftarrow>Val a} e2-\<succ> {λVal:i:. avar G i a ..; R}|] ==> G,A\<turnstile>{Normal P} e1.[e2]=\<succ> {R}" --{* expressions *} | NewC: "[|G,A\<turnstile>{Normal P} .Init C. {Alloc G (CInst C) Q}|] ==> G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}" | NewA: "[|G,A\<turnstile>{Normal P} .init_comp_ty T. {Q}; G,A\<turnstile>{Q} e-\<succ> {λVal:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}|] ==> G,A\<turnstile>{Normal P} New T[e]-\<succ> {R}" | Cast: "[|G,A\<turnstile>{Normal P} e-\<succ> {λVal:v:. λs.. abupd (raise_if (¬G,s\<turnstile>v fits T) ClassCast) .; Q\<leftarrow>Val v}|] ==> G,A\<turnstile>{Normal P} Cast T e-\<succ> {Q}" | Inst: "[|G,A\<turnstile>{Normal P} e-\<succ> {λVal:v:. λs.. Q\<leftarrow>Val (Bool (v≠Null ∧ G,s\<turnstile>v fits RefT T))}|] ==> G,A\<turnstile>{Normal P} e InstOf T-\<succ> {Q}" | Lit: "G,A\<turnstile>{Normal (P\<leftarrow>Val v)} Lit v-\<succ> {P}" | UnOp: "[|G,A\<turnstile>{Normal P} e-\<succ> {λVal:v:. Q\<leftarrow>Val (eval_unop unop v)}|] ==> G,A\<turnstile>{Normal P} UnOp unop e-\<succ> {Q}" | BinOp: "[|G,A\<turnstile>{Normal P} e1-\<succ> {Q}; ∀v1. G,A\<turnstile>{Q\<leftarrow>Val v1} (if need_second_arg binop v1 then (In1l e2) else (In1r Skip))\<succ> {λVal:v2:. R\<leftarrow>Val (eval_binop binop v1 v2)}|] ==> G,A\<turnstile>{Normal P} BinOp binop e1 e2-\<succ> {R}" | Super:" G,A\<turnstile>{Normal (λs.. P\<leftarrow>Val (val_this s))} Super-\<succ> {P}" | Acc: "[|G,A\<turnstile>{Normal P} va=\<succ> {λVar:(v,f):. Q\<leftarrow>Val v}|] ==> G,A\<turnstile>{Normal P} Acc va-\<succ> {Q}" | Ass: "[|G,A\<turnstile>{Normal P} va=\<succ> {Q}; ∀vf. G,A\<turnstile>{Q\<leftarrow>Var vf} e-\<succ> {λVal:v:. assign (snd vf) v .; R}|] ==> G,A\<turnstile>{Normal P} va:=e-\<succ> {R}" | Cond: "[|G,A \<turnstile>{Normal P} e0-\<succ> {P'}; ∀b. G,A\<turnstile>{P'\<leftarrow>=b} (if b then e1 else e2)-\<succ> {Q}|] ==> G,A\<turnstile>{Normal P} e0 ? e1 : e2-\<succ> {Q}" | Call: "[|G,A\<turnstile>{Normal P} e-\<succ> {Q}; ∀a. G,A\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {R a}; ∀a vs invC declC l. G,A\<turnstile>{(R a\<leftarrow>Vals vs ∧. (λs. declC=invocation_declclass G mode (store s) a statT (|name=mn,parTs=pTs|)), ∧ invC = invocation_class mode (store s) a statT ∧ l = locals (store s)) ;. init_lvars G declC (|name=mn,parTs=pTs|)), mode a vs) ∧. (λs. normal s --> G\<turnstile>mode->invC\<preceq>statT)} Methd declC (|name=mn,parTs=pTs|)),-\<succ> {set_lvars l .; S}|] ==> G,A\<turnstile>{Normal P} {accC,statT,mode}e·mn({pTs}args)-\<succ> {S}" | Methd:"[|G,A∪ {{P} Methd-\<succ> {Q} | ms} |\<turnstile> {{P} body G-\<succ> {Q} | ms}|] ==> G,A|\<turnstile>{{P} Methd-\<succ> {Q} | ms}" | Body: "[|G,A\<turnstile>{Normal P} .Init D. {Q}; G,A\<turnstile>{Q} .c. {λs.. abupd (absorb Ret) .; R\<leftarrow>(In1 (the (locals s Result)))}|] ==> G,A\<turnstile>{Normal P} Body D c-\<succ> {R}" --{* expression lists *} | Nil: "G,A\<turnstile>{Normal (P\<leftarrow>Vals [])} []\<doteq>\<succ> {P}" | Cons: "[|G,A\<turnstile>{Normal P} e-\<succ> {Q}; ∀v. G,A\<turnstile>{Q\<leftarrow>Val v} es\<doteq>\<succ> {λVals:vs:. R\<leftarrow>Vals (v#vs)}|] ==> G,A\<turnstile>{Normal P} e#es\<doteq>\<succ> {R}" --{* statements *} | Skip: "G,A\<turnstile>{Normal (P\<leftarrow>♦)} .Skip. {P}" | Expr: "[|G,A\<turnstile>{Normal P} e-\<succ> {Q\<leftarrow>♦}|] ==> G,A\<turnstile>{Normal P} .Expr e. {Q}" | Lab: "[|G,A\<turnstile>{Normal P} .c. {abupd (absorb l) .; Q}|] ==> G,A\<turnstile>{Normal P} .l• c. {Q}" | Comp: "[|G,A\<turnstile>{Normal P} .c1. {Q}; G,A\<turnstile>{Q} .c2. {R}|] ==> G,A\<turnstile>{Normal P} .c1;;c2. {R}" | If: "[|G,A \<turnstile>{Normal P} e-\<succ> {P'}; ∀b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c1 else c2). {Q}|] ==> G,A\<turnstile>{Normal P} .If(e) c1 Else c2. {Q}" (* unfolding variant of Loop, not needed here LoopU:"[|G,A \<turnstile>{Normal P} e-\<succ> {P'}; ∀b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c;;While(e) c else Skip).{Q}|] ==> G,A\<turnstile>{Normal P} .While(e) c. {Q}" *) | Loop: "[|G,A\<turnstile>{P} e-\<succ> {P'}; G,A\<turnstile>{Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}|] ==> G,A\<turnstile>{P} .l• While(e) c. {(P'\<leftarrow>=False)\<down>=♦}" | Jmp: "G,A\<turnstile>{Normal (abupd (λa. (Some (Jump j))) .; P\<leftarrow>♦)} .Jmp j. {P}" | Throw:"[|G,A\<turnstile>{Normal P} e-\<succ> {λVal:a:. abupd (throw a) .; Q\<leftarrow>♦}|] ==> G,A\<turnstile>{Normal P} .Throw e. {Q}" | Try: "[|G,A\<turnstile>{Normal P} .c1. {SXAlloc G Q}; G,A\<turnstile>{Q ∧. (λs. G,s\<turnstile>catch C) ;. new_xcpt_var vn} .c2. {R}; (Q ∧. (λs. ¬G,s\<turnstile>catch C)) => R|] ==> G,A\<turnstile>{Normal P} .Try c1 Catch(C vn) c2. {R}" | Fin: "[|G,A\<turnstile>{Normal P} .c1. {Q}; ∀x. G,A\<turnstile>{Q ∧. (λs. x = fst s) ;. abupd (λx. None)} .c2. {abupd (abrupt_if (x≠None) x) .; R}|] ==> G,A\<turnstile>{Normal P} .c1 Finally c2. {R}" | Done: "G,A\<turnstile>{Normal (P\<leftarrow>♦ ∧. initd C)} .Init C. {P}" | Init: "[|the (class G C) = c; G,A\<turnstile>{Normal ((P ∧. Not o initd C) ;. supd (init_class_obj G C))} .(if C = Object then Skip else Init (super c)). {Q}; ∀l. G,A\<turnstile>{Q ∧. (λs. l = locals (store s)) ;. set_lvars empty} .init c. {set_lvars l .; R}|] ==> G,A\<turnstile>{Normal (P ∧. Not o initd C)} .Init C. {R}" -- {* Some dummy rules for the intermediate terms @{text Callee}, @{text InsInitE}, @{text InsInitV}, @{text FinA} only used by the smallstep semantics. *} | InsInitV: " G,A\<turnstile>{Normal P} InsInitV c v=\<succ> {Q}" | InsInitE: " G,A\<turnstile>{Normal P} InsInitE c e-\<succ> {Q}" | Callee: " G,A\<turnstile>{Normal P} Callee l e-\<succ> {Q}" | FinA: " G,A\<turnstile>{Normal P} .FinA a c. {Q}" (* axioms *) constdefs adapt_pre :: "'a assn => 'a assn => 'a assn => 'a assn" "adapt_pre P Q Q'≡λY s Z. ∀Y' s'. ∃Z'. P Y s Z' ∧ (Q Y' s' Z' --> Q' Y' s' Z)" section "rules derived by induction" lemma cut_valid: "[|G,A'|\<Turnstile>ts; G,A|\<Turnstile>A'|] ==> G,A|\<Turnstile>ts" apply (unfold ax_valids_def) apply fast done (*if cut is available Goal "[|G,A'|\<turnstile>ts; A' ⊆ A; ∀P Q t. {P} t\<succ> {Q} ∈ A' --> (∃T. (G,L)\<turnstile>t::T) |] ==> G,A|\<turnstile>ts" b y etac ax_derivs.cut 1; b y eatac ax_derivs.asm 1 1; qed "ax_thin"; *) lemma ax_thin [rule_format (no_asm)]: "G,(A'::'a triple set)|\<turnstile>(ts::'a triple set) ==> ∀A. A' ⊆ A --> G,A|\<turnstile>ts" apply (erule ax_derivs.induct) apply (tactic "ALLGOALS(EVERY'[clarify_tac @{claset}, REPEAT o smp_tac 1])") apply (rule ax_derivs.empty) apply (erule (1) ax_derivs.insert) apply (fast intro: ax_derivs.asm) (*apply (fast intro: ax_derivs.cut) *) apply (fast intro: ax_derivs.weaken) apply (rule ax_derivs.conseq, intro strip, tactic "smp_tac 3 1",clarify, tactic "smp_tac 1 1",rule exI, rule exI, erule (1) conjI) (* 37 subgoals *) prefer 18 (* Methd *) apply (rule ax_derivs.Methd, drule spec, erule mp, fast) apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) (thms "ax_derivs.intros"))) *}) apply auto done lemma ax_thin_insert: "G,(A::'a triple set)\<turnstile>(t::'a triple) ==> G,insert x A\<turnstile>t" apply (erule ax_thin) apply fast done lemma subset_mtriples_iff: "ts ⊆ {{P} mb-\<succ> {Q} | ms} = (∃ms'. ms'⊆ms ∧ ts = {{P} mb-\<succ> {Q} | ms'})" apply (unfold mtriples_def) apply (rule subset_image_iff) done lemma weaken: "G,(A::'a triple set)|\<turnstile>(ts'::'a triple set) ==> !ts. ts ⊆ ts' --> G,A|\<turnstile>ts" apply (erule ax_derivs.induct) (*42 subgoals*) apply (tactic "ALLGOALS strip_tac") apply (tactic {* ALLGOALS(REPEAT o (EVERY'[dtac (thm "subset_singletonD"), etac disjE, fast_tac (claset() addSIs [thm "ax_derivs.empty"])]))*}) apply (tactic "TRYALL hyp_subst_tac") apply (simp, rule ax_derivs.empty) apply (drule subset_insertD) apply (blast intro: ax_derivs.insert) apply (fast intro: ax_derivs.asm) (*apply (blast intro: ax_derivs.cut) *) apply (fast intro: ax_derivs.weaken) apply (rule ax_derivs.conseq, clarify, tactic "smp_tac 3 1", blast(* unused *)) (*37 subgoals*) apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) (thms "ax_derivs.intros")) THEN_ALL_NEW fast_tac @{claset}) *}) (*1 subgoal*) apply (clarsimp simp add: subset_mtriples_iff) apply (rule ax_derivs.Methd) apply (drule spec) apply (erule impE) apply (rule exI) apply (erule conjI) apply (rule HOL.refl) oops (* dead end, Methd is to blame *) section "rules derived from conseq" text {* In the following rules we often have to give some type annotations like: @{term "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q}"}. Given only the term above without annotations, Isabelle would infer a more general type were we could have different types of auxiliary variables in the assumption set (@{term A}) and in the triple itself (@{term P} and @{term Q}). But @{text "ax_derivs.Methd"} enforces the same type in the inductive definition of the derivation. So we have to restrict the types to be able to apply the rules. *} lemma conseq12: "[|G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'}; ∀Y s Z. P Y s Z --> (∀Y' s'. (∀Y Z'. P' Y s Z' --> Q' Y' s' Z') --> Q Y' s' Z)|] ==> G,A\<turnstile>{P ::'a assn} t\<succ> {Q }" apply (rule ax_derivs.conseq) apply clarsimp apply blast done -- {* Nice variant, since it is so symmetric we might be able to memorise it. *} lemma conseq12': "[|G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'}; ∀s Y' s'. (∀Y Z. P' Y s Z --> Q' Y' s' Z) --> (∀Y Z. P Y s Z --> Q Y' s' Z)|] ==> G,A\<turnstile>{P::'a assn } t\<succ> {Q }" apply (erule conseq12) apply fast done lemma conseq12_from_conseq12': "[|G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'}; ∀Y s Z. P Y s Z --> (∀Y' s'. (∀Y Z'. P' Y s Z' --> Q' Y' s' Z') --> Q Y' s' Z)|] ==> G,A\<turnstile>{P::'a assn} t\<succ> {Q }" apply (erule conseq12') apply blast done lemma conseq1: "[|G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q}; P => P'|] ==> G,A\<turnstile>{P::'a assn} t\<succ> {Q}" apply (erule conseq12) apply blast done lemma conseq2: "[|G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q'}; Q' => Q|] ==> G,A\<turnstile>{P::'a assn} t\<succ> {Q}" apply (erule conseq12) apply blast done lemma ax_escape: "[|∀Y s Z. P Y s Z --> G,(A::'a triple set)\<turnstile>{λY' s' (Z'::'a). (Y',s') = (Y,s)} t\<succ> {λY s Z'. Q Y s Z} |] ==> G,A\<turnstile>{P::'a assn} t\<succ> {Q::'a assn}" apply (rule ax_derivs.conseq) apply force done (* unused *) lemma ax_constant: "[| C ==> G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q}|] ==> G,A\<turnstile>{λY s Z. C ∧ P Y s Z} t\<succ> {Q}" apply (rule ax_escape (* unused *)) apply clarify apply (rule conseq12) apply fast apply auto done (*alternative (more direct) proof: apply (rule ax_derivs.conseq) *)(* unused *)(* apply (fast) *) lemma ax_impossible [intro]: "G,(A::'a triple set)\<turnstile>{λY s Z. False} t\<succ> {Q::'a assn}" apply (rule ax_escape) apply clarify done (* unused *) lemma ax_nochange_lemma: "[|P Y s; All (op = w)|] ==> P w s" apply auto done lemma ax_nochange: "G,(A::(res × state) triple set)\<turnstile>{λY s Z. (Y,s)=Z} t\<succ> {λY s Z. (Y,s)=Z} ==> G,A\<turnstile>{P::(res × state) assn} t\<succ> {P}" apply (erule conseq12) apply auto apply (erule (1) ax_nochange_lemma) done (* unused *) lemma ax_trivial: "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {λY s Z. True}" apply (rule ax_derivs.conseq(* unused *)) apply auto done (* unused *) lemma ax_disj: "[|G,(A::'a triple set)\<turnstile>{P1::'a assn} t\<succ> {Q1}; G,A\<turnstile>{P2::'a assn} t\<succ> {Q2}|] ==> G,A\<turnstile>{λY s Z. P1 Y s Z ∨ P2 Y s Z} t\<succ> {λY s Z. Q1 Y s Z ∨ Q2 Y s Z}" apply (rule ax_escape (* unused *)) apply safe apply (erule conseq12, fast)+ done (* unused *) lemma ax_supd_shuffle: "(∃Q. G,(A::'a triple set)\<turnstile>{P::'a assn} .c1. {Q} ∧ G,A\<turnstile>{Q ;. f} .c2. {R}) = (∃Q'. G,A\<turnstile>{P} .c1. {f .; Q'} ∧ G,A\<turnstile>{Q'} .c2. {R})" apply (best elim!: conseq1 conseq2) done lemma ax_cases: " [|G,(A::'a triple set)\<turnstile>{P ∧. C} t\<succ> {Q::'a assn}; G,A\<turnstile>{P ∧. Not o C} t\<succ> {Q}|] ==> G,A\<turnstile>{P} t\<succ> {Q}" apply (unfold peek_and_def) apply (rule ax_escape) apply clarify apply (case_tac "C s") apply (erule conseq12, force)+ done (*alternative (more direct) proof: apply (rule rtac ax_derivs.conseq) *)(* unused *)(* apply clarify apply (case_tac "C s") apply force+ *) lemma ax_adapt: "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q} ==> G,A\<turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}" apply (unfold adapt_pre_def) apply (erule conseq12) apply fast done lemma adapt_pre_adapts: "G,(A::'a triple set)\<Turnstile>{P::'a assn} t\<succ> {Q} --> G,A\<Turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}" apply (unfold adapt_pre_def) apply (simp add: ax_valids_def triple_valid_def2) apply fast done lemma adapt_pre_weakest: "∀G (A::'a triple set) t. G,A\<Turnstile>{P} t\<succ> {Q} --> G,A\<Turnstile>{P'} t\<succ> {Q'} ==> P' => adapt_pre P Q (Q'::'a assn)" apply (unfold adapt_pre_def) apply (drule spec) apply (drule_tac x = "{}" in spec) apply (drule_tac x = "In1r Skip" in spec) apply (simp add: ax_valids_def triple_valid_def2) oops lemma peek_and_forget1_Normal: "G,(A::'a triple set)\<turnstile>{Normal P} t\<succ> {Q::'a assn} ==> G,A\<turnstile>{Normal (P ∧. p)} t\<succ> {Q}" apply (erule conseq1) apply (simp (no_asm)) done lemma peek_and_forget1: "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q} ==> G,A\<turnstile>{P ∧. p} t\<succ> {Q}" apply (erule conseq1) apply (simp (no_asm)) done lemmas ax_NormalD = peek_and_forget1 [of _ _ _ _ _ normal] lemma peek_and_forget2: "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q ∧. p} ==> G,A\<turnstile>{P} t\<succ> {Q}" apply (erule conseq2) apply (simp (no_asm)) done lemma ax_subst_Val_allI: "∀v. G,(A::'a triple set)\<turnstile>{(P' v )\<leftarrow>Val v} t\<succ> {(Q v)::'a assn} ==> ∀v. G,A\<turnstile>{(λw:. P' (the_In1 w))\<leftarrow>Val v} t\<succ> {Q v}" apply (force elim!: conseq1) done lemma ax_subst_Var_allI: "∀v. G,(A::'a triple set)\<turnstile>{(P' v )\<leftarrow>Var v} t\<succ> {(Q v)::'a assn} ==> ∀v. G,A\<turnstile>{(λw:. P' (the_In2 w))\<leftarrow>Var v} t\<succ> {Q v}" apply (force elim!: conseq1) done lemma ax_subst_Vals_allI: "(∀v. G,(A::'a triple set)\<turnstile>{( P' v )\<leftarrow>Vals v} t\<succ> {(Q v)::'a assn}) ==> ∀v. G,A\<turnstile>{(λw:. P' (the_In3 w))\<leftarrow>Vals v} t\<succ> {Q v}" apply (force elim!: conseq1) done section "alternative axioms" lemma ax_Lit2: "G,(A::'a triple set)\<turnstile>{Normal P::'a assn} Lit v-\<succ> {Normal (P\<down>=Val v)}" apply (rule ax_derivs.Lit [THEN conseq1]) apply force done lemma ax_Lit2_test_complete: "G,(A::'a triple set)\<turnstile>{Normal (P\<leftarrow>Val v)::'a assn} Lit v-\<succ> {P}" apply (rule ax_Lit2 [THEN conseq2]) apply force done lemma ax_LVar2: "G,(A::'a triple set)\<turnstile>{Normal P::'a assn} LVar vn=\<succ> {Normal (λs.. P\<down>=Var (lvar vn s))}" apply (rule ax_derivs.LVar [THEN conseq1]) apply force done lemma ax_Super2: "G,(A::'a triple set)\<turnstile> {Normal P::'a assn} Super-\<succ> {Normal (λs.. P\<down>=Val (val_this s))}" apply (rule ax_derivs.Super [THEN conseq1]) apply force done lemma ax_Nil2: "G,(A::'a triple set)\<turnstile>{Normal P::'a assn} []\<doteq>\<succ> {Normal (P\<down>=Vals [])}" apply (rule ax_derivs.Nil [THEN conseq1]) apply force done section "misc derived structural rules" (* unused *) lemma ax_finite_mtriples_lemma: "[|F ⊆ ms; finite ms; ∀(C,sig)∈ms. G,(A::'a triple set)\<turnstile>{Normal (P C sig)::'a assn} mb C sig-\<succ> {Q C sig}|] ==> G,A|\<turnstile>{{P} mb-\<succ> {Q} | F}" apply (frule (1) finite_subset) apply (erule rev_mp) apply (erule thin_rl) apply (erule finite_induct) apply (unfold mtriples_def) apply (clarsimp intro!: ax_derivs.empty ax_derivs.insert)+ apply force done lemmas ax_finite_mtriples = ax_finite_mtriples_lemma [OF subset_refl] lemma ax_derivs_insertD: "G,(A::'a triple set)|\<turnstile>insert (t::'a triple) ts ==> G,A\<turnstile>t ∧ G,A|\<turnstile>ts" apply (fast intro: ax_derivs.weaken) done lemma ax_methods_spec: "[|G,(A::'a triple set)|\<turnstile>split f ` ms; (C,sig) ∈ ms|]==> G,A\<turnstile>((f C sig)::'a triple)" apply (erule ax_derivs.weaken) apply (force del: image_eqI intro: rev_image_eqI) done (* this version is used to avoid using the cut rule *) lemma ax_finite_pointwise_lemma [rule_format]: "[|F ⊆ ms; finite ms|] ==> ((∀(C,sig)∈F. G,(A::'a triple set)\<turnstile>(f C sig::'a triple)) --> (∀(C,sig)∈ms. G,A\<turnstile>(g C sig::'a triple))) --> G,A|\<turnstile>split f ` F --> G,A|\<turnstile>split g ` F" apply (frule (1) finite_subset) apply (erule rev_mp) apply (erule thin_rl) apply (erule finite_induct) apply clarsimp+ apply (drule ax_derivs_insertD) apply (rule ax_derivs.insert) apply (simp (no_asm_simp) only: split_tupled_all) apply (auto elim: ax_methods_spec) done lemmas ax_finite_pointwise = ax_finite_pointwise_lemma [OF subset_refl] lemma ax_no_hazard: "G,(A::'a triple set)\<turnstile>{P ∧. type_ok G t} t\<succ> {Q::'a assn} ==> G,A\<turnstile>{P} t\<succ> {Q}" apply (erule ax_cases) apply (rule ax_derivs.hazard [THEN conseq1]) apply force done lemma ax_free_wt: "(∃T L C. (|prg=G,cls=C,lcl=L|)),\<turnstile>t::T) --> G,(A::'a triple set)\<turnstile>{Normal P} t\<succ> {Q::'a assn} ==> G,A\<turnstile>{Normal P} t\<succ> {Q}" apply (rule ax_no_hazard) apply (rule ax_escape) apply clarify apply (erule mp [THEN conseq12]) apply (auto simp add: type_ok_def) done ML_setup {* bind_thms ("ax_Abrupts", sum3_instantiate @{thm ax_derivs.Abrupt}) *} declare ax_Abrupts [intro!] lemmas ax_Normal_cases = ax_cases [of _ _ _ normal] lemma ax_Skip [intro!]: "G,(A::'a triple set)\<turnstile>{P\<leftarrow>♦} .Skip. {P::'a assn}" apply (rule ax_Normal_cases) apply (rule ax_derivs.Skip) apply fast done lemmas ax_SkipI = ax_Skip [THEN conseq1, standard] section "derived rules for methd call" lemma ax_Call_known_DynT: "[|G\<turnstile>IntVir->C\<preceq>statT; ∀a vs l. G,A\<turnstile>{(R a\<leftarrow>Vals vs ∧. (λs. l = locals (store s)) ;. init_lvars G C (|name=mn,parTs=pTs|)), IntVir a vs)} Methd C (|name=mn,parTs=pTs|)),-\<succ> {set_lvars l .; S}; ∀a. G,A\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {R a ∧. (λs. C = obj_class (the (heap (store s) (the_Addr a))) ∧ C = invocation_declclass G IntVir (store s) a statT (|name=mn,parTs=pTs|)), )}; G,(A::'a triple set)\<turnstile>{Normal P} e-\<succ> {Q::'a assn}|] ==> G,A\<turnstile>{Normal P} {accC,statT,IntVir}e·mn({pTs}args)-\<succ> {S}" apply (erule ax_derivs.Call) apply safe apply (erule spec) apply (rule ax_escape, clarsimp) apply (drule spec, drule spec, drule spec,erule conseq12) apply force done lemma ax_Call_Static: "[|∀a vs l. G,A\<turnstile>{R a\<leftarrow>Vals vs ∧. (λs. l = locals (store s)) ;. init_lvars G C (|name=mn,parTs=pTs|)), Static any_Addr vs} Methd C (|name=mn,parTs=pTs|)),-\<succ> {set_lvars l .; S}; G,A\<turnstile>{Normal P} e-\<succ> {Q}; ∀ a. G,(A::'a triple set)\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {(R::val => 'a assn) a ∧. (λ s. C=invocation_declclass G Static (store s) a statT (|name=mn,parTs=pTs|)),)} |] ==> G,A\<turnstile>{Normal P} {accC,statT,Static}e·mn({pTs}args)-\<succ> {S}" apply (erule ax_derivs.Call) apply safe apply (erule spec) apply (rule ax_escape, clarsimp) apply (erule_tac V = "?P --> ?Q" in thin_rl) apply (drule spec,drule spec,drule spec, erule conseq12) apply (force simp add: init_lvars_def Let_def) done lemma ax_Methd1: "[|G,A∪{{P} Methd-\<succ> {Q} | ms}|\<turnstile> {{P} body G-\<succ> {Q} | ms}; (C,sig)∈ ms|] ==> G,A\<turnstile>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}" apply (drule ax_derivs.Methd) apply (unfold mtriples_def) apply (erule (1) ax_methods_spec) done lemma ax_MethdN: "G,insert({Normal P} Methd C sig-\<succ> {Q}) A\<turnstile> {Normal P} body G C sig-\<succ> {Q} ==> G,A\<turnstile>{Normal P} Methd C sig-\<succ> {Q}" apply (rule ax_Methd1) apply (rule_tac [2] singletonI) apply (unfold mtriples_def) apply clarsimp done lemma ax_StatRef: "G,(A::'a triple set)\<turnstile>{Normal (P\<leftarrow>Val Null)} StatRef rt-\<succ> {P::'a assn}" apply (rule ax_derivs.Cast) apply (rule ax_Lit2 [THEN conseq2]) apply clarsimp done section "rules derived from Init and Done" lemma ax_InitS: "[|the (class G C) = c; C ≠ Object; ∀l. G,A\<turnstile>{Q ∧. (λs. l = locals (store s)) ;. set_lvars empty} .init c. {set_lvars l .; R}; G,A\<turnstile>{Normal ((P ∧. Not o initd C) ;. supd (init_class_obj G C))} .Init (super c). {Q}|] ==> G,(A::'a triple set)\<turnstile>{Normal (P ∧. Not o initd C)} .Init C. {R::'a assn}" apply (erule ax_derivs.Init) apply (simp (no_asm_simp)) apply assumption done lemma ax_Init_Skip_lemma: "∀l. G,(A::'a triple set)\<turnstile>{P\<leftarrow>♦ ∧. (λs. l = locals (store s)) ;. set_lvars l'} .Skip. {(set_lvars l .; P)::'a assn}" apply (rule allI) apply (rule ax_SkipI) apply clarsimp done lemma ax_triv_InitS: "[|the (class G C) = c;init c = Skip; C ≠ Object; P\<leftarrow>♦ => (supd (init_class_obj G C) .; P); G,A\<turnstile>{Normal (P ∧. initd C)} .Init (super c). {(P ∧. initd C)\<leftarrow>♦}|] ==> G,(A::'a triple set)\<turnstile>{Normal P\<leftarrow>♦} .Init C. {(P ∧. initd C)::'a assn}" apply (rule_tac C = "initd C" in ax_cases) apply (rule conseq1, rule ax_derivs.Done, clarsimp) apply (simp (no_asm)) apply (erule (1) ax_InitS) apply simp apply (rule ax_Init_Skip_lemma) apply (erule conseq1) apply force done lemma ax_Init_Object: "wf_prog G ==> G,(A::'a triple set)\<turnstile> {Normal ((supd (init_class_obj G Object) .; P\<leftarrow>♦) ∧. Not o initd Object)} .Init Object. {(P ∧. initd Object)::'a assn}" apply (rule ax_derivs.Init) apply (drule class_Object, force) apply (simp_all (no_asm)) apply (rule_tac [2] ax_Init_Skip_lemma) apply (rule ax_SkipI, force) done lemma ax_triv_Init_Object: "[|wf_prog G; (P::'a assn) => (supd (init_class_obj G Object) .; P)|] ==> G,(A::'a triple set)\<turnstile>{Normal P\<leftarrow>♦} .Init Object. {P ∧. initd Object}" apply (rule_tac C = "initd Object" in ax_cases) apply (rule conseq1, rule ax_derivs.Done, clarsimp) apply (erule ax_Init_Object [THEN conseq1]) apply force done section "introduction rules for Alloc and SXAlloc" lemma ax_SXAlloc_Normal: "G,(A::'a triple set)\<turnstile>{P::'a assn} .c. {Normal Q} ==> G,A\<turnstile>{P} .c. {SXAlloc G Q}" apply (erule conseq2) apply (clarsimp elim!: sxalloc_elim_cases simp add: split_tupled_all) done lemma ax_Alloc: "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Normal (λY (x,s) Z. (∀a. new_Addr (heap s) = Some a --> Q (Val (Addr a)) (Norm(init_obj G (CInst C) (Heap a) s)) Z)) ∧. heap_free (Suc (Suc 0))} ==> G,A\<turnstile>{P} t\<succ> {Alloc G (CInst C) Q}" apply (erule conseq2) apply (auto elim!: halloc_elim_cases) done lemma ax_Alloc_Arr: "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {λVal:i:. Normal (λY (x,s) Z. ¬the_Intg i<0 ∧ (∀a. new_Addr (heap s) = Some a --> Q (Val (Addr a)) (Norm (init_obj G (Arr T (the_Intg i)) (Heap a) s)) Z)) ∧. heap_free (Suc (Suc 0))} ==> G,A\<turnstile>{P} t\<succ> {λVal:i:. abupd (check_neg i) .; Alloc G (Arr T(the_Intg i)) Q}" apply (erule conseq2) apply (auto elim!: halloc_elim_cases) done lemma ax_SXAlloc_catch_SXcpt: "[|G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {(λY (x,s) Z. x=Some (Xcpt (Std xn)) ∧ (∀a. new_Addr (heap s) = Some a --> Q Y (Some (Xcpt (Loc a)),init_obj G (CInst (SXcpt xn)) (Heap a) s) Z)) ∧. heap_free (Suc (Suc 0))}|] ==> G,A\<turnstile>{P} t\<succ> {SXAlloc G (λY s Z. Q Y s Z ∧ G,s\<turnstile>catch SXcpt xn)}" apply (erule conseq2) apply (auto elim!: sxalloc_elim_cases halloc_elim_cases) done end
lemma assn_imp_def2:
(P => Q) = (∀Y s Z. P Y s Z --> Q Y s Z)
lemma peek_and_def2:
(P ∧. p) Y s = (λZ. P Y s Z ∧ p s)
lemma peek_and_Not:
(P ∧. (λs. ¬ f s)) = (P ∧. Not o f)
lemma peek_and_and:
(P ∧. p ∧. p) = (P ∧. p)
lemma peek_and_commut:
(P ∧. p ∧. q) = (P ∧. q ∧. p)
lemma peek_and_Normal:
(Normal P ∧. p) = Normal (P ∧. p)
lemma assn_supd_def2:
(P ;. f) Y s' Z = (∃s. P Y s Z ∧ s' = f s)
lemma supd_assn_def2:
(f .; P) Y s = P Y (f s)
lemma supd_assn_supdD:
(f .; Q ;. f) Y s Z ==> Q Y s Z
lemma supd_assn_supdI:
Q Y s Z ==> (f .; Q ;. f) Y s Z
lemma subst_res_def2:
(P\<leftarrow>w) Y = P w
lemma subst_subst_res:
P\<leftarrow>w\<leftarrow>v = P\<leftarrow>w
lemma peek_and_subst_res:
(P ∧. p)\<leftarrow>w = (P\<leftarrow>w ∧. p)
lemma subst_Bool_def2:
(P\<leftarrow>=b) Y s Z = (∃v. P (In1 v) s Z ∧ (normal s --> the_Bool v = b))
lemma subst_Bool_the_BoolI:
P (In1 b) s Z ==> (P\<leftarrow>=the_Bool b) Y s Z
lemma peek_res_def2:
peek_res P Y = P Y Y
lemma peek_res_subst_res:
peek_res P\<leftarrow>w = P w\<leftarrow>w
lemma peek_subst_res_allI:
(!!a. T a (P (f a)\<leftarrow>f a)) ==> ∀a. T a (peek_res P\<leftarrow>f a)
lemma ign_res_def2:
P\<down> Y s Z = (∃Y. P Y s Z)
lemma ign_ign_res:
P\<down>\<down> = P\<down>
lemma ign_subst_res:
P\<down>\<leftarrow>w = P\<down>
lemma peek_and_ign_res:
(P ∧. p)\<down> = (P\<down> ∧. p)
lemma peek_st_def2:
peek_st Pf Y s = Pf (snd s) Y s
lemma peek_st_triv:
(λs.. P) = P
lemma peek_st_st:
(λs.. peek_st (P s)) = (λs.. P s s)
lemma peek_st_split:
peek_st P = (λY s. P (snd s) Y s)
lemma peek_st_subst_res:
peek_st P\<leftarrow>w = (λs.. P s\<leftarrow>w)
lemma peek_st_Normal:
(λs.. Normal (P s)) = Normal (peek_st P)
lemma ign_res_eq_def2:
(P\<down>=w) Y s Z = ((∃Y. P Y s Z) ∧ Y = w)
lemma ign_ign_res_eq:
(P\<down>=w)\<down> = P\<down>
lemma ign_res_eq_subst_res:
P\<down>=w\<leftarrow>w = P\<down>
lemma subst_Bool_ign_res_eq:
(P\<leftarrow>=b\<down>=x) Y s Z = ((P\<leftarrow>=b) Y s Z ∧ Y = x)
lemma RefVar_def2:
(vf ..; P) Y s = P (In2 (fst (vf s))) (snd (vf s))
lemma Alloc_def2:
Alloc G otag P Y s Z =
(∀s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a-> s' -->
P (In1 (Addr a)) s' Z)
lemma SXAlloc_def2:
SXAlloc G P Y s Z = (∀s'. G\<turnstile>s \<midarrow>sxalloc-> s' --> P Y s' Z)
lemma inj_triple:
inj (λ(P, t, Q). {P} t> {Q})
lemma triple_inj_eq:
({P} t> {Q} = {P'} t'> {Q'}) = (P = P' ∧ t = t' ∧ Q = Q')
lemma triple_valid_def2:
G\<Turnstile>n:{P} t> {Q} =
(∀Y s Z.
P Y s Z -->
(∃L. (normal s -->
(∃C T A.
(| prg = G, cls = C, lcl = L, ... = () |)|-t::T ∧
(| prg = G, cls = C, lcl = L,
... = () |)\<turnstile> dom (locals (snd s)) »t» A)) ∧
s::\<preceq>(G, L)) -->
(∀Y' s'.
G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n-> (Y', s') -->
Q Y' s' Z))
lemma cut_valid:
[| G,A'|\<Turnstile>ts; G,A|\<Turnstile>A' |] ==> G,A|\<Turnstile>ts
lemma ax_thin:
[| G,A'|\<turnstile>ts; A' ⊆ A |] ==> G,A|\<turnstile>ts
lemma ax_thin_insert:
G,A\<turnstile>t ==> G,insert x A\<turnstile>t
lemma subset_mtriples_iff:
(ts ⊆ {{P} mb-\<succ> {Q} | ms}) = (∃ms'⊆ms. ts = {{P} mb-\<succ> {Q} | ms'})
lemma conseq12:
[| G,A\<turnstile>{P'} t> {Q'};
∀Y s Z.
P Y s Z --> (∀Y' s'. (∀Y Z'. P' Y s Z' --> Q' Y' s' Z') --> Q Y' s' Z) |]
==> G,A\<turnstile>{P} t> {Q}
lemma conseq12':
[| G,A\<turnstile>{P'} t> {Q'};
∀s Y' s'.
(∀Y Z. P' Y s Z --> Q' Y' s' Z) --> (∀Y Z. P Y s Z --> Q Y' s' Z) |]
==> G,A\<turnstile>{P} t> {Q}
lemma conseq12_from_conseq12':
[| G,A\<turnstile>{P'} t> {Q'};
∀Y s Z.
P Y s Z --> (∀Y' s'. (∀Y Z'. P' Y s Z' --> Q' Y' s' Z') --> Q Y' s' Z) |]
==> G,A\<turnstile>{P} t> {Q}
lemma conseq1:
[| G,A\<turnstile>{P'} t> {Q}; P => P' |] ==> G,A\<turnstile>{P} t> {Q}
lemma conseq2:
[| G,A\<turnstile>{P} t> {Q'}; Q' => Q |] ==> G,A\<turnstile>{P} t> {Q}
lemma ax_escape:
∀Y s Z.
P Y s Z --> G,A\<turnstile>{λY' s' Z'. (Y', s') = (Y, s)} t>
{λY s Z'. Q Y s Z}
==> G,A\<turnstile>{P} t> {Q}
lemma ax_constant:
(C ==> G,A\<turnstile>{P} t> {Q})
==> G,A\<turnstile>{λY s Z. C ∧ P Y s Z} t> {Q}
lemma ax_impossible:
G,A\<turnstile>{λY s Z. False} t> {Q}
lemma ax_nochange_lemma:
[| P Y s; All (op = w) |] ==> P w s
lemma ax_nochange:
G,A\<turnstile>{λY s. op = (Y, s)} t> {λY s. op = (Y, s)}
==> G,A\<turnstile>{P} t> {P}
lemma ax_trivial:
G,A\<turnstile>{P} t> {λY s Z. True}
lemma ax_disj:
[| G,A\<turnstile>{P1.0} t> {Q1.0}; G,A\<turnstile>{P2.0} t> {Q2.0} |]
==> G,A\<turnstile>{λY s Z. P1.0 Y s Z ∨ P2.0 Y s Z} t>
{λY s Z. Q1.0 Y s Z ∨ Q2.0 Y s Z}
lemma ax_supd_shuffle:
(∃Q. G,A\<turnstile>{P} .c1.0. {Q} ∧ G,A\<turnstile>{Q ;. f} .c2.0. {R}) =
(∃Q'. G,A\<turnstile>{P} .c1.0. {f .; Q'} ∧ G,A\<turnstile>{Q'} .c2.0. {R})
lemma ax_cases:
[| G,A\<turnstile>{P ∧. C} t> {Q}; G,A\<turnstile>{P ∧. Not o C} t> {Q} |]
==> G,A\<turnstile>{P} t> {Q}
lemma ax_adapt:
G,A\<turnstile>{P} t> {Q} ==> G,A\<turnstile>{adapt_pre P Q Q'} t> {Q'}
lemma adapt_pre_adapts:
G,A|={P} t> {Q} --> G,A|={adapt_pre P Q Q'} t> {Q'}
lemma peek_and_forget1_Normal:
G,A\<turnstile>{Normal P} t> {Q} ==> G,A\<turnstile>{Normal (P ∧. p)} t> {Q}
lemma peek_and_forget1:
G,A\<turnstile>{P} t> {Q} ==> G,A\<turnstile>{P ∧. p} t> {Q}
lemma ax_NormalD:
G,A\<turnstile>{P} t> {Q} ==> G,A\<turnstile>{Normal P} t> {Q}
lemma peek_and_forget2:
G,A\<turnstile>{P} t> {Q ∧. p} ==> G,A\<turnstile>{P} t> {Q}
lemma ax_subst_Val_allI:
∀v. G,A\<turnstile>{P' v\<leftarrow>In1 v} t> {Q v}
==> ∀v. G,A\<turnstile>{(λw:. P' (the_In1 w))\<leftarrow>In1 v} t> {Q v}
lemma ax_subst_Var_allI:
∀v. G,A\<turnstile>{P' v\<leftarrow>In2 v} t> {Q v}
==> ∀v. G,A\<turnstile>{(λw:. P' (the_In2 w))\<leftarrow>In2 v} t> {Q v}
lemma ax_subst_Vals_allI:
∀v. G,A\<turnstile>{P' v\<leftarrow>In3 v} t> {Q v}
==> ∀v. G,A\<turnstile>{(λw:. P' (the_In3 w))\<leftarrow>In3 v} t> {Q v}
lemma ax_Lit2:
G,A\<turnstile>{Normal P} Lit v-> {Normal (P\<down>=In1 v)}
lemma ax_Lit2_test_complete:
G,A\<turnstile>{Normal (P\<leftarrow>In1 v)} Lit v-> {P}
lemma ax_LVar2:
G,A\<turnstile>{Normal P} LVar vn=> {Normal (λs.. P\<down>=In2 (lvar vn s))}
lemma ax_Super2:
G,A\<turnstile>{Normal P} Super-> {Normal (λs.. P\<down>=In1 (val_this s))}
lemma ax_Nil2:
G,A\<turnstile>{Normal P} []#> {Normal (P\<down>=In3 [])}
lemma ax_finite_mtriples_lemma:
[| F ⊆ ms; finite ms;
∀(C, sig)∈ms. G,A\<turnstile>{Normal (P C sig)} mb C sig-> {Q C sig} |]
==> G,A|\<turnstile>{{P} mb-\<succ> {Q} | F}
lemma ax_finite_mtriples:
[| finite F;
∀(C, sig)∈F. G,A\<turnstile>{Normal (P C sig)} mb C sig-> {Q C sig} |]
==> G,A|\<turnstile>{{P} mb-\<succ> {Q} | F}
lemma ax_derivs_insertD:
G,A|\<turnstile>insert t ts ==> G,A\<turnstile>t ∧ G,A|\<turnstile>ts
lemma ax_methods_spec:
[| G,A|\<turnstile>(λ(x, y). f x y) ` ms; (C, sig) ∈ ms |]
==> G,A\<turnstile>f C sig
lemma ax_finite_pointwise_lemma:
[| F ⊆ ms; finite ms;
!!x. [| !!x. x ∈ F ==> (λ(C, sig). G,A\<turnstile>f C sig) x; x ∈ ms |]
==> (λ(C, sig). G,A\<turnstile>g C sig) x;
G,A|\<turnstile>(λ(x, y). f x y) ` F |]
==> G,A|\<turnstile>(λ(x, y). g x y) ` F
lemma ax_finite_pointwise:
[| finite F;
!!x. [| !!x. x ∈ F ==> (λ(C, sig). G,A\<turnstile>f C sig) x; x ∈ F |]
==> (λ(C, sig). G,A\<turnstile>g C sig) x;
G,A|\<turnstile>(λ(x, y). f x y) ` F |]
==> G,A|\<turnstile>(λ(x, y). g x y) ` F
lemma ax_no_hazard:
G,A\<turnstile>{P ∧. type_ok G t} t> {Q} ==> G,A\<turnstile>{P} t> {Q}
lemma ax_free_wt:
(∃T L C. (| prg = G, cls = C, lcl = L, ... = () |)|-t::T) -->
G,A\<turnstile>{Normal P} t> {Q}
==> G,A\<turnstile>{Normal P} t> {Q}
theorem ax_Abrupts:
G,A\<turnstile>{P\<leftarrow>In1 arbitrary ∧. Not o normal} x-> {P}
G,A\<turnstile>{P\<leftarrow>In2 arbitrary ∧. Not o normal} x=> {P}
G,A\<turnstile>{P\<leftarrow>In3 arbitrary ∧. Not o normal} x#> {P}
G,A\<turnstile>{P\<leftarrow>dummy_res ∧. Not o normal} .x. {P}
lemma ax_Normal_cases:
[| G,A\<turnstile>{Normal P} t> {Q};
G,A\<turnstile>{P ∧. Not o normal} t> {Q} |]
==> G,A\<turnstile>{P} t> {Q}
lemma ax_Skip:
G,A\<turnstile>{P\<leftarrow>dummy_res} .Skip. {P}
lemma ax_SkipI:
P => Q\<leftarrow>dummy_res ==> G,A\<turnstile>{P} .Skip. {Q}
lemma ax_Call_known_DynT:
[| G\<turnstile>IntVir->C\<preceq>statT;
∀a vs l.
G,A\<turnstile>{R a\<leftarrow>In3 vs ∧. (λs. l = locals (snd s)) ;.
init_lvars G C (| name = mn, parTs = pTs, ... = () |)
IntVir a vs}
Methd C (| name = mn, parTs = pTs, ... = () |)-> {set_lvars l .; S};
∀a. G,A\<turnstile>{Q\<leftarrow>In1 a} args#>
{R a ∧.
(λs. C = obj_class (lookup_obj (snd s) a) ∧
C = invocation_declclass G IntVir (snd s) a statT
(| name = mn, parTs = pTs, ... = () |))};
G,A\<turnstile>{Normal P} e-> {Q} |]
==> G,A\<turnstile>{Normal P} {accC,statT,IntVir}e·mn( {pTs}args)-> {S}
lemma ax_Call_Static:
[| ∀a vs l.
G,A\<turnstile>{R a\<leftarrow>In3 vs ∧. (λs. l = locals (snd s)) ;.
init_lvars G C (| name = mn, parTs = pTs, ... = () |)
Static any_Addr vs}
Methd C (| name = mn, parTs = pTs, ... = () |)-> {set_lvars l .; S};
G,A\<turnstile>{Normal P} e-> {Q};
∀a. G,A\<turnstile>{Q\<leftarrow>In1 a} args#>
{R a ∧.
(λs. C = invocation_declclass G Static (snd s) a statT
(| name = mn, parTs = pTs, ... = () |))} |]
==> G,A\<turnstile>{Normal P} {accC,statT,Static}e·mn( {pTs}args)-> {S}
lemma ax_Methd1:
[| G,A ∪ {{P} Methd-\<succ> {Q} | ms}|\<turnstile>{{P} body G-\<succ> {Q} | ms};
(C, sig) ∈ ms |]
==> G,A\<turnstile>{Normal (P C sig)} Methd C sig-> {Q C sig}
lemma ax_MethdN:
G,insert ({Normal P} Methd C sig-> {Q}) A\<turnstile>{Normal P} body G C sig->
{Q}
==> G,A\<turnstile>{Normal P} Methd C sig-> {Q}
lemma ax_StatRef:
G,A\<turnstile>{Normal (P\<leftarrow>In1 Null)} StatRef rt-> {P}
lemma ax_InitS:
[| the (class G C) = c; C ≠ Object;
∀l. G,A\<turnstile>{Q ∧. (λs. l = locals (snd s)) ;. set_lvars empty}
.init c. {set_lvars l .; R};
G,A\<turnstile>{Normal (P ∧. Not o initd C ;. supd (init_class_obj G C))}
.Init (super c). {Q} |]
==> G,A\<turnstile>{Normal (P ∧. Not o initd C)} .Init C. {R}
lemma ax_Init_Skip_lemma:
∀l. G,A\<turnstile>{P\<leftarrow>dummy_res ∧. (λs. l = locals (snd s)) ;.
set_lvars l'}
.Skip. {set_lvars l .; P}
lemma ax_triv_InitS:
[| the (class G C) = c; init c = Skip; C ≠ Object;
P\<leftarrow>dummy_res => (supd (init_class_obj G C) .; P);
G,A\<turnstile>{Normal (P ∧. initd C)} .Init (super c).
{(P ∧. initd C)\<leftarrow>dummy_res} |]
==> G,A\<turnstile>{Normal P\<leftarrow>dummy_res} .Init C. {P ∧. initd C}
lemma ax_Init_Object:
wf_prog G
==> G,A\<turnstile>{Normal
(supd (init_class_obj G Object) .;
P\<leftarrow>dummy_res ∧.
Not o initd Object)}
.Init Object. {P ∧. initd Object}
lemma ax_triv_Init_Object:
[| wf_prog G; P => (supd (init_class_obj G Object) .; P) |]
==> G,A\<turnstile>{Normal P\<leftarrow>dummy_res} .Init Object.
{P ∧. initd Object}
lemma ax_SXAlloc_Normal:
G,A\<turnstile>{P} .c. {Normal Q} ==> G,A\<turnstile>{P} .c. {SXAlloc G Q}
lemma ax_Alloc:
G,A\<turnstile>{P} t>
{Normal
(λY (x, s) Z.
∀a. new_Addr (heap s) = Some a -->
Q (In1 (Addr a)) (Norm (init_obj G (CInst C) (Inl a) s)) Z) ∧.
heap_free (Suc (Suc 0))}
==> G,A\<turnstile>{P} t> {Alloc G (CInst C) Q}
lemma ax_Alloc_Arr:
G,A\<turnstile>{P} t>
{λVal:i:. Normal
(λY (x, s) Z.
¬ the_Intg i < 0 ∧
(∀a. new_Addr (heap s) = Some a -->
Q (In1 (Addr a))
(Norm (init_obj G (Arr T (the_Intg i)) (Inl a) s)) Z)) ∧.
heap_free (Suc (Suc 0))}
==> G,A\<turnstile>{P} t>
{λVal:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) Q}
lemma ax_SXAlloc_catch_SXcpt:
G,A\<turnstile>{P} t>
{(λY (x, s) Z.
x = Some (Xcpt (Std xn)) ∧
(∀a. new_Addr (heap s) = Some a -->
Q Y (Some (Xcpt (Loc a)), init_obj G (CInst (SXcpt xn)) (Inl a) s)
Z)) ∧.
heap_free (Suc (Suc 0))}
==> G,A\<turnstile>{P} t>
{SXAlloc G (λY s Z. Q Y s Z ∧ G,s\<turnstile>catch SXcpt xn)}