Theory AxSem
subsection ‹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}
›
type_synonym res = vals
abbreviation (input)
Val where "Val x == In1 x"
abbreviation (input)
Var where "Var x == In2 x"
abbreviation (input)
Vals where "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) ∘ CONST the_In1"
"λVar:v . b" == "(λv. b) ∘ CONST the_In2"
"λVals:v. b" == "(λv. b) ∘ CONST the_In3"
type_synonym 'a assn = "res ⇒ state ⇒ 'a ⇒ bool"
translations
(type) "'a assn" <= (type) "vals ⇒ state ⇒ 'a ⇒ bool"
definition
assn_imp :: "'a assn ⇒ 'a assn ⇒ bool" (infixr "⇒" 25)
where "(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
subsubsection "assertion transformers"
subsection "peek-and"
definition
peek_and :: "'a assn ⇒ (state ⇒ bool) ⇒ 'a assn" (infixl "∧." 13)
where "(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 ∘ 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
abbreviation
Normal :: "'a assn ⇒ 'a assn"
where "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"
definition
assn_supd :: "'a assn ⇒ (state ⇒ state) ⇒ 'a assn" (infixl ";." 13)
where "(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"
definition
supd_assn :: "(state ⇒ state) ⇒ 'a assn ⇒ 'a assn" (infixr ".;" 13)
where "(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"
definition
subst_res :: "'a assn ⇒ res ⇒ 'a assn" ("_←_" [60,61] 60)
where "P←w = (λY. P w)"
lemma subst_res_def2 [simp]: "(P←w) Y = P w"
apply (unfold subst_res_def)
apply (simp (no_asm))
done
lemma subst_subst_res [simp]: "P←w←v = P←w"
apply (rule ext)
apply (simp (no_asm))
done
lemma peek_and_subst_res [simp]: "(P ∧. p)←w = (P←w ∧. p)"
apply (rule ext)
apply (rule ext)
apply (simp (no_asm))
done
subsection "subst-Bool"
definition
subst_Bool :: "'a assn ⇒ bool ⇒ 'a assn" ("_←=_" [60,61] 60)
where "P←=b = (λY s Z. ∃v. P (Val v) s Z ∧ (normal s ⟶ the_Bool v=b))"
lemma subst_Bool_def2 [simp]:
"(P←=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←=the_Bool b) Y s Z"
apply auto
done
subsection "peek-res"
definition
peek_res :: "(res ⇒ 'a assn) ⇒ 'a assn"
where "peek_res Pf = (λY. Pf Y Y)"
syntax
"_peek_res" :: "pttrn ⇒ 'a assn ⇒ 'a assn" ("λ_:. _" [0,3] 3)
translations
"λw:. P" == "CONST 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←w = P w←w"
apply (rule ext)
apply (simp (no_asm))
done
lemma peek_subst_res_allI:
"(⋀a. T a (P (f a)←f a)) ⟹ ∀a. T a (peek_res P←f a)"
apply (rule allI)
apply (simp (no_asm))
apply fast
done
subsection "ign-res"
definition
ign_res :: "'a assn ⇒ 'a assn" ("_↓" [1000] 1000)
where "P↓ = (λY s Z. ∃Y. P Y s Z)"
lemma ign_res_def2 [simp]: "P↓ Y s Z = (∃Y. P Y s Z)"
apply (unfold ign_res_def)
apply (simp (no_asm))
done
lemma ign_ign_res [simp]: "P↓↓ = P↓"
apply (rule ext)
apply (rule ext)
apply (rule ext)
apply (simp (no_asm))
done
lemma ign_subst_res [simp]: "P↓←w = P↓"
apply (rule ext)
apply (rule ext)
apply (rule ext)
apply (simp (no_asm))
done
lemma peek_and_ign_res [simp]: "(P ∧. p)↓ = (P↓ ∧. p)"
apply (rule ext)
apply (rule ext)
apply (rule ext)
apply (simp (no_asm))
done
subsection "peek-st"
definition
peek_st :: "(st ⇒ 'a assn) ⇒ 'a assn"
where "peek_st P = (λY s. P (store s) Y s)"
syntax
"_peek_st" :: "pttrn ⇒ 'a assn ⇒ 'a assn" ("λ_.. _" [0,3] 3)
translations
"λs.. P" == "CONST 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)←w = (λs.. P s←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"
definition
ign_res_eq :: "'a assn ⇒ res ⇒ 'a assn" ("_↓=_" [60,61] 60)
where "P↓=w ≡ (λY:. P↓ ∧. (λs. Y=w))"
lemma ign_res_eq_def2 [simp]: "(P↓=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↓=w)↓ = P↓"
apply (rule ext)
apply (rule ext)
apply (rule ext)
apply (simp (no_asm))
done
lemma ign_res_eq_subst_res: "P↓=w←w = P↓"
apply (rule ext)
apply (rule ext)
apply (rule ext)
apply (simp (no_asm))
done
lemma subst_Bool_ign_res_eq: "((P←=b)↓=x) Y s Z = ((P←=b) Y s Z ∧ Y=x)"
apply (simp (no_asm))
done
subsection "RefVar"
definition
RefVar :: "(state ⇒ vvar × state) ⇒ 'a assn ⇒ 'a assn" (infixr "..;" 13)
where "(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"
definition
Alloc :: "prog ⇒ obj_tag ⇒ 'a assn ⇒ 'a assn"
where "Alloc G otag P = (λY s Z. ∀s' a. G⊢s ─halloc otag≻a→ s'⟶ P (Val (Addr a)) s' Z)"
definition
SXAlloc :: "prog ⇒ 'a assn ⇒ 'a assn"
where "SXAlloc G P = (λY s Z. ∀s'. G⊢s ─sxalloc→ s' ⟶ P Y s' Z)"
lemma Alloc_def2 [simp]: "Alloc G otag P Y s Z =
(∀s' a. G⊢s ─halloc otag≻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⊢s ─sxalloc→ s' ⟶ P Y s' Z)"
apply (unfold SXAlloc_def)
apply (simp (no_asm))
done
subsubsection "validity"
definition
type_ok :: "prog ⇒ term ⇒ state ⇒ bool" where
"type_ok G t s =
(∃L T C A. (normal s ⟶ ⦇prg=G,cls=C,lcl=L⦈⊢t∷T ∧
⦇prg=G,cls=C,lcl=L⦈⊢dom (locals (store s))»t»A )
∧ s∷≼(G,L))"