(* Title: HOL/Bali/Eval.thy ID: $Id: Eval.thy,v 1.20 2008/03/29 18:14:01 wenzelm Exp $ Author: David von Oheimb *) header {* Operational evaluation (big-step) semantics of Java expressions and statements *} theory Eval imports State DeclConcepts begin text {* improvements over Java Specification 1.0: \begin{itemize} \item dynamic method lookup does not need to consider the return type (cf.15.11.4.4) \item throw raises a NullPointer exception if a null reference is given, and each throw of a standard exception yield a fresh exception object (was not specified) \item if there is not enough memory even to allocate an OutOfMemory exception, evaluation/execution fails, i.e. simply stops (was not specified) \item array assignment checks lhs (and may throw exceptions) before evaluating rhs \item fixed exact positions of class initializations (immediate at first active use) \end{itemize} design issues: \begin{itemize} \item evaluation vs. (single-step) transition semantics evaluation semantics chosen, because: \begin{itemize} \item[++] less verbose and therefore easier to read (and to handle in proofs) \item[+] more abstract \item[+] intermediate values (appearing in recursive rules) need not be stored explicitly, e.g. no call body construct or stack of invocation frames containing local variables and return addresses for method calls needed \item[+] convenient rule induction for subject reduction theorem \item[-] no interleaving (for parallelism) can be described \item[-] stating a property of infinite executions requires the meta-level argument that this property holds for any finite prefixes of it (e.g. stopped using a counter that is decremented to zero and then throwing an exception) \end{itemize} \item unified evaluation for variables, expressions, expression lists, statements \item the value entry in statement rules is redundant \item the value entry in rules is irrelevant in case of exceptions, but its full inclusion helps to make the rule structure independent of exception occurence. \item as irrelevant value entries are ignored, it does not matter if they are unique. For simplicity, (fixed) arbitrary values are preferred over "free" values. \item the rule format is such that the start state may contain an exception. \begin{itemize} \item[++] faciliates exception handling \item[+] symmetry \end{itemize} \item the rules are defined carefully in order to be applicable even in not type-correct situations (yielding undefined values), e.g. @{text "the_Addr (Val (Bool b)) = arbitrary"}. \begin{itemize} \item[++] fewer rules \item[-] less readable because of auxiliary functions like @{text the_Addr} \end{itemize} Alternative: "defensive" evaluation throwing some InternalError exception in case of (impossible, for correct programs) type mismatches \item there is exactly one rule per syntactic construct \begin{itemize} \item[+] no redundancy in case distinctions \end{itemize} \item halloc fails iff there is no free heap address. When there is only one free heap address left, it returns an OutOfMemory exception. In this way it is guaranteed that when an OutOfMemory exception is thrown for the first time, there is a free location on the heap to allocate it. \item the allocation of objects that represent standard exceptions is deferred until execution of any enclosing catch clause, which is transparent to the program. \begin{itemize} \item[-] requires an auxiliary execution relation \item[++] avoids copies of allocation code and awkward case distinctions (whether there is enough memory to allocate the exception) in evaluation rules \end{itemize} \item unfortunately @{text new_Addr} is not directly executable because of Hilbert operator. \end{itemize} simplifications: \begin{itemize} \item local variables are initialized with default values (no definite assignment) \item garbage collection not considered, therefore also no finalizers \item stack overflow and memory overflow during class initialization not modelled \item exceptions in initializations not replaced by ExceptionInInitializerError \end{itemize} *} types vvar = "val × (val => state => state)" vals = "(val, vvar, val list) sum3" translations "vvar" <= (type) "val × (val => state => state)" "vals" <= (type)"(val, vvar, val list) sum3" text {* To avoid redundancy and to reduce the number of rules, there is only one evaluation rule for each syntactic term. This is also true for variables (e.g. see the rules below for @{text LVar}, @{text FVar} and @{text AVar}). So evaluation of a variable must capture both possible further uses: read (rule @{text Acc}) or write (rule @{text Ass}) to the variable. Therefor a variable evaluates to a special value @{term vvar}, which is a pair, consisting of the current value (for later read access) and an update function (for later write access). Because during assignment to an array variable an exception may occur if the types don't match, the update function is very generic: it transforms the full state. This generic update function causes some technical trouble during some proofs (e.g. type safety, correctness of definite assignment). There we need to prove some additional invariant on this update function to prove the assignment correct, since the update function could potentially alter the whole state in an arbitrary manner. This invariant must be carried around through the whole induction. So for future approaches it may be better not to take such a generic update function, but only to store the address and the kind of variable (array (+ element type), local variable or field) for later assignment. *} syntax (xsymbols) dummy_res :: "vals" ("♦") translations "♦" == "In1 Unit" syntax val_inj_vals:: "expr => term" ("⌊_⌋e" 1000) var_inj_vals:: "var => term" ("⌊_⌋v" 1000) lst_inj_vals:: "expr list => term" ("⌊_⌋l" 1000) translations "⌊e⌋e" \<rightharpoonup> "In1 e" "⌊v⌋v" \<rightharpoonup> "In2 v" "⌊es⌋l" \<rightharpoonup> "In3 es" constdefs arbitrary3 :: "('al + 'ar, 'b, 'c) sum3 => vals" "arbitrary3 ≡ sum3_case (In1 o sum_case (λx. arbitrary) (λx. Unit)) (λx. In2 arbitrary) (λx. In3 arbitrary)" lemma [simp]: "arbitrary3 (In1l x) = In1 arbitrary" by (simp add: arbitrary3_def) lemma [simp]: "arbitrary3 (In1r x) = ♦" by (simp add: arbitrary3_def) lemma [simp]: "arbitrary3 (In2 x) = In2 arbitrary" by (simp add: arbitrary3_def) lemma [simp]: "arbitrary3 (In3 x) = In3 arbitrary" by (simp add: arbitrary3_def) section "exception throwing and catching" constdefs throw :: "val => abopt => abopt" "throw a' x ≡ abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)" lemma throw_def2: "throw a' x = abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)" apply (unfold throw_def) apply (simp (no_asm)) done constdefs fits :: "prog => st => val => ty => bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60) "G,s\<turnstile>a' fits T ≡ (∃rt. T=RefT rt) --> a'=Null ∨ G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T" lemma fits_Null [simp]: "G,s\<turnstile>Null fits T" by (simp add: fits_def) lemma fits_Addr_RefT [simp]: "G,s\<turnstile>Addr a fits RefT t = G\<turnstile>obj_ty (the (heap s a))\<preceq>RefT t" by (simp add: fits_def) lemma fitsD: "!!X. G,s\<turnstile>a' fits T ==> (∃pt. T = PrimT pt) ∨ (∃t. T = RefT t) ∧ a' = Null ∨ (∃t. T = RefT t) ∧ a' ≠ Null ∧ G\<turnstile>obj_ty (lookup_obj s a')\<preceq>T" apply (unfold fits_def) apply (case_tac "∃pt. T = PrimT pt") apply simp_all apply (case_tac "T") defer apply (case_tac "a' = Null") apply simp_all apply iprover done constdefs catch ::"prog => state => qtname => bool" ("_,_\<turnstile>catch _"[61,61,61]60) "G,s\<turnstile>catch C≡∃xc. abrupt s=Some (Xcpt xc) ∧ G,store s\<turnstile>Addr (the_Loc xc) fits Class C" lemma catch_Norm [simp]: "¬G,Norm s\<turnstile>catch tn" apply (unfold catch_def) apply (simp (no_asm)) done lemma catch_XcptLoc [simp]: "G,(Some (Xcpt (Loc a)),s)\<turnstile>catch C = G,s\<turnstile>Addr a fits Class C" apply (unfold catch_def) apply (simp (no_asm)) done lemma catch_Jump [simp]: "¬G,(Some (Jump j),s)\<turnstile>catch tn" apply (unfold catch_def) apply (simp (no_asm)) done lemma catch_Error [simp]: "¬G,(Some (Error e),s)\<turnstile>catch tn" apply (unfold catch_def) apply (simp (no_asm)) done constdefs new_xcpt_var :: "vname => state => state" "new_xcpt_var vn ≡ λ(x,s). Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)" lemma new_xcpt_var_def2 [simp]: "new_xcpt_var vn (x,s) = Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)" apply (unfold new_xcpt_var_def) apply (simp (no_asm)) done section "misc" constdefs assign :: "('a => state => state) => 'a => state => state" "assign f v ≡ λ(x,s). let (x',s') = (if x = None then f v else id) (x,s) in (x',if x' = None then s' else s)" (* lemma assign_Norm_Norm [simp]: "f v (|abrupt=None,store=s|)), = (|abrupt=None,store=s'|)), ==> assign f v (|abrupt=None,store=s|)), = (|abrupt=None,store=s'|))," by (simp add: assign_def Let_def) *) lemma assign_Norm_Norm [simp]: "f v (Norm s) = Norm s' ==> assign f v (Norm s) = Norm s'" by (simp add: assign_def Let_def) (* lemma assign_Norm_Some [simp]: "[|abrupt (f v (|abrupt=None,store=s|)),) = Some y|] ==> assign f v (|abrupt=None,store=s|)), = (|abrupt=Some y,store =s|))," by (simp add: assign_def Let_def split_beta) *) lemma assign_Norm_Some [simp]: "[|abrupt (f v (Norm s)) = Some y|] ==> assign f v (Norm s) = (Some y,s)" by (simp add: assign_def Let_def split_beta) lemma assign_Some [simp]: "assign f v (Some x,s) = (Some x,s)" by (simp add: assign_def Let_def split_beta) lemma assign_Some1 [simp]: "¬ normal s ==> assign f v s = s" by (auto simp add: assign_def Let_def split_beta) lemma assign_supd [simp]: "assign (λv. supd (f v)) v (x,s) = (x, if x = None then f v s else s)" apply auto done lemma assign_raise_if [simp]: "assign (λv (x,s). ((raise_if (b s v) xcpt) x, f v s)) v (x, s) = (raise_if (b s v) xcpt x, if x=None ∧ ¬b s v then f v s else s)" apply (case_tac "x = None") apply auto done (* lemma assign_raise_if [simp]: "assign (λv s. (|abrupt=(raise_if (b (store s) v) xcpt) (abrupt s), store = f v (store s)|)),) v s = (|abrupt=raise_if (b (store s) v) xcpt (abrupt s), store= if (abrupt s)=None ∧ ¬b (store s) v then f v (store s) else (store s)|))," apply (case_tac "abrupt s = None") apply auto done *) constdefs init_comp_ty :: "ty => stmt" "init_comp_ty T ≡ if (∃C. T = Class C) then Init (the_Class T) else Skip" lemma init_comp_ty_PrimT [simp]: "init_comp_ty (PrimT pt) = Skip" apply (unfold init_comp_ty_def) apply (simp (no_asm)) done constdefs invocation_class :: "inv_mode => st => val => ref_ty => qtname" "invocation_class m s a' statT ≡ (case m of Static => if (∃ statC. statT = ClassT statC) then the_Class (RefT statT) else Object | SuperM => the_Class (RefT statT) | IntVir => obj_class (lookup_obj s a'))" invocation_declclass::"prog => inv_mode => st => val => ref_ty => sig => qtname" "invocation_declclass G m s a' statT sig ≡ declclass (the (dynlookup G statT (invocation_class m s a' statT) sig))" lemma invocation_class_IntVir [simp]: "invocation_class IntVir s a' statT = obj_class (lookup_obj s a')" by (simp add: invocation_class_def) lemma dynclass_SuperM [simp]: "invocation_class SuperM s a' statT = the_Class (RefT statT)" by (simp add: invocation_class_def) lemma invocation_class_Static [simp]: "invocation_class Static s a' statT = (if (∃ statC. statT = ClassT statC) then the_Class (RefT statT) else Object)" by (simp add: invocation_class_def) constdefs init_lvars :: "prog => qtname => sig => inv_mode => val => val list => state => state" "init_lvars G C sig mode a' pvs ≡ λ (x,s). let m = mthd (the (methd G C sig)); l = λ k. (case k of EName e => (case e of VNam v => (empty ((pars m)[\<mapsto>]pvs)) v | Res => None) | This => (if mode=Static then None else Some a')) in set_lvars l (if mode = Static then x else np a' x,s)" lemma init_lvars_def2: --{* better suited for simplification *} "init_lvars G C sig mode a' pvs (x,s) = set_lvars (λ k. (case k of EName e => (case e of VNam v => (empty ((pars (mthd (the (methd G C sig))))[\<mapsto>]pvs)) v | Res => None) | This => (if mode=Static then None else Some a'))) (if mode = Static then x else np a' x,s)" apply (unfold init_lvars_def) apply (simp (no_asm) add: Let_def) done constdefs body :: "prog => qtname => sig => expr" "body G C sig ≡ let m = the (methd G C sig) in Body (declclass m) (stmt (mbody (mthd m)))" lemma body_def2: --{* better suited for simplification *} "body G C sig = Body (declclass (the (methd G C sig))) (stmt (mbody (mthd (the (methd G C sig)))))" apply (unfold body_def Let_def) apply auto done section "variables" constdefs lvar :: "lname => st => vvar" "lvar vn s ≡ (the (locals s vn), λv. supd (lupd(vn\<mapsto>v)))" fvar :: "qtname => bool => vname => val => state => vvar × state" "fvar C stat fn a' s ≡ let (oref,xf) = if stat then (Stat C,id) else (Heap (the_Addr a'),np a'); n = Inl (fn,C); f = (λv. supd (upd_gobj oref n v)) in ((the (values (the (globs (store s) oref)) n),f),abupd xf s)" avar :: "prog => val => val => state => vvar × state" "avar G i' a' s ≡ let oref = Heap (the_Addr a'); i = the_Intg i'; n = Inr i; (T,k,cs) = the_Arr (globs (store s) oref); f = (λv (x,s). (raise_if (¬G,s\<turnstile>v fits T) ArrStore x ,upd_gobj oref n v s)) in ((the (cs n),f) ,abupd (raise_if (¬i in_bounds k) IndOutBound o np a') s)" lemma fvar_def2: --{* better suited for simplification *} "fvar C stat fn a' s = ((the (values (the (globs (store s) (if stat then Stat C else Heap (the_Addr a')))) (Inl (fn,C))) ,(λv. supd (upd_gobj (if stat then Stat C else Heap (the_Addr a')) (Inl (fn,C)) v))) ,abupd (if stat then id else np a') s) " apply (unfold fvar_def) apply (simp (no_asm) add: Let_def split_beta) done lemma avar_def2: --{* better suited for simplification *} "avar G i' a' s = ((the ((snd(snd(the_Arr (globs (store s) (Heap (the_Addr a')))))) (Inr (the_Intg i'))) ,(λv (x,s'). (raise_if (¬G,s'\<turnstile>v fits (fst(the_Arr (globs (store s) (Heap (the_Addr a')))))) ArrStore x ,upd_gobj (Heap (the_Addr a')) (Inr (the_Intg i')) v s'))) ,abupd (raise_if (¬(the_Intg i') in_bounds (fst(snd(the_Arr (globs (store s) (Heap (the_Addr a'))))))) IndOutBound o np a') s)" apply (unfold avar_def) apply (simp (no_asm) add: Let_def split_beta) done constdefs check_field_access:: "prog => qtname => qtname => vname => bool => val => state => state" "check_field_access G accC statDeclC fn stat a' s ≡ let oref = if stat then Stat statDeclC else Heap (the_Addr a'); dynC = case oref of Heap a => obj_class (the (globs (store s) oref)) | Stat C => C; f = (the (table_of (DeclConcepts.fields G dynC) (fn,statDeclC))) in abupd (error_if (¬ G\<turnstile>Field fn (statDeclC,f) in dynC dyn_accessible_from accC) AccessViolation) s" constdefs check_method_access:: "prog => qtname => ref_ty => inv_mode => sig => val => state => state" "check_method_access G accC statT mode sig a' s ≡ let invC = invocation_class mode (store s) a' statT; dynM = the (dynlookup G statT invC sig) in abupd (error_if (¬ G\<turnstile>Methd sig dynM in invC dyn_accessible_from accC) AccessViolation) s" section "evaluation judgments" inductive halloc :: "[prog,state,obj_tag,loc,state]=>bool" ("_\<turnstile>_ \<midarrow>halloc _\<succ>_-> _"[61,61,61,61,61]60) for G::prog where --{* allocating objects on the heap, cf. 12.5 *} Abrupt: "G\<turnstile>(Some x,s) \<midarrow>halloc oi\<succ>arbitrary-> (Some x,s)" | New: "[|new_Addr (heap s) = Some a; (x,oi') = (if atleast_free (heap s) (Suc (Suc 0)) then (None,oi) else (Some (Xcpt (Loc a)),CInst (SXcpt OutOfMemory)))|] ==> G\<turnstile>Norm s \<midarrow>halloc oi\<succ>a-> (x,init_obj G oi' (Heap a) s)" inductive sxalloc :: "[prog,state,state]=>bool" ("_\<turnstile>_ \<midarrow>sxalloc-> _"[61,61,61]60) for G::prog where --{* allocating exception objects for standard exceptions (other than OutOfMemory) *} Norm: "G\<turnstile> Norm s \<midarrow>sxalloc-> Norm s" | Jmp: "G\<turnstile>(Some (Jump j), s) \<midarrow>sxalloc-> (Some (Jump j), s)" | Error: "G\<turnstile>(Some (Error e), s) \<midarrow>sxalloc-> (Some (Error e), s)" | XcptL: "G\<turnstile>(Some (Xcpt (Loc a) ),s) \<midarrow>sxalloc-> (Some (Xcpt (Loc a)),s)" | SXcpt: "[|G\<turnstile>Norm s0 \<midarrow>halloc (CInst (SXcpt xn))\<succ>a-> (x,s1)|] ==> G\<turnstile>(Some (Xcpt (Std xn)),s0) \<midarrow>sxalloc-> (Some (Xcpt (Loc a)),s1)" inductive eval :: "[prog,state,term,vals,state]=>bool" ("_\<turnstile>_ \<midarrow>_\<succ>-> '(_, _')" [61,61,80,0,0]60) and exec ::"[prog,state,stmt ,state]=>bool"("_\<turnstile>_ \<midarrow>_-> _" [61,61,65, 61]60) and evar ::"[prog,state,var ,vvar,state]=>bool"("_\<turnstile>_ \<midarrow>_=\<succ>_-> _"[61,61,90,61,61]60) and eval'::"[prog,state,expr ,val ,state]=>bool"("_\<turnstile>_ \<midarrow>_-\<succ>_-> _"[61,61,80,61,61]60) and evals::"[prog,state,expr list , val list ,state]=>bool"("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_-> _"[61,61,61,61,61]60) for G::prog where "G\<turnstile>s \<midarrow>c -> s' ≡ G\<turnstile>s \<midarrow>In1r c\<succ>-> (♦, s')" | "G\<turnstile>s \<midarrow>e-\<succ>v -> s' ≡ G\<turnstile>s \<midarrow>In1l e\<succ>-> (In1 v, s')" | "G\<turnstile>s \<midarrow>e=\<succ>vf-> s' ≡ G\<turnstile>s \<midarrow>In2 e\<succ>-> (In2 vf, s')" | "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v -> s' ≡ G\<turnstile>s \<midarrow>In3 e\<succ>-> (In3 v, s')" --{* propagation of abrupt completion *} --{* cf. 14.1, 15.5 *} | Abrupt: "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>-> (arbitrary3 t, (Some xc, s))" --{* execution of statements *} --{* cf. 14.5 *} | Skip: "G\<turnstile>Norm s \<midarrow>Skip-> Norm s" --{* cf. 14.7 *} | Expr: "[|G\<turnstile>Norm s0 \<midarrow>e-\<succ>v-> s1|] ==> G\<turnstile>Norm s0 \<midarrow>Expr e-> s1" | Lab: "[|G\<turnstile>Norm s0 \<midarrow>c -> s1|] ==> G\<turnstile>Norm s0 \<midarrow>l• c-> abupd (absorb l) s1" --{* cf. 14.2 *} | Comp: "[|G\<turnstile>Norm s0 \<midarrow>c1 -> s1; G\<turnstile> s1 \<midarrow>c2 -> s2|] ==> G\<turnstile>Norm s0 \<midarrow>c1;; c2-> s2" --{* cf. 14.8.2 *} | If: "[|G\<turnstile>Norm s0 \<midarrow>e-\<succ>b-> s1; G\<turnstile> s1\<midarrow>(if the_Bool b then c1 else c2)-> s2|] ==> G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 -> s2" --{* cf. 14.10, 14.10.1 *} --{* A continue jump from the while body @{term c} is handled by this rule. If a continue jump with the proper label was invoked inside @{term c} this label (Cont l) is deleted out of the abrupt component of the state before the iterative evaluation of the while statement. A break jump is handled by the Lab Statement @{text "Lab l (while…)"}. *} | Loop: "[|G\<turnstile>Norm s0 \<midarrow>e-\<succ>b-> s1; if the_Bool b then (G\<turnstile>s1 \<midarrow>c-> s2 ∧ G\<turnstile>(abupd (absorb (Cont l)) s2) \<midarrow>l• While(e) c-> s3) else s3 = s1|] ==> G\<turnstile>Norm s0 \<midarrow>l• While(e) c-> s3" | Jmp: "G\<turnstile>Norm s \<midarrow>Jmp j-> (Some (Jump j), s)" --{* cf. 14.16 *} | Throw: "[|G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'-> s1|] ==> G\<turnstile>Norm s0 \<midarrow>Throw e-> abupd (throw a') s1" --{* cf. 14.18.1 *} | Try: "[|G\<turnstile>Norm s0 \<midarrow>c1-> s1; G\<turnstile>s1 \<midarrow>sxalloc-> s2; if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2-> s3 else s3 = s2|] ==> G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(C vn) c2-> s3" --{* cf. 14.18.2 *} | Fin: "[|G\<turnstile>Norm s0 \<midarrow>c1-> (x1,s1); G\<turnstile>Norm s1 \<midarrow>c2-> s2; s3=(if (∃ err. x1=Some (Error err)) then (x1,s1) else abupd (abrupt_if (x1≠None) x1) s2) |] ==> G\<turnstile>Norm s0 \<midarrow>c1 Finally c2-> s3" --{* cf. 12.4.2, 8.5 *} | Init: "[|the (class G C) = c; if inited C (globs s0) then s3 = Norm s0 else (G\<turnstile>Norm (init_class_obj G C s0) \<midarrow>(if C = Object then Skip else Init (super c))-> s1 ∧ G\<turnstile>set_lvars empty s1 \<midarrow>init c-> s2 ∧ s3 = restore_lvars s1 s2)|] ==> G\<turnstile>Norm s0 \<midarrow>Init C-> s3" --{* This class initialisation rule is a little bit inaccurate. Look at the exact sequence: (1) The current class object (the static fields) are initialised (@{text init_class_obj}), (2) the superclasses are initialised, (3) the static initialiser of the current class is invoked. More precisely we should expect another ordering, namely 2 1 3. But we can't just naively toggle 1 and 2. By calling @{text init_class_obj} before initialising the superclasses, we also implicitly record that we have started to initialise the current class (by setting an value for the class object). This becomes crucial for the completeness proof of the axiomatic semantics @{text "AxCompl.thy"}. Static initialisation requires an induction on the number of classes not yet initialised (or to be more precise, classes were the initialisation has not yet begun). So we could first assign a dummy value to the class before superclass initialisation and afterwards set the correct values. But as long as we don't take memory overflow into account when allocating class objects, we can leave things as they are for convenience. *} --{* evaluation of expressions *} --{* cf. 15.8.1, 12.4.1 *} | NewC: "[|G\<turnstile>Norm s0 \<midarrow>Init C-> s1; G\<turnstile> s1 \<midarrow>halloc (CInst C)\<succ>a-> s2|] ==> G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a-> s2" --{* cf. 15.9.1, 12.4.1 *} | NewA: "[|G\<turnstile>Norm s0 \<midarrow>init_comp_ty T-> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'-> s2; G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a-> s3|] ==> G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a-> s3" --{* cf. 15.15 *} | Cast: "[|G\<turnstile>Norm s0 \<midarrow>e-\<succ>v-> s1; s2 = abupd (raise_if (¬G,store s1\<turnstile>v fits T) ClassCast) s1|] ==> G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v-> s2" --{* cf. 15.19.2 *} | Inst: "[|G\<turnstile>Norm s0 \<midarrow>e-\<succ>v-> s1; b = (v≠Null ∧ G,store s1\<turnstile>v fits RefT T)|] ==> G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b-> s1" --{* cf. 15.7.1 *} | Lit: "G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v-> Norm s" | UnOp: "[|G\<turnstile>Norm s0 \<midarrow>e-\<succ>v-> s1|] ==> G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>(eval_unop unop v)-> s1" | BinOp: "[|G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1-> s1; G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then (In1l e2) else (In1r Skip)) \<succ>-> (In1 v2, s2) |] ==> G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)-> s2" --{* cf. 15.10.2 *} | Super: "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s-> Norm s" --{* cf. 15.2 *} | Acc: "[|G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)-> s1|] ==> G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v-> s1" --{* cf. 15.25.1 *} | Ass: "[|G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)-> s1; G\<turnstile> s1 \<midarrow>e-\<succ>v -> s2|] ==> G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v-> assign f v s2" --{* cf. 15.24 *} | Cond: "[|G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b-> s1; G\<turnstile> s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v-> s2|] ==> G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v-> s2" -- {* The interplay of @{term Call}, @{term Methd} and @{term Body}: Method invocation is split up into these three rules: \begin{itemize} \item [@{term Call}] Calculates the target address and evaluates the arguments of the method, and then performs dynamic or static lookup of the method, corresponding to the call mode. Then the @{term Methd} rule is evaluated on the calculated declaration class of the method invocation. \item [@{term Methd}] A syntactic bridge for the folded method body. It is used by the axiomatic semantics to add the proper hypothesis for recursive calls of the method. \item [@{term Body}] An extra syntactic entity for the unfolded method body was introduced to properly trigger class initialisation. Without class initialisation we could just evaluate the body statement. \end{itemize} *} --{* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5 *} | Call: "[|G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'-> s1; G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs-> s2; D = invocation_declclass G mode (store s2) a' statT (|name=mn,parTs=pTs|)),; s3=init_lvars G D (|name=mn,parTs=pTs|)), mode a' vs s2; s3' = check_method_access G accC statT mode (|name=mn,parTs=pTs|)), a' s3; G\<turnstile>s3' \<midarrow>Methd D (|name=mn,parTs=pTs|)),-\<succ>v-> s4|] ==> G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e·mn({pTs}args)-\<succ>v-> (restore_lvars s2 s4)" --{* The accessibility check is after @{term init_lvars}, to keep it simple. @{term init_lvars} already tests for the absence of a null-pointer reference in case of an instance method invocation. *} | Methd: "[|G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v-> s1|] ==> G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v-> s1" | Body: "[|G\<turnstile>Norm s0 \<midarrow>Init D-> s1; G\<turnstile>s1 \<midarrow>c-> s2; s3 = (if (∃ l. abrupt s2 = Some (Jump (Break l)) ∨ abrupt s2 = Some (Jump (Cont l))) then abupd (λ x. Some (Error CrossMethodJump)) s2 else s2)|] ==> G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result) ->abupd (absorb Ret) s3" --{* cf. 14.15, 12.4.1 *} --{* We filter out a break/continue in @{term s2}, so that we can proof definite assignment correct, without the need of conformance of the state. By this the different parts of the typesafety proof can be disentangled a little. *} --{* evaluation of variables *} --{* cf. 15.13.1, 15.7.2 *} | LVar: "G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s-> Norm s" --{* cf. 15.10.1, 12.4.1 *} | FVar: "[|G\<turnstile>Norm s0 \<midarrow>Init statDeclC-> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a-> s2; (v,s2') = fvar statDeclC stat fn a s2; s3 = check_field_access G accC statDeclC fn stat a s2' |] ==> G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v-> s3" --{* The accessibility check is after @{term fvar}, to keep it simple. @{term fvar} already tests for the absence of a null-pointer reference in case of an instance field *} --{* cf. 15.12.1, 15.25.1 *} | AVar: "[|G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a-> s1; G\<turnstile>s1 \<midarrow>e2-\<succ>i-> s2; (v,s2') = avar G i a s2|] ==> G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v-> s2'" --{* evaluation of expression lists *} --{* cf. 15.11.4.2 *} | Nil: "G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]-> Norm s0" --{* cf. 15.6.4 *} | Cons: "[|G\<turnstile>Norm s0 \<midarrow>e -\<succ> v -> s1; G\<turnstile> s1 \<midarrow>es\<doteq>\<succ>vs-> s2|] ==> G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs-> s2" (* Rearrangement of premisses: [0,1(Abrupt),2(Skip),8(Jmp),4(Lab),30(Nil),31(Cons),27(LVar),17(Cast),18(Inst), 17(Lit),18(UnOp),19(BinOp),20(Super),21(Acc),3(Expr),5(Comp),25(Methd),26(Body),23(Cond),6(If), 7(Loop),11(Fin),9(Throw),13(NewC),14(NewA),12(Init),22(Ass),10(Try),28(FVar), 29(AVar),24(Call)] *) ML {* bind_thm ("eval_induct_", rearrange_prems [0,1,2,8,4,30,31,27,15,16, 17,18,19,20,21,3,5,25,26,23,6, 7,11,9,13,14,12,22,10,28, 29,24] @{thm eval.induct}) *} lemmas eval_induct = eval_induct_ [split_format and and and and and and and and and and and and and and s1 (* Acc *) and and s2 (* Comp *) and and and and and and s2 (* Fin *) and and s2 (* NewC *)] declare split_if [split del] split_if_asm [split del] option.split [split del] option.split_asm [split del] inductive_cases halloc_elim_cases: "G\<turnstile>(Some xc,s) \<midarrow>halloc oi\<succ>a-> s'" "G\<turnstile>(Norm s) \<midarrow>halloc oi\<succ>a-> s'" inductive_cases sxalloc_elim_cases: "G\<turnstile> Norm s \<midarrow>sxalloc-> s'" "G\<turnstile>(Some (Jump j),s) \<midarrow>sxalloc-> s'" "G\<turnstile>(Some (Error e),s) \<midarrow>sxalloc-> s'" "G\<turnstile>(Some (Xcpt (Loc a )),s) \<midarrow>sxalloc-> s'" "G\<turnstile>(Some (Xcpt (Std xn)),s) \<midarrow>sxalloc-> s'" inductive_cases sxalloc_cases: "G\<turnstile>s \<midarrow>sxalloc-> s'" lemma sxalloc_elim_cases2: "[|G\<turnstile>s \<midarrow>sxalloc-> s'; !!s. [|s' = Norm s|] ==> P; !!j s. [|s' = (Some (Jump j),s)|] ==> P; !!e s. [|s' = (Some (Error e),s)|] ==> P; !!a s. [|s' = (Some (Xcpt (Loc a)),s)|] ==> P |] ==> P" apply cut_tac apply (erule sxalloc_cases) apply blast+ done declare not_None_eq [simp del] (* IntDef.Zero_def [simp del] *) declare split_paired_All [simp del] split_paired_Ex [simp del] declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} inductive_cases eval_cases: "G\<turnstile>s \<midarrow>t\<succ>-> (v, s')" inductive_cases eval_elim_cases [cases set]: "G\<turnstile>(Some xc,s) \<midarrow>t \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1r Skip \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1r (Jmp j) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1r (l• c) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In3 ([]) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In3 (e#es) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l (Lit w) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l (UnOp unop e) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l (BinOp binop e1 e2) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In2 (LVar vn) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l (Cast T e) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l (e InstOf T) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l (Super) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l (Acc va) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1r (Expr e) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1r (c1;; c2) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1l (Methd C sig) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1l (Body D c) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1r (l• While(e) c) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1r (Throw e) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In1l (NewC C) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l (New T[e]) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l (Ass va e) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2) \<succ>-> (x, s')" "G\<turnstile>Norm s \<midarrow>In2 ({accC,statDeclC,stat}e..fn) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In2 (e1.[e2]) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1l ({accC,statT,mode}e·mn({pT}p)) \<succ>-> (v, s')" "G\<turnstile>Norm s \<midarrow>In1r (Init C) \<succ>-> (x, s')" declare not_None_eq [simp] (* IntDef.Zero_def [simp] *) declare split_paired_All [simp] split_paired_Ex [simp] declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} declare split_if [split] split_if_asm [split] option.split [split] option.split_asm [split] lemma eval_Inj_elim: "G\<turnstile>s \<midarrow>t\<succ>-> (w,s') ==> case t of In1 ec => (case ec of Inl e => (∃v. w = In1 v) | Inr c => w = ♦) | In2 e => (∃v. w = In2 v) | In3 e => (∃v. w = In3 v)" apply (erule eval_cases) apply auto apply (induct_tac "t") apply (induct_tac "a") apply auto done text {* The following simplification procedures set up the proper injections of terms and their corresponding values in the evaluation relation: E.g. an expression (injection @{term In1l} into terms) always evaluates to ordinary values (injection @{term In1} into generalised values @{term vals}). *} lemma eval_expr_eq: "G\<turnstile>s \<midarrow>In1l t\<succ>-> (w, s') = (∃v. w=In1 v ∧ G\<turnstile>s \<midarrow>t-\<succ>v -> s')" by (auto, frule eval_Inj_elim, auto) lemma eval_var_eq: "G\<turnstile>s \<midarrow>In2 t\<succ>-> (w, s') = (∃vf. w=In2 vf ∧ G\<turnstile>s \<midarrow>t=\<succ>vf-> s')" by (auto, frule eval_Inj_elim, auto) lemma eval_exprs_eq: "G\<turnstile>s \<midarrow>In3 t\<succ>-> (w, s') = (∃vs. w=In3 vs ∧ G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs-> s')" by (auto, frule eval_Inj_elim, auto) lemma eval_stmt_eq: "G\<turnstile>s \<midarrow>In1r t\<succ>-> (w, s') = (w=♦ ∧ G\<turnstile>s \<midarrow>t -> s')" by (auto, frule eval_Inj_elim, auto, frule eval_Inj_elim, auto) simproc_setup eval_expr ("G\<turnstile>s \<midarrow>In1l t\<succ>-> (w, s')") = {* fn _ => fn _ => fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE | _ => SOME (mk_meta_eq @{thm eval_expr_eq})) *} simproc_setup eval_var ("G\<turnstile>s \<midarrow>In2 t\<succ>-> (w, s')") = {* fn _ => fn _ => fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE | _ => SOME (mk_meta_eq @{thm eval_var_eq})) *} simproc_setup eval_exprs ("G\<turnstile>s \<midarrow>In3 t\<succ>-> (w, s')") = {* fn _ => fn _ => fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE | _ => SOME (mk_meta_eq @{thm eval_exprs_eq})) *} simproc_setup eval_stmt ("G\<turnstile>s \<midarrow>In1r t\<succ>-> (w, s')") = {* fn _ => fn _ => fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE | _ => SOME (mk_meta_eq @{thm eval_stmt_eq})) *} ML {* bind_thms ("AbruptIs", sum3_instantiate @{thm eval.Abrupt}) *} declare halloc.Abrupt [intro!] eval.Abrupt [intro!] AbruptIs [intro!] text{* @{text Callee},@{text InsInitE}, @{text InsInitV}, @{text FinA} are only used in smallstep semantics, not in the bigstep semantics. So their is no valid evaluation of these terms *} lemma eval_Callee: "G\<turnstile>Norm s\<midarrow>Callee l e-\<succ>v-> s' = False" proof - { fix s t v s' assume eval: "G\<turnstile>s \<midarrow>t\<succ>-> (v,s')" and normal: "normal s" and callee: "t=In1l (Callee l e)" then have "False" by induct auto } then show ?thesis by (cases s') fastsimp qed lemma eval_InsInitE: "G\<turnstile>Norm s\<midarrow>InsInitE c e-\<succ>v-> s' = False" proof - { fix s t v s' assume eval: "G\<turnstile>s \<midarrow>t\<succ>-> (v,s')" and normal: "normal s" and callee: "t=In1l (InsInitE c e)" then have "False" by induct auto } then show ?thesis by (cases s') fastsimp qed lemma eval_InsInitV: "G\<turnstile>Norm s\<midarrow>InsInitV c w=\<succ>v-> s' = False" proof - { fix s t v s' assume eval: "G\<turnstile>s \<midarrow>t\<succ>-> (v,s')" and normal: "normal s" and callee: "t=In2 (InsInitV c w)" then have "False" by induct auto } then show ?thesis by (cases s') fastsimp qed lemma eval_FinA: "G\<turnstile>Norm s\<midarrow>FinA a c-> s' = False" proof - { fix s t v s' assume eval: "G\<turnstile>s \<midarrow>t\<succ>-> (v,s')" and normal: "normal s" and callee: "t=In1r (FinA a c)" then have "False" by induct auto } then show ?thesis by (cases s') fastsimp qed lemma eval_no_abrupt_lemma: "!!s s'. G\<turnstile>s \<midarrow>t\<succ>-> (w,s') ==> normal s' --> normal s" by (erule eval_cases, auto) lemma eval_no_abrupt: "G\<turnstile>(x,s) \<midarrow>t\<succ>-> (w,Norm s') = (x = None ∧ G\<turnstile>Norm s \<midarrow>t\<succ>-> (w,Norm s'))" apply auto apply (frule eval_no_abrupt_lemma, auto)+ done simproc_setup eval_no_abrupt ("G\<turnstile>(x,s) \<midarrow>e\<succ>-> (w,Norm s')") = {* fn _ => fn _ => fn ct => (case Thm.term_of ct of (_ $ _ $ (Const (@{const_name Pair}, _) $ (Const (@{const_name None}, _)) $ _) $ _ $ _ $ _) => NONE | _ => SOME (mk_meta_eq @{thm eval_no_abrupt})) *} lemma eval_abrupt_lemma: "G\<turnstile>s \<midarrow>t\<succ>-> (v,s') ==> abrupt s=Some xc --> s'= s ∧ v = arbitrary3 t" by (erule eval_cases, auto) lemma eval_abrupt: " G\<turnstile>(Some xc,s) \<midarrow>t\<succ>-> (w,s') = (s'=(Some xc,s) ∧ w=arbitrary3 t ∧ G\<turnstile>(Some xc,s) \<midarrow>t\<succ>-> (arbitrary3 t,(Some xc,s)))" apply auto apply (frule eval_abrupt_lemma, auto)+ done simproc_setup eval_abrupt ("G\<turnstile>(Some xc,s) \<midarrow>e\<succ>-> (w,s')") = {* fn _ => fn _ => fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ _ $ (Const (@{const_name Pair}, _) $ (Const (@{const_name Some}, _) $ _)$ _)) => NONE | _ => SOME (mk_meta_eq @{thm eval_abrupt})) *} lemma LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else arbitrary)-> s" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Lit) lemma SkipI [intro!]: "G\<turnstile>s \<midarrow>Skip-> s" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Skip) lemma ExprI: "G\<turnstile>s \<midarrow>e-\<succ>v-> s' ==> G\<turnstile>s \<midarrow>Expr e-> s'" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Expr) lemma CompI: "[|G\<turnstile>s \<midarrow>c1-> s1; G\<turnstile>s1 \<midarrow>c2-> s2|] ==> G\<turnstile>s \<midarrow>c1;; c2-> s2" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Comp) lemma CondI: "!!s1. [|G\<turnstile>s \<midarrow>e-\<succ>b-> s1; G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v-> s2|] ==> G\<turnstile>s \<midarrow>e ? e1 : e2-\<succ>(if normal s1 then v else arbitrary)-> s2" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Cond) lemma IfI: "[|G\<turnstile>s \<midarrow>e-\<succ>v-> s1; G\<turnstile>s1 \<midarrow>(if the_Bool v then c1 else c2)-> s2|] ==> G\<turnstile>s \<midarrow>If(e) c1 Else c2-> s2" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.If) lemma MethdI: "G\<turnstile>s \<midarrow>body G C sig-\<succ>v-> s' ==> G\<turnstile>s \<midarrow>Methd C sig-\<succ>v-> s'" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Methd) lemma eval_Call: "[|G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'-> s1; G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>pvs-> s2; D = invocation_declclass G mode (store s2) a' statT (|name=mn,parTs=pTs|)),; s3 = init_lvars G D (|name=mn,parTs=pTs|)), mode a' pvs s2; s3' = check_method_access G accC statT mode (|name=mn,parTs=pTs|)), a' s3; G\<turnstile>s3'\<midarrow>Methd D (|name=mn,parTs=pTs|)),-\<succ> v-> s4; s4' = restore_lvars s2 s4|] ==> G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e·mn({pTs}ps)-\<succ>v-> s4'" apply (drule eval.Call, assumption) apply (rule HOL.refl) apply simp+ done lemma eval_Init: "[|if inited C (globs s0) then s3 = Norm s0 else G\<turnstile>Norm (init_class_obj G C s0) \<midarrow>(if C = Object then Skip else Init (super (the (class G C))))-> s1 ∧ G\<turnstile>set_lvars empty s1 \<midarrow>(init (the (class G C)))-> s2 ∧ s3 = restore_lvars s1 s2|] ==> G\<turnstile>Norm s0 \<midarrow>Init C-> s3" apply (rule eval.Init) apply auto done lemma init_done: "initd C s ==> G\<turnstile>s \<midarrow>Init C-> s" apply (case_tac "s", simp) apply (case_tac "a") apply safe apply (rule eval_Init) apply auto done lemma eval_StatRef: "G\<turnstile>s \<midarrow>StatRef rt-\<succ>(if abrupt s=None then Null else arbitrary)-> s" apply (case_tac "s", simp) apply (case_tac "a = None") apply (auto del: eval.Abrupt intro!: eval.intros) done lemma SkipD [dest!]: "G\<turnstile>s \<midarrow>Skip-> s' ==> s' = s" apply (erule eval_cases) by auto lemma Skip_eq [simp]: "G\<turnstile>s \<midarrow>Skip-> s' = (s = s')" by auto (*unused*) lemma init_retains_locals [rule_format (no_asm)]: "G\<turnstile>s \<midarrow>t\<succ>-> (w,s') ==> (∀C. t=In1r (Init C) --> locals (store s) = locals (store s'))" apply (erule eval.induct) apply (simp (no_asm_use) split del: split_if_asm option.split_asm)+ apply auto done lemma halloc_xcpt [dest!]: "!!s'. G\<turnstile>(Some xc,s) \<midarrow>halloc oi\<succ>a-> s' ==> s'=(Some xc,s)" apply (erule_tac halloc_elim_cases) by auto (* G\<turnstile>(x,(h,l)) \<midarrow>e\<succ>v-> (x',(h',l'))) ==> l This = l' This" G\<turnstile>(x,(h,l)) \<midarrow>s -> (x',(h',l'))) ==> l This = l' This" *) lemma eval_Methd: "G\<turnstile>s \<midarrow>In1l(body G C sig)\<succ>-> (w,s') ==> G\<turnstile>s \<midarrow>In1l(Methd C sig)\<succ>-> (w,s')" apply (case_tac "s") apply (case_tac "a") apply clarsimp+ apply (erule eval.Methd) apply (drule eval_abrupt_lemma) apply force done lemma eval_Body: "[|G\<turnstile>Norm s0 \<midarrow>Init D-> s1; G\<turnstile>s1 \<midarrow>c-> s2; res=the (locals (store s2) Result); s3 = (if (∃ l. abrupt s2 = Some (Jump (Break l)) ∨ abrupt s2 = Some (Jump (Cont l))) then abupd (λ x. Some (Error CrossMethodJump)) s2 else s2); s4=abupd (absorb Ret) s3|] ==> G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>res->s4" by (auto elim: eval.Body) lemma eval_binop_arg2_indep: "¬ need_second_arg binop v1 ==> eval_binop binop v1 x = eval_binop binop v1 y" by (cases binop) (simp_all add: need_second_arg_def) lemma eval_BinOp_arg2_indepI: assumes eval_e1: "G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1-> s1" and no_need: "¬ need_second_arg binop v1" shows "G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)-> s1" (is "?EvalBinOp v2") proof - from eval_e1 have "?EvalBinOp Unit" by (rule eval.BinOp) (simp add: no_need) moreover from no_need have "eval_binop binop v1 Unit = eval_binop binop v1 v2" by (simp add: eval_binop_arg2_indep) ultimately show ?thesis by simp qed section "single valued" lemma unique_halloc [rule_format (no_asm)]: "G\<turnstile>s \<midarrow>halloc oi\<succ>a -> s' ==> G\<turnstile>s \<midarrow>halloc oi\<succ>a' -> s'' --> a' = a ∧ s'' = s'" apply (erule halloc.induct) apply (auto elim!: halloc_elim_cases split del: split_if split_if_asm) apply (drule trans [THEN sym], erule sym) defer apply (drule trans [THEN sym], erule sym) apply auto done lemma single_valued_halloc: "single_valued {((s,oi),(a,s')). G\<turnstile>s \<midarrow>halloc oi\<succ>a -> s'}" apply (unfold single_valued_def) by (clarsimp, drule (1) unique_halloc, auto) lemma unique_sxalloc [rule_format (no_asm)]: "G\<turnstile>s \<midarrow>sxalloc-> s' ==> G\<turnstile>s \<midarrow>sxalloc-> s'' --> s'' = s'" apply (erule sxalloc.induct) apply (auto dest: unique_halloc elim!: sxalloc_elim_cases split del: split_if split_if_asm) done lemma single_valued_sxalloc: "single_valued {(s,s'). G\<turnstile>s \<midarrow>sxalloc-> s'}" apply (unfold single_valued_def) apply (blast dest: unique_sxalloc) done lemma split_pairD: "(x,y) = p ==> x = fst p & y = snd p" by auto lemma unique_eval [rule_format (no_asm)]: "G\<turnstile>s \<midarrow>t\<succ>-> (w, s') ==> (∀w' s''. G\<turnstile>s \<midarrow>t\<succ>-> (w', s'') --> w' = w ∧ s'' = s')" apply (erule eval_induct) apply (tactic {* ALLGOALS (EVERY' [strip_tac, rotate_tac ~1, eresolve_tac (thms "eval_elim_cases")]) *}) (* 31 subgoals *) prefer 28 (* Try *) apply (simp (no_asm_use) only: split add: split_if_asm) (* 34 subgoals *) prefer 30 (* Init *) apply (case_tac "inited C (globs s0)", (simp only: if_True if_False simp_thms)+) prefer 26 (* While *) apply (simp (no_asm_use) only: split add: split_if_asm, blast) (* 36 subgoals *) apply (blast dest: unique_sxalloc unique_halloc split_pairD)+ done (* unused *) lemma single_valued_eval: "single_valued {((s, t), (v, s')). G\<turnstile>s \<midarrow>t\<succ>-> (v, s')}" apply (unfold single_valued_def) by (clarify, drule (1) unique_eval, auto) end
lemma
arbitrary3 (In1l x) = In1 arbitrary
lemma
arbitrary3 (In1r x) = dummy_res
lemma
arbitrary3 (In2 x) = In2 arbitrary
lemma
arbitrary3 (In3 x) = In3 arbitrary
lemma throw_def2:
throw a' x = abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) ((np a') x)
lemma fits_Null:
G,s\<turnstile>Null fits T
lemma fits_Addr_RefT:
G,s\<turnstile>Addr a fits RefT t =
G\<turnstile>obj_ty (the (heap s a))\<preceq>RefT t
lemma fitsD:
G,s\<turnstile>a' fits T
==> (∃pt. T = PrimT pt) ∨
(∃t. T = RefT t) ∧ a' = Null ∨
(∃t. T = RefT t) ∧
a' ≠ Null ∧ G\<turnstile>obj_ty (lookup_obj s a')\<preceq>T
lemma catch_Norm:
¬ G,Norm s\<turnstile>catch tn
lemma catch_XcptLoc:
G,(Some (Xcpt (Loc a)), s)\<turnstile>catch C =
G,s\<turnstile>Addr a fits Class C
lemma catch_Jump:
¬ G,(Some (Jump j), s)\<turnstile>catch tn
lemma catch_Error:
¬ G,(Some (Error e), s)\<turnstile>catch tn
lemma new_xcpt_var_def2:
new_xcpt_var vn (x, s) =
Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)
lemma assign_Norm_Norm:
f v (Norm s) = Norm s' ==> assign f v (Norm s) = Norm s'
lemma assign_Norm_Some:
fst (f v (Norm s)) = Some y ==> assign f v (Norm s) = (Some y, s)
lemma assign_Some:
assign f v (Some x, s) = (Some x, s)
lemma assign_Some1:
¬ normal s ==> assign f v s = s
lemma assign_supd:
assign (λv. supd (f v)) v (x, s) = (x, if x = None then f v s else s)
lemma assign_raise_if:
assign (λv (x, s). ((raise_if (b s v) xcpt) x, f v s)) v (x, s) =
((raise_if (b s v) xcpt) x, if x = None ∧ ¬ b s v then f v s else s)
lemma init_comp_ty_PrimT:
init_comp_ty (PrimT pt) = Skip
lemma invocation_class_IntVir:
invocation_class IntVir s a' statT = obj_class (lookup_obj s a')
lemma dynclass_SuperM:
invocation_class SuperM s a' statT = the_Class (RefT statT)
lemma invocation_class_Static:
invocation_class Static s a' statT =
(if ∃statC. statT = ClassT statC then the_Class (RefT statT) else Object)
lemma init_lvars_def2:
init_lvars G C sig mode a' pvs (x, s) =
(set_lvars
(lname_case (ename_case [pars (mthd (the (methd G C sig))) [|->] pvs] None)
(if mode = Static then None else Some a')))
(if mode = Static then x else (np a') x, s)
lemma body_def2:
body G C sig =
Body (declclass (the (methd G C sig)))
(stmt (mbody (mthd (the (methd G C sig)))))
lemma fvar_def2:
fvar C stat fn a' s =
((the (values (the (globs (snd s) (if stat then Inr C else Inl (the_Addr a'))))
(Inl (fn, C))),
λv. supd (upd_gobj (if stat then Inr C else Inl (the_Addr a')) (Inl (fn, C))
v)),
abupd (if stat then id else np a') s)
lemma avar_def2:
avar G i' a' s =
((the (snd (snd (the_Arr (globs (snd s) (Inl (the_Addr a')))))
(Inr (the_Intg i'))),
λv (x, s').
((raise_if
(¬ G,s'\<turnstile>v fits fst (the_Arr
(globs (snd s) (Inl (the_Addr a')))))
ArrStore)
x,
upd_gobj (Inl (the_Addr a')) (Inr (the_Intg i')) v s')),
abupd
(raise_if
(¬ the_Intg i'
in_bounds fst (snd (the_Arr (globs (snd s) (Inl (the_Addr a'))))))
IndOutBound o
np a')
s)
lemma eval_induct:
[| G\<turnstile>x1.0 \<midarrow>x2.0\<succ>-> (x3.0, x4.0);
!!xc s t. P (Some xc, s) t (arbitrary3 t) (Some xc, s);
!!s. P (Norm s) (In1r Skip) dummy_res (Norm s);
!!s j. P (Norm s) (In1r (Jmp j)) dummy_res (Some (Jump j), s);
!!s0 c s1 l.
[| G\<turnstile>Norm s0 \<midarrow>c-> s1;
P (Norm s0) (In1r c) dummy_res s1 |]
==> P (Norm s0) (In1r (l• c)) dummy_res (abupd (absorb l) s1);
!!s0. P (Norm s0) (In3 []) (In3 []) (Norm s0);
!!s0 e v s1 es vs s2.
[| G\<turnstile>Norm s0 \<midarrow>e-\<succ>v-> s1;
P (Norm s0) (In1l e) (In1 v) s1;
G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs-> s2;
P s1 (In3 es) (In3 vs) s2 |]
==> P (Norm s0) (In3 (e # es)) (In3 (v # vs)) s2;
!!s vn. P (Norm s) (In2 (LVar vn)) (In2 (lvar vn s)) (Norm s);
!!s0 e v s1 s2 T.
[| G\<turnstile>Norm s0 \<midarrow>e-\<succ>v-> s1;
P (Norm s0) (In1l e) (In1 v) s1;
s2 = abupd (raise_if (¬ G,snd s1\<turnstile>v fits T) ClassCast) s1 |]
==> P (Norm s0) (In1l (Cast T e)) (In1 v) s2;
!!s0 e v s1 b T.
[| G\<turnstile>Norm s0 \<midarrow>e-\<succ>v-> s1;
P (Norm s0) (In1l e) (In1 v) s1;
b = (v ≠ Null ∧ G,snd s1\<turnstile>v fits RefT T) |]
==> P (Norm s0) (In1l (e InstOf T)) (In1 (Bool b)) s1;
!!s v. P (Norm s) (In1l (Lit v)) (In1 v) (Norm s);
!!s0 e v s1 unop.
[| G\<turnstile>Norm s0 \<midarrow>e-\<succ>v-> s1;
P (Norm s0) (In1l e) (In1 v) s1 |]
==> P (Norm s0) (In1l (UnOp unop e)) (In1 (eval_unop unop v)) s1;
!!s0 e1 v1 s1 binop e2 v2 s2.
[| G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1-> s1;
P (Norm s0) (In1l e1) (In1 v1) s1;
G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
else In1r Skip)\<succ>-> (In1 v2, s2);
P s1 (if need_second_arg binop v1 then In1l e2 else In1r Skip) (In1 v2)
s2 |]
==> P (Norm s0) (In1l (BinOp binop e1 e2)) (In1 (eval_binop binop v1 v2))
s2;
!!s. P (Norm s) (In1l Super) (In1 (val_this s)) (Norm s);
!!s0 va v f x y.
[| G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v, f)-> (x, y);
P (Norm s0) (In2 va) (In2 (v, f)) (x, y) |]
==> P (Norm s0) (In1l (Acc va)) (In1 v) (x, y);
!!s0 e v s1.
[| G\<turnstile>Norm s0 \<midarrow>e-\<succ>v-> s1;
P (Norm s0) (In1l e) (In1 v) s1 |]
==> P (Norm s0) (In1r (Expr e)) dummy_res s1;
!!s0 c1 s1 c2 x y.
[| G\<turnstile>Norm s0 \<midarrow>c1-> s1;
P (Norm s0) (In1r c1) dummy_res s1;
G\<turnstile>s1 \<midarrow>c2-> (x, y);
P s1 (In1r c2) dummy_res (x, y) |]
==> P (Norm s0) (In1r (c1;; c2)) dummy_res (x, y);
!!s0 D sig v s1.
[| G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v-> s1;
P (Norm s0) (In1l (body G D sig)) (In1 v) s1 |]
==> P (Norm s0) (In1l (Methd D sig)) (In1 v) s1;
!!s0 D s1 c s2 s3.
[| G\<turnstile>Norm s0 \<midarrow>Init D-> s1;
P (Norm s0) (In1r (Init D)) dummy_res s1;
G\<turnstile>s1 \<midarrow>c-> s2; P s1 (In1r c) dummy_res s2;
s3 =
(if ∃l. fst s2 = Some (Jump (Break l)) ∨ fst s2 = Some (Jump (Cont l))
then abupd (λx. Some (Error CrossMethodJump)) s2 else s2) |]
==> P (Norm s0) (In1l (Body D c)) (In1 (the (locals (snd s2) Result)))
(abupd (absorb Ret) s3);
!!s0 e0 b s1 e1 e2 v s2.
[| G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b-> s1;
P (Norm s0) (In1l e0) (In1 b) s1;
G\<turnstile>s1 \<midarrow>(if the_Bool b then e1
else e2)-\<succ>v-> s2;
P s1 (In1l (if the_Bool b then e1 else e2)) (In1 v) s2 |]
==> P (Norm s0) (In1l (e0 ? e1 : e2)) (In1 v) s2;
!!s0 e b s1 c1 c2 s2.
[| G\<turnstile>Norm s0 \<midarrow>e-\<succ>b-> s1;
P (Norm s0) (In1l e) (In1 b) s1;
G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)-> s2;
P s1 (In1r (if the_Bool b then c1 else c2)) dummy_res s2 |]
==> P (Norm s0) (In1r (If(e) c1 Else c2)) dummy_res s2;
!!s0 e b s1 c s2 l s3.
[| G\<turnstile>Norm s0 \<midarrow>e-\<succ>b-> s1;
P (Norm s0) (In1l e) (In1 b) s1;
if the_Bool b
then (G\<turnstile>s1 \<midarrow>c-> s2 ∧ P s1 (In1r c) dummy_res s2) ∧
G\<turnstile>abupd (absorb (Cont l))
s2 \<midarrow>l• While(e) c-> s3 ∧
P (abupd (absorb (Cont l)) s2) (In1r (l• While(e) c)) dummy_res s3
else s3 = s1 |]
==> P (Norm s0) (In1r (l• While(e) c)) dummy_res s3;
!!s0 c1 x1 s1 c2 x y.
[| G\<turnstile>Norm s0 \<midarrow>c1-> (x1, s1);
P (Norm s0) (In1r c1) dummy_res (x1, s1);
G\<turnstile>Norm s1 \<midarrow>c2-> (x, y);
P (Norm s1) (In1r c2) dummy_res (x, y) |]
==> P (Norm s0) (In1r (c1 Finally c2)) dummy_res
(if ∃err. x1 = Some (Error err) then (x1, s1)
else abupd (abrupt_if (x1 ≠ None) x1) (x, y));
!!s0 e a' s1.
[| G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'-> s1;
P (Norm s0) (In1l e) (In1 a') s1 |]
==> P (Norm s0) (In1r (Throw e)) dummy_res (abupd (throw a') s1);
!!s0 C s1 a x y.
[| G\<turnstile>Norm s0 \<midarrow>Init C-> s1;
P (Norm s0) (In1r (Init C)) dummy_res s1;
G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a-> (x, y) |]
==> P (Norm s0) (In1l (NewC C)) (In1 (Addr a)) (x, y);
!!s0 T s1 e i' s2 a s3.
[| G\<turnstile>Norm s0 \<midarrow>init_comp_ty T-> s1;
P (Norm s0) (In1r (init_comp_ty T)) dummy_res s1;
G\<turnstile>s1 \<midarrow>e-\<succ>i'-> s2; P s1 (In1l e) (In1 i') s2;
G\<turnstile>abupd (check_neg i')
s2 \<midarrow>halloc Arr T (the_Intg i')\<succ>a-> s3 |]
==> P (Norm s0) (In1l (New T[e])) (In1 (Addr a)) s3;
!!C c s0 s3 s1 s2.
[| the (class G C) = c;
if inited C (globs s0) then s3 = Norm s0
else (G\<turnstile>Norm ((init_class_obj G C)
s0) \<midarrow>(if C = Object then Skip
else Init (super c))-> s1 ∧
P (Norm ((init_class_obj G C) s0))
(In1r (if C = Object then Skip else Init (super c))) dummy_res
s1) ∧
(G\<turnstile>(set_lvars empty) s1 \<midarrow>init c-> s2 ∧
P ((set_lvars empty) s1) (In1r (init c)) dummy_res s2) ∧
s3 = (set_lvars (locals (snd s1))) s2 |]
==> P (Norm s0) (In1r (Init C)) dummy_res s3;
!!s0 va w f s1 e v s2.
[| G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w, f)-> s1;
P (Norm s0) (In2 va) (In2 (w, f)) s1;
G\<turnstile>s1 \<midarrow>e-\<succ>v-> s2; P s1 (In1l e) (In1 v) s2 |]
==> P (Norm s0) (In1l (va:=e)) (In1 v) (assign f v s2);
!!s0 c1 s1 s2 C vn c2 s3.
[| G\<turnstile>Norm s0 \<midarrow>c1-> s1;
P (Norm s0) (In1r c1) dummy_res s1;
G\<turnstile>s1 \<midarrow>sxalloc-> s2;
if G,s2\<turnstile>catch C
then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2-> s3 ∧
P (new_xcpt_var vn s2) (In1r c2) dummy_res s3
else s3 = s2 |]
==> P (Norm s0) (In1r (Try c1 Catch(C vn) c2)) dummy_res s3;
!!s0 statDeclC s1 e a s2 v s2' stat fn s3 accC.
[| G\<turnstile>Norm s0 \<midarrow>Init statDeclC-> s1;
P (Norm s0) (In1r (Init statDeclC)) dummy_res s1;
G\<turnstile>s1 \<midarrow>e-\<succ>a-> s2; P s1 (In1l e) (In1 a) s2;
(v, s2') = fvar statDeclC stat fn a s2;
s3 = check_field_access G accC statDeclC fn stat a s2' |]
==> P (Norm s0) (In2 ({accC,statDeclC,stat}e..fn)) (In2 v) s3;
!!s0 e1 a s1 e2 i s2 v s2'.
[| G\<turnstile>Norm s0 \<midarrow>e1-\<succ>a-> s1;
P (Norm s0) (In1l e1) (In1 a) s1;
G\<turnstile>s1 \<midarrow>e2-\<succ>i-> s2; P s1 (In1l e2) (In1 i) s2;
(v, s2') = avar G i a s2 |]
==> P (Norm s0) (In2 (e1.[e2])) (In2 v) s2';
!!s0 e a' s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4.
[| G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'-> s1;
P (Norm s0) (In1l e) (In1 a') s1;
G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs-> s2;
P s1 (In3 args) (In3 vs) s2;
D = invocation_declclass G mode (snd s2) a' statT
(| name = mn, parTs = pTs |);
s3 = init_lvars G D (| name = mn, parTs = pTs |) mode a' vs s2;
s3' =
check_method_access G accC statT mode (| name = mn, parTs = pTs |) a'
s3;
G\<turnstile>s3' \<midarrow>Methd D
(| name = mn,
parTs = pTs |)-\<succ>v-> s4;
P s3' (In1l (Methd D (| name = mn, parTs = pTs |))) (In1 v) s4 |]
==> P (Norm s0) (In1l ({accC,statT,mode}e·mn( {pTs}args))) (In1 v)
((set_lvars (locals (snd s2))) s4) |]
==> P x1.0 x2.0 x3.0 x4.0
lemma sxalloc_elim_cases2:
[| G\<turnstile>s \<midarrow>sxalloc-> s'; !!s. s' = Norm s ==> P;
!!j s. s' = (Some (Jump j), s) ==> P; !!e s. s' = (Some (Error e), s) ==> P;
!!a s. s' = (Some (Xcpt (Loc a)), s) ==> P |]
==> P
lemma eval_Inj_elim:
G\<turnstile>s \<midarrow>t\<succ>-> (w, s')
==> case t of In1l e => ∃v. w = In1 v | In1r c => w = dummy_res
| In2 e => ∃v. w = In2 v | In3 e => ∃v. w = In3 v
lemma eval_expr_eq:
G\<turnstile>s \<midarrow>In1l t\<succ>-> (w, s') =
(∃v. w = In1 v ∧ G\<turnstile>s \<midarrow>t-\<succ>v-> s')
lemma eval_var_eq:
G\<turnstile>s \<midarrow>In2 t\<succ>-> (w, s') =
(∃vf. w = In2 vf ∧ G\<turnstile>s \<midarrow>t=\<succ>vf-> s')
lemma eval_exprs_eq:
G\<turnstile>s \<midarrow>In3 t\<succ>-> (w, s') =
(∃vs. w = In3 vs ∧ G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs-> s')
lemma eval_stmt_eq:
G\<turnstile>s \<midarrow>In1r t\<succ>-> (w, s') =
(w = dummy_res ∧ G\<turnstile>s \<midarrow>t-> s')
lemma eval_Callee:
G\<turnstile>Norm s \<midarrow>Callee l e-\<succ>v-> s' = False
lemma eval_InsInitE:
G\<turnstile>Norm s \<midarrow>InsInitE c e-\<succ>v-> s' = False
lemma eval_InsInitV:
G\<turnstile>Norm s \<midarrow>InsInitV c w=\<succ>v-> s' = False
lemma eval_FinA:
G\<turnstile>Norm s \<midarrow>FinA a c-> s' = False
lemma eval_no_abrupt_lemma:
G\<turnstile>s \<midarrow>t\<succ>-> (w, s') ==> normal s' --> normal s
lemma eval_no_abrupt:
G\<turnstile>(x, s) \<midarrow>t\<succ>-> (w, Norm s') =
(x = None ∧ G\<turnstile>Norm s \<midarrow>t\<succ>-> (w, Norm s'))
lemma eval_abrupt_lemma:
G\<turnstile>s \<midarrow>t\<succ>-> (v, s')
==> fst s = Some xc --> s' = s ∧ v = arbitrary3 t
lemma eval_abrupt:
G\<turnstile>(Some xc, s) \<midarrow>t\<succ>-> (w, s') =
(s' = (Some xc, s) ∧
w = arbitrary3 t ∧
G\<turnstile>(Some xc, s) \<midarrow>t\<succ>-> (arbitrary3 t, (Some xc, s)))
lemma LitI:
G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else arbitrary)-> s
lemma SkipI:
G\<turnstile>s \<midarrow>Skip-> s
lemma ExprI:
G\<turnstile>s \<midarrow>e-\<succ>v-> s'
==> G\<turnstile>s \<midarrow>Expr e-> s'
lemma CompI:
[| G\<turnstile>s \<midarrow>c1.0-> s1.0;
G\<turnstile>s1.0 \<midarrow>c2.0-> s2.0 |]
==> G\<turnstile>s \<midarrow>c1.0;; c2.0-> s2.0
lemma CondI:
[| G\<turnstile>s \<midarrow>e-\<succ>b-> s1.0;
G\<turnstile>s1.0 \<midarrow>(if the_Bool b then e1.0
else e2.0)-\<succ>v-> s2.0 |]
==> G\<turnstile>s \<midarrow>e ? e1.0 : e2.0-\<succ>(if normal s1.0 then v
else arbitrary)-> s2.0
lemma IfI:
[| G\<turnstile>s \<midarrow>e-\<succ>v-> s1.0;
G\<turnstile>s1.0 \<midarrow>(if the_Bool v then c1.0 else c2.0)-> s2.0 |]
==> G\<turnstile>s \<midarrow>If(e) c1.0 Else c2.0-> s2.0
lemma MethdI:
G\<turnstile>s \<midarrow>body G C sig-\<succ>v-> s'
==> G\<turnstile>s \<midarrow>Methd C sig-\<succ>v-> s'
lemma eval_Call:
[| G\<turnstile>Norm s0.0 \<midarrow>e-\<succ>a'-> s1.0;
G\<turnstile>s1.0 \<midarrow>ps\<doteq>\<succ>pvs-> s2.0;
D = invocation_declclass G mode (snd s2.0) a' statT
(| name = mn, parTs = pTs |);
s3.0 = init_lvars G D (| name = mn, parTs = pTs |) mode a' pvs s2.0;
s3' =
check_method_access G accC statT mode (| name = mn, parTs = pTs |) a' s3.0;
G\<turnstile>s3' \<midarrow>Methd D
(| name = mn, parTs = pTs |)-\<succ>v-> s4.0;
s4' = (set_lvars (locals (snd s2.0))) s4.0 |]
==> G\<turnstile>Norm s0.0 \<midarrow>{accC,statT,mode}e·mn( {pTs}ps)-\<succ>v-> s4'
lemma eval_Init:
if inited C (globs s0.0) then s3.0 = Norm s0.0
else G\<turnstile>Norm ((init_class_obj G C)
s0.0) \<midarrow>(if C = Object then Skip
else Init (super (the (class G C))))-> s1.0 ∧
G\<turnstile>(set_lvars empty)
s1.0 \<midarrow>init (the (class G C))-> s2.0 ∧
s3.0 = (set_lvars (locals (snd s1.0))) s2.0
==> G\<turnstile>Norm s0.0 \<midarrow>Init C-> s3.0
lemma init_done:
initd C s ==> G\<turnstile>s \<midarrow>Init C-> s
lemma eval_StatRef:
G\<turnstile>s \<midarrow>StatRef
rt-\<succ>(if fst s = None then Null
else arbitrary)-> s
lemma SkipD:
G\<turnstile>s \<midarrow>Skip-> s' ==> s' = s
lemma Skip_eq:
G\<turnstile>s \<midarrow>Skip-> s' = (s = s')
lemma init_retains_locals:
[| G\<turnstile>s \<midarrow>t\<succ>-> (w, s'); t = In1r (Init C) |]
==> locals (snd s) = locals (snd s')
lemma halloc_xcpt:
G\<turnstile>(Some xc, s) \<midarrow>halloc oi\<succ>a-> s'
==> s' = (Some xc, s)
lemma eval_Methd:
G\<turnstile>s \<midarrow>In1l (body G C sig)\<succ>-> (w, s')
==> G\<turnstile>s \<midarrow>In1l (Methd C sig)\<succ>-> (w, s')
lemma eval_Body:
[| G\<turnstile>Norm s0.0 \<midarrow>Init D-> s1.0;
G\<turnstile>s1.0 \<midarrow>c-> s2.0; res = the (locals (snd s2.0) Result);
s3.0 =
(if ∃l. fst s2.0 = Some (Jump (Break l)) ∨ fst s2.0 = Some (Jump (Cont l))
then abupd (λx. Some (Error CrossMethodJump)) s2.0 else s2.0);
s4.0 = abupd (absorb Ret) s3.0 |]
==> G\<turnstile>Norm s0.0 \<midarrow>Body D c-\<succ>res-> s4.0
lemma eval_binop_arg2_indep:
¬ need_second_arg binop v1.0
==> eval_binop binop v1.0 x = eval_binop binop v1.0 y
lemma eval_BinOp_arg2_indepI:
[| G\<turnstile>Norm s0.0 \<midarrow>e1.0-\<succ>v1.0-> s1.0;
¬ need_second_arg binop v1.0 |]
==> G\<turnstile>Norm s0.0 \<midarrow>BinOp binop e1.0
e2.0-\<succ>eval_binop binop v1.0
v2.0-> s1.0
lemma unique_halloc:
[| G\<turnstile>s \<midarrow>halloc oi\<succ>a-> s';
G\<turnstile>s \<midarrow>halloc oi\<succ>a'-> s'' |]
==> a' = a ∧ s'' = s'
lemma single_valued_halloc:
single_valuedP
(λ(s, oi) (a, s'). G\<turnstile>s \<midarrow>halloc oi\<succ>a-> s')
lemma unique_sxalloc:
[| G\<turnstile>s \<midarrow>sxalloc-> s';
G\<turnstile>s \<midarrow>sxalloc-> s'' |]
==> s'' = s'
lemma single_valued_sxalloc:
single_valuedP (sxalloc G)
lemma split_pairD:
(x, y) = p ==> x = fst p ∧ y = snd p
lemma unique_eval:
[| G\<turnstile>s \<midarrow>t\<succ>-> (w, s');
G\<turnstile>s \<midarrow>t\<succ>-> (w', s'') |]
==> w' = w ∧ s'' = s'
lemma single_valued_eval:
single_valuedP (λ(s, t) (v, s'). G\<turnstile>s \<midarrow>t\<succ>-> (v, s'))