diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Evaln.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Evaln.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,373 @@ +(* Title: isabelle/Bali/Evaln.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1999 Technische Universitaet Muenchen +*) +header {* Operational evaluation (big-step) semantics of Java expressions and + statements +*} + +theory Evaln = Eval: + +text {* +Variant of eval relation with counter for bounded recursive depth +Evaln could completely replace Eval. +*} + +consts + + evaln :: "prog \ (state \ term \ nat \ vals \ state) set" + +syntax + + evaln :: "[prog, state, term, nat, vals * state] => bool" + ("_|-_ -_>-_-> _" [61,61,80, 61,61] 60) + evarn :: "[prog, state, var , vvar , nat, state] => bool" + ("_|-_ -_=>_-_-> _" [61,61,90,61,61,61] 60) + eval_n:: "[prog, state, expr , val , nat, state] => bool" + ("_|-_ -_->_-_-> _" [61,61,80,61,61,61] 60) + evalsn:: "[prog, state, expr list, val list, nat, state] => bool" + ("_|-_ -_#>_-_-> _" [61,61,61,61,61,61] 60) + execn :: "[prog, state, stmt , nat, state] => bool" + ("_|-_ -_-_-> _" [61,61,65, 61,61] 60) + +syntax (xsymbols) + + evaln :: "[prog, state, term, nat, vals \ state] \ bool" + ("_\_ \_\\_\ _" [61,61,80, 61,61] 60) + evarn :: "[prog, state, var , vvar , nat, state] \ bool" + ("_\_ \_=\_\_\ _" [61,61,90,61,61,61] 60) + eval_n:: "[prog, state, expr , val , nat, state] \ bool" + ("_\_ \_-\_\_\ _" [61,61,80,61,61,61] 60) + evalsn:: "[prog, state, expr list, val list, nat, state] \ bool" + ("_\_ \_\\_\_\ _" [61,61,61,61,61,61] 60) + execn :: "[prog, state, stmt , nat, state] \ bool" + ("_\_ \_\_\ _"   [61,61,65, 61,61] 60) + +translations + + "G\s \t \\n\ w___s' " == "(s,t,n,w___s') \ evaln G" + "G\s \t \\n\ (w, s')" <= "(s,t,n,w, s') \ evaln G" + "G\s \t \\n\ (w,x,s')" <= "(s,t,n,w,x,s') \ evaln G" + "G\s \c \n\ (x,s')" <= "G\s \In1r c\\n\ (\ ,x,s')" + "G\s \c \n\ s' " == "G\s \In1r c\\n\ (\ , s')" + "G\s \e-\v \n\ (x,s')" <= "G\s \In1l e\\n\ (In1 v ,x,s')" + "G\s \e-\v \n\ s' " == "G\s \In1l e\\n\ (In1 v , s')" + "G\s \e=\vf \n\ (x,s')" <= "G\s \In2 e\\n\ (In2 vf,x,s')" + "G\s \e=\vf \n\ s' " == "G\s \In2 e\\n\ (In2 vf, s')" + "G\s \e\\v \n\ (x,s')" <= "G\s \In3 e\\n\ (In3 v ,x,s')" + "G\s \e\\v \n\ s' " == "G\s \In3 e\\n\ (In3 v , s')" + + +inductive "evaln G" intros + +(* propagation of abrupt completion *) + + Abrupt: "G\(Some xc,s) \t\\n\ (arbitrary3 t,(Some xc,s))" + + +(* evaluation of variables *) + + LVar: "G\Norm s \LVar vn=\lvar vn s\n\ Norm s" + + FVar: "\G\Norm s0 \Init C\n\ s1; G\s1 \e-\a'\n\ s2; + (v,s2') = fvar C stat fn a' s2\ \ + G\Norm s0 \{C,stat}e..fn=\v\n\ s2'" + + AVar: "\G\ Norm s0 \e1-\a\n\ s1 ; G\s1 \e2-\i\n\ s2; + (v,s2') = avar G i a s2\ \ + G\Norm s0 \e1.[e2]=\v\n\ s2'" + + + + +(* evaluation of expressions *) + + NewC: "\G\Norm s0 \Init C\n\ s1; + G\ s1 \halloc (CInst C)\a\ s2\ \ + G\Norm s0 \NewC C-\Addr a\n\ s2" + + NewA: "\G\Norm s0 \init_comp_ty T\n\ s1; G\s1 \e-\i'\n\ s2; + G\abupd (check_neg i') s2 \halloc (Arr T (the_Intg i'))\a\ s3\ \ + G\Norm s0 \New T[e]-\Addr a\n\ s3" + + Cast: "\G\Norm s0 \e-\v\n\ s1; + s2 = abupd (raise_if (\G,snd s1\v fits T) ClassCast) s1\ \ + G\Norm s0 \Cast T e-\v\n\ s2" + + Inst: "\G\Norm s0 \e-\v\n\ s1; + b = (v\Null \ G,store s1\v fits RefT T)\ \ + G\Norm s0 \e InstOf T-\Bool b\n\ s1" + + Lit: "G\Norm s \Lit v-\v\n\ Norm s" + + Super: "G\Norm s \Super-\val_this s\n\ Norm s" + + Acc: "\G\Norm s0 \va=\(v,f)\n\ s1\ \ + G\Norm s0 \Acc va-\v\n\ s1" + + Ass: "\G\Norm s0 \va=\(w,f)\n\ s1; + G\ s1 \e-\v \n\ s2\ \ + G\Norm s0 \va:=e-\v\n\ assign f v s2" + + Cond: "\G\Norm s0 \e0-\b\n\ s1; + G\ s1 \(if the_Bool b then e1 else e2)-\v\n\ s2\ \ + G\Norm s0 \e0 ? e1 : e2-\v\n\ s2" + + Call: + "\G\Norm s0 \e-\a'\n\ s1; G\s1 \args\\vs\n\ s2; + D = invocation_declclass G mode (store s2) a' statT \name=mn,parTs=pTs\; + G\init_lvars G D \name=mn,parTs=pTs\ mode a' vs s2 + \Methd D \name=mn,parTs=pTs\-\v\n\ s3\ + \ G\Norm s0 \{statT,mode}e\mn({pTs}args)-\v\n\ (restore_lvars s2 s3)" + + Methd:"\G\Norm s0 \body G D sig-\v\n\ s1\ \ + G\Norm s0 \Methd D sig-\v\Suc n\ s1" + + Body: "\G\Norm s0\Init D\n\ s1; G\s1 \c\n\ s2\\ + G\Norm s0 \Body D c-\the (locals (store s2) Result)\n\abupd (absorb Ret) s2" + +(* evaluation of expression lists *) + + Nil: + "G\Norm s0 \[]\\[]\n\ Norm s0" + + Cons: "\G\Norm s0 \e -\ v \n\ s1; + G\ s1 \es\\vs\n\ s2\ \ + G\Norm s0 \e#es\\v#vs\n\ s2" + + +(* execution of statements *) + + Skip: "G\Norm s \Skip\n\ Norm s" + + Expr: "\G\Norm s0 \e-\v\n\ s1\ \ + G\Norm s0 \Expr e\n\ s1" + + Lab: "\G\Norm s0 \c \n\ s1\ \ + G\Norm s0 \l\ c\n\ abupd (absorb (Break l)) s1" + + Comp: "\G\Norm s0 \c1 \n\ s1; + G\ s1 \c2 \n\ s2\ \ + G\Norm s0 \c1;; c2\n\ s2" + + If: "\G\Norm s0 \e-\b\n\ s1; + G\ s1\(if the_Bool b then c1 else c2)\n\ s2\ \ + G\Norm s0 \If(e) c1 Else c2 \n\ s2" + + Loop: "\G\Norm s0 \e-\b\n\ s1; + if normal s1 \ the_Bool b + then (G\s1 \c\n\ s2 \ + G\(abupd (absorb (Cont l)) s2) \l\ While(e) c\n\ s3) + else s3 = s1\ \ + G\Norm s0 \l\ While(e) c\n\ s3" + + Do: "G\Norm s \Do j\n\ (Some (Jump j), s)" + + Throw:"\G\Norm s0 \e-\a'\n\ s1\ \ + G\Norm s0 \Throw e\n\ abupd (throw a') s1" + + Try: "\G\Norm s0 \c1\n\ s1; G\s1 \sxalloc\ s2; + if G,s2\catch tn then G\new_xcpt_var vn s2 \c2\n\ s3 else s3 = s2\ + \ + G\Norm s0 \Try c1 Catch(tn vn) c2\n\ s3" + + Fin: "\G\Norm s0 \c1\n\ (x1,s1); + G\Norm s1 \c2\n\ s2\ \ + G\Norm s0 \c1 Finally c2\n\ abupd (abrupt_if (x1\None) x1) s2" + + Init: "\the (class G C) = c; + if inited C (globs s0) then s3 = Norm s0 + else (G\Norm (init_class_obj G C s0) + \(if C = Object then Skip else Init (super c))\n\ s1 \ + G\set_lvars empty s1 \init c\n\ s2 \ + s3 = restore_lvars s1 s2)\ + \ + G\Norm s0 \Init C\n\ s3" +monos + if_def2 + +lemma evaln_eval: "\ws. G\s \t\\n\ ws \ G\s \t\\ ws" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (erule evaln.induct) +apply (rule eval.intros, (assumption+)?,(force split del: split_if)?)+ +done + + +lemma Suc_le_D_lemma: "\Suc n <= m'; (\m. n <= m \ P (Suc m)) \ \ P m'" +apply (frule Suc_le_D) +apply fast +done + +lemma evaln_nonstrict [rule_format (no_asm), elim]: + "\ws. G\s \t\\n\ ws \ \m. n\m \ G\s \t\\m\ ws" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (erule evaln.induct) +apply (tactic {* ALLGOALS (EVERY'[strip_tac, TRY o etac (thm "Suc_le_D_lemma"), + REPEAT o smp_tac 1, + resolve_tac (thms "evaln.intros") THEN_ALL_NEW TRY o atac]) *}) +(* 3 subgoals *) +apply (auto split del: split_if) +done + +lemmas evaln_nonstrict_Suc = evaln_nonstrict [OF _ le_refl [THEN le_SucI]] + +lemma evaln_max2: "\G\s1 \t1\\n1\ ws1; G\s2 \t2\\n2\ ws2\ \ + G\s1 \t1\\max n1 n2\ ws1 \ G\s2 \t2\\max n1 n2\ ws2" +apply (fast intro: le_maxI1 le_maxI2) +done + +lemma evaln_max3: +"\G\s1 \t1\\n1\ ws1; G\s2 \t2\\n2\ ws2; G\s3 \t3\\n3\ ws3\ \ + G\s1 \t1\\max (max n1 n2) n3\ ws1 \ + G\s2 \t2\\max (max n1 n2) n3\ ws2 \ + G\s3 \t3\\max (max n1 n2) n3\ ws3" +apply (drule (1) evaln_max2, erule thin_rl) +apply (fast intro!: le_maxI1 le_maxI2) +done + +lemma eval_evaln: "\ws. G\s \t\\ ws \ (\n. G\s \t\\n\ ws)" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (erule eval.induct) +apply (tactic {* ALLGOALS + (asm_full_simp_tac (HOL_basic_ss addsplits [split_if_asm])) *}) +apply (tactic {* ALLGOALS (EVERY'[ + REPEAT o eresolve_tac [exE, conjE], rtac exI, + TRY o datac (thm "evaln_max3") 2, REPEAT o etac conjE, + resolve_tac (thms "evaln.intros") THEN_ALL_NEW + force_tac (HOL_cs, HOL_ss)]) *}) +done + +declare split_if [split del] split_if_asm [split del] + option.split [split del] option.split_asm [split del] +inductive_cases evaln_cases: "G\s \t\\n\ vs'" + +inductive_cases evaln_elim_cases: + "G\(Some xc, s) \t \\n\ vs'" + "G\Norm s \In1r Skip \\n\ xs'" + "G\Norm s \In1r (Do j) \\n\ xs'" + "G\Norm s \In1r (l\ c) \\n\ xs'" + "G\Norm s \In3 ([]) \\n\ vs'" + "G\Norm s \In3 (e#es) \\n\ vs'" + "G\Norm s \In1l (Lit w) \\n\ vs'" + "G\Norm s \In2 (LVar vn) \\n\ vs'" + "G\Norm s \In1l (Cast T e) \\n\ vs'" + "G\Norm s \In1l (e InstOf T) \\n\ vs'" + "G\Norm s \In1l (Super) \\n\ vs'" + "G\Norm s \In1l (Acc va) \\n\ vs'" + "G\Norm s \In1r (Expr e) \\n\ xs'" + "G\Norm s \In1r (c1;; c2) \\n\ xs'" + "G\Norm s \In1l (Methd C sig) \\n\ xs'" + "G\Norm s \In1l (Body D c) \\n\ xs'" + "G\Norm s \In1l (e0 ? e1 : e2) \\n\ vs'" + "G\Norm s \In1r (If(e) c1 Else c2) \\n\ xs'" + "G\Norm s \In1r (l\ While(e) c) \\n\ xs'" + "G\Norm s \In1r (c1 Finally c2) \\n\ xs'" + "G\Norm s \In1r (Throw e) \\n\ xs'" + "G\Norm s \In1l (NewC C) \\n\ vs'" + "G\Norm s \In1l (New T[e]) \\n\ vs'" + "G\Norm s \In1l (Ass va e) \\n\ vs'" + "G\Norm s \In1r (Try c1 Catch(tn vn) c2) \\n\ xs'" + "G\Norm s \In2 ({C,stat}e..fn) \\n\ vs'" + "G\Norm s \In2 (e1.[e2]) \\n\ vs'" + "G\Norm s \In1l ({statT,mode}e\mn({pT}p)) \\n\ vs'" + "G\Norm s \In1r (Init C) \\n\ xs'" +declare split_if [split] split_if_asm [split] + option.split [split] option.split_asm [split] + +lemma evaln_Inj_elim: "G\s \t\\n\ (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 evaln_cases , auto) +apply (induct_tac "t") +apply (induct_tac "a") +apply auto +done + +ML_setup {* +fun enf nam inj rhs = +let + val name = "evaln_" ^ nam ^ "_eq" + val lhs = "G\s \" ^ inj ^ " t\\n\ (w, s')" + val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") + (K [Auto_tac, ALLGOALS (ftac (thm "evaln_Inj_elim")) THEN Auto_tac]) + fun is_Inj (Const (inj,_) $ _) = true + | is_Inj _ = false + fun pred (_ $ (Const ("Pair",_) $ _ $ (Const ("Pair", _) $ _ $ + (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ x $ _ )))) $ _ ) = is_Inj x +in + make_simproc name lhs pred (thm name) +end; + +val evaln_expr_proc = enf "expr" "In1l" "\v. w=In1 v \ G\s \t-\v \n\ s'"; +val evaln_var_proc = enf "var" "In2" "\vf. w=In2 vf \ G\s \t=\vf\n\ s'"; +val evaln_exprs_proc= enf "exprs""In3" "\vs. w=In3 vs \ G\s \t\\vs\n\ s'"; +val evaln_stmt_proc = enf "stmt" "In1r" " w=\ \ G\s \t \n\ s'"; +Addsimprocs [evaln_expr_proc,evaln_var_proc,evaln_exprs_proc,evaln_stmt_proc]; + +bind_thms ("evaln_AbruptIs", sum3_instantiate (thm "evaln.Abrupt")) +*} +declare evaln_AbruptIs [intro!] + +lemma evaln_abrupt_lemma: "G\s \e\\n\ (v,s') \ + fst s = Some xc \ s' = s \ v = arbitrary3 e" +apply (erule evaln_cases , auto) +done + +lemma evaln_abrupt: + "\s'. G\(Some xc,s) \e\\n\ (w,s') = (s' = (Some xc,s) \ + w=arbitrary3 e \ G\(Some xc,s) \e\\n\ (arbitrary3 e,(Some xc,s)))" +apply auto +apply (frule evaln_abrupt_lemma, auto)+ +done + +ML {* +local + fun is_Some (Const ("Pair",_) $ (Const ("Option.option.Some",_) $ _)$ _) =true + | is_Some _ = false + fun pred (_ $ (Const ("Pair",_) $ + _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $ + (Const ("Pair", _) $ _ $ x)))) $ _ ) = is_Some x +in + val evaln_abrupt_proc = + make_simproc "evaln_abrupt" "G\(Some xc,s) \e\\n\ (w,s')" pred (thm "evaln_abrupt") +end; +Addsimprocs [evaln_abrupt_proc] +*} + +lemma evaln_LitI: "G\s \Lit v-\(if normal s then v else arbitrary)\n\ s" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: evaln.Lit) + +lemma CondI: + "\s1. \G\s \e-\b\n\ s1; G\s1 \(if the_Bool b then e1 else e2)-\v\n\ s2\ \ + G\s \e ? e1 : e2-\(if normal s1 then v else arbitrary)\n\ s2" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: evaln.Cond) + +lemma evaln_SkipI [intro!]: "G\s \Skip\n\ s" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: evaln.Skip) + +lemma evaln_ExprI: "G\s \e-\v\n\ s' \ G\s \Expr e\n\ s'" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: evaln.Expr) + +lemma evaln_CompI: "\G\s \c1\n\ s1; G\s1 \c2\n\ s2\ \ G\s \c1;; c2\n\ s2" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: evaln.Comp) + +lemma evaln_IfI: + "\G\s \e-\v\n\ s1; G\s1 \(if the_Bool v then c1 else c2)\n\ s2\ \ + G\s \If(e) c1 Else c2\n\ s2" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: evaln.If) + +lemma evaln_SkipD [dest!]: "G\s \Skip\n\ s' \ s' = s" +by (erule evaln_cases, auto) + +lemma evaln_Skip_eq [simp]: "G\s \Skip\n\ s' = (s = s')" +apply auto +done + +end