# HG changeset patch # User schirmer # Date 1012233619 -3600 # Node ID 00d4a435777f03d09c2ef071a766fe3cd94f55c7 # Parent de505273c971374d9de5fd4393f4279833664146 Isabelle/Bali sources; diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/AxCompl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/AxCompl.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,645 @@ +(* Title: isabelle/Bali/AxCompl.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1999 Technische Universitaet Muenchen +*) + +header {* +Completeness proof for Axiomatic semantics of Java expressions and statements +*} + +theory AxCompl = AxSem: + +text {* +design issues: +\begin{itemize} +\item proof structured by Most General Formulas (-> Thomas Kleymann) +\end{itemize} +*} +section "set of not yet initialzed classes" + +constdefs + + nyinitcls :: "prog \ state \ qtname set" + "nyinitcls G s \ {C. is_class G C \ \ initd C s}" + +lemma nyinitcls_subset_class: "nyinitcls G s \ {C. is_class G C}" +apply (unfold nyinitcls_def) +apply fast +done + +lemmas finite_nyinitcls [simp] = + finite_is_class [THEN nyinitcls_subset_class [THEN finite_subset], standard] + +lemma card_nyinitcls_bound: "card (nyinitcls G s) \ card {C. is_class G C}" +apply (rule nyinitcls_subset_class [THEN finite_is_class [THEN card_mono]]) +done + +lemma nyinitcls_set_locals_cong [simp]: + "nyinitcls G (x,set_locals l s) = nyinitcls G (x,s)" +apply (unfold nyinitcls_def) +apply (simp (no_asm)) +done + +lemma nyinitcls_abrupt_cong [simp]: "nyinitcls G (f x, y) = nyinitcls G (x, y)" +apply (unfold nyinitcls_def) +apply (simp (no_asm)) +done + +lemma nyinitcls_abupd_cong [simp]:"!!s. nyinitcls G (abupd f s) = nyinitcls G s" +apply (unfold nyinitcls_def) +apply (simp (no_asm_simp) only: split_tupled_all) +apply (simp (no_asm)) +done + +lemma card_nyinitcls_abrupt_congE [elim!]: + "card (nyinitcls G (x, s)) \ n \ card (nyinitcls G (y, s)) \ n" +apply (unfold nyinitcls_def) +apply auto +done + +lemma nyinitcls_new_xcpt_var [simp]: +"nyinitcls G (new_xcpt_var vn s) = nyinitcls G s" +apply (unfold nyinitcls_def) +apply (induct_tac "s") +apply (simp (no_asm)) +done + +lemma nyinitcls_init_lvars [simp]: + "nyinitcls G ((init_lvars G C sig mode a' pvs) s) = nyinitcls G s" +apply (induct_tac "s") +apply (simp (no_asm) add: init_lvars_def2 split add: split_if) +done + +lemma nyinitcls_emptyD: "\nyinitcls G s = {}; is_class G C\ \ initd C s" +apply (unfold nyinitcls_def) +apply fast +done + +lemma card_Suc_lemma: "\card (insert a A) \ Suc n; a\A; finite A\ \ card A \ n" +apply (rotate_tac 1) +apply clarsimp +done + +lemma nyinitcls_le_SucD: +"\card (nyinitcls G (x,s)) \ Suc n; \inited C (globs s); class G C=Some y\ \ + card (nyinitcls G (x,init_class_obj G C s)) \ n" +apply (subgoal_tac + "nyinitcls G (x,s) = insert C (nyinitcls G (x,init_class_obj G C s))") +apply clarsimp +apply (erule thin_rl) +apply (rule card_Suc_lemma [OF _ _ finite_nyinitcls]) +apply (auto dest!: not_initedD elim!: + simp add: nyinitcls_def inited_def split add: split_if_asm) +done + +ML {* bind_thm("inited_gext'",permute_prems 0 1 (thm "inited_gext")) *} + +lemma nyinitcls_gext: "snd s\|snd s' \ nyinitcls G s' \ nyinitcls G s" +apply (unfold nyinitcls_def) +apply (force dest!: inited_gext') +done + +lemma card_nyinitcls_gext: + "\snd s\|snd s'; card (nyinitcls G s) \ n\\ card (nyinitcls G s') \ n" +apply (rule le_trans) +apply (rule card_mono) +apply (rule finite_nyinitcls) +apply (erule nyinitcls_gext) +apply assumption +done + + +section "init-le" + +constdefs + init_le :: "prog \ nat \ state \ bool" ("_\init\_" [51,51] 50) + "G\init\n \ \s. card (nyinitcls G s) \ n" + +lemma init_le_def2 [simp]: "(G\init\n) s = (card (nyinitcls G s)\n)" +apply (unfold init_le_def) +apply auto +done + +lemma All_init_leD: "\n::nat. G,A\{P \. G\init\n} t\ {Q} \ G,A\{P} t\ {Q}" +apply (drule spec) +apply (erule conseq1) +apply clarsimp +apply (rule card_nyinitcls_bound) +done + +section "Most General Triples and Formulas" + +constdefs + + remember_init_state :: "state assn" ("\") + "\ \ \Y s Z. s = Z" + +lemma remember_init_state_def2 [simp]: "\ Y = op =" +apply (unfold remember_init_state_def) +apply (simp (no_asm)) +done + +consts + + MGF ::"[state assn, term, prog] \ state triple" ("{_} _\ {_\}"[3,65,3]62) + MGFn::"[nat , term, prog] \ state triple" ("{=:_} _\ {_\}"[3,65,3]62) + +defs + + + MGF_def: + "{P} t\ {G\} \ {P} t\ {\Y s' s. G\s \t\\ (Y,s')}" + + MGFn_def: + "{=:n} t\ {G\} \ {\ \. G\init\n} t\ {G\}" + +(* unused *) +lemma MGF_valid: "G,{}\{\} t\ {G\}" +apply (unfold MGF_def) +apply (force dest!: evaln_eval simp add: ax_valids_def triple_valid_def2) +done + +lemma MGF_res_eq_lemma [simp]: + "(\Y' Y s. Y = Y' \ P s \ Q s) = (\s. P s \ Q s)" +apply auto +done + +lemma MGFn_def2: +"G,A\{=:n} t\ {G\} = G,A\{\ \. G\init\n} + t\ {\Y s' s. G\s \t\\ (Y,s')}" +apply (unfold MGFn_def MGF_def) +apply fast +done + +lemma MGF_MGFn_iff: "G,A\{\} t\ {G\} = (\n. G,A\{=:n} t\ {G\})" +apply (simp (no_asm_use) add: MGFn_def2 MGF_def) +apply safe +apply (erule_tac [2] All_init_leD) +apply (erule conseq1) +apply clarsimp +done + +lemma MGFnD: +"G,A\{=:n} t\ {G\} \ + G,A\{(\Y' s' s. s' = s \ P s) \. G\init\n} + t\ {(\Y' s' s. G\s\t\\(Y',s') \ P s) \. G\init\n}" +apply (unfold init_le_def) +apply (simp (no_asm_use) add: MGFn_def2) +apply (erule conseq12) +apply clarsimp +apply (erule (1) eval_gext [THEN card_nyinitcls_gext]) +done +lemmas MGFnD' = MGFnD [of _ _ _ _ "\x. True"] + +lemma MGFNormalI: "G,A\{Normal \} t\ {G\} \ + G,(A::state triple set)\{\::state assn} t\ {G\}" +apply (unfold MGF_def) +apply (rule ax_Normal_cases) +apply (erule conseq1) +apply clarsimp +apply (rule ax_derivs.Abrupt [THEN conseq1]) +apply (clarsimp simp add: Let_def) +done + +lemma MGFNormalD: "G,A\{\} t\ {G\} \ G,A\{Normal \} t\ {G\}" +apply (unfold MGF_def) +apply (erule conseq1) +apply clarsimp +done + +lemma MGFn_NormalI: +"G,(A::state triple set)\{Normal((\Y' s' s. s'=s \ normal s) \. G\init\n)}t\ + {\Y s' s. G\s \t\\ (Y,s')} \ G,A\{=:n}t\{G\}" +apply (simp (no_asm_use) add: MGFn_def2) +apply (rule ax_Normal_cases) +apply (erule conseq1) +apply clarsimp +apply (rule ax_derivs.Abrupt [THEN conseq1]) +apply (clarsimp simp add: Let_def) +done + +lemma MGFn_free_wt: + "(\T L C. \prg=G,cls=C,lcl=L\\t\T) + \ G,(A::state triple set)\{=:n} t\ {G\} + \ G,A\{=:n} t\ {G\}" +apply (rule MGFn_NormalI) +apply (rule ax_free_wt) +apply (auto elim: conseq12 simp add: MGFn_def MGF_def) +done + + +section "main lemmas" + +declare fun_upd_apply [simp del] +declare splitI2 [rule del] (*prevents ugly renaming of state variables*) + +ML_setup {* +Delsimprocs [eval_expr_proc, eval_var_proc, eval_exprs_proc, eval_stmt_proc] +*} (*prevents modifying rhs of MGF*) +ML {* +val eval_css = (claset() delrules [thm "eval.Abrupt"] addSIs (thms "eval.intros") + delrules[thm "eval.Expr", thm "eval.Init", thm "eval.Try"] + addIs [thm "eval.Expr", thm "eval.Init"] + addSEs[thm "eval.Try"] delrules[equalityCE], + simpset() addsimps [split_paired_all,Let_def] + addsimprocs [eval_expr_proc,eval_var_proc,eval_exprs_proc,eval_stmt_proc]); +val eval_Force_tac = force_tac eval_css; + +val wt_prepare_tac = EVERY'[ + rtac (thm "MGFn_free_wt"), + clarsimp_tac (claset() addSEs (thms "wt_elim_cases"), simpset())] +val compl_prepare_tac = EVERY'[rtac (thm "MGFn_NormalI"), Simp_tac] +val forw_hyp_tac = EVERY'[etac (thm "MGFnD'" RS thm "conseq12"), Clarsimp_tac] +val forw_hyp_eval_Force_tac = + EVERY'[TRY o rtac allI, forw_hyp_tac, eval_Force_tac] +*} + +lemma MGFn_Init: "\m. Suc m\n \ (\t. G,A\{=:m} t\ {G\}) \ + G,(A::state triple set)\{=:n} In1r (Init C)\ {G\}" +apply (tactic "wt_prepare_tac 1") +(* requires is_class G C two times for nyinitcls *) +apply (tactic "compl_prepare_tac 1") +apply (rule_tac C = "initd C" in ax_cases) +apply (rule ax_derivs.Done [THEN conseq1]) +apply (clarsimp intro!: init_done) +apply (rule_tac y = n in nat.exhaust, clarsimp) +apply (rule ax_impossible [THEN conseq1]) +apply (force dest!: nyinitcls_emptyD) +apply clarsimp +apply (drule_tac x = "nat" in spec) +apply clarsimp +apply (rule_tac Q = " (\Y s' (x,s) . G\ (x,init_class_obj G C s) \ (if C=Object then Skip else Init (super (the (class G C))))\ s' \ x=None \ \inited C (globs s)) \. G\init\nat" in ax_derivs.Init) +apply simp +apply (rule_tac P' = "Normal ((\Y s' s. s' = supd (init_class_obj G C) s \ normal s \ \ initd C s) \. G\init\nat) " in conseq1) +prefer 2 +apply (force elim!: nyinitcls_le_SucD) +apply (simp split add: split_if, rule conjI, clarify) +apply (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1") +apply clarify +apply (drule spec) +apply (erule MGFnD' [THEN conseq12]) +apply (tactic "force_tac (claset(), simpset() addsimprocs[eval_stmt_proc]) 1") +apply (rule allI) +apply (drule spec) +apply (erule MGFnD' [THEN conseq12]) +apply clarsimp +apply (tactic {* pair_tac "sa" 1 *}) +apply (tactic"clarsimp_tac (claset(), simpset() addsimprocs[eval_stmt_proc]) 1") +apply (rule eval_Init, force+) +done +lemmas MGFn_InitD = MGFn_Init [THEN MGFnD, THEN ax_NormalD] + +lemma MGFn_Call: +"\\C sig. G,(A::state triple set)\{=:n} In1l (Methd C sig)\ {G\}; + G,A\{=:n} In1l e\ {G\}; G,A\{=:n} In3 ps\ {G\}\ \ + G,A\{=:n} In1l ({statT,mode}e\mn({pTs'}ps))\ {G\}" +apply (tactic "wt_prepare_tac 1") (* required for equating mode = invmode m e *) +apply (tactic "compl_prepare_tac 1") +apply (rule_tac R = "\a'. (\Y (x2,s2) (x,s) . x = None \ (\s1 pvs. G\Norm s \e-\a'\ s1 \ Y = In3 pvs \ G\s1 \ps\\pvs\ (x2,s2))) \. G\init\n" in ax_derivs.Call) +apply (erule MGFnD [THEN ax_NormalD]) +apply safe +apply (erule_tac V = "All ?P" in thin_rl, tactic "forw_hyp_eval_Force_tac 1") +apply (drule spec, drule spec) +apply (erule MGFnD' [THEN conseq12]) +apply (tactic "clarsimp_tac eval_css 1") +apply (erule (1) eval_Call) +apply (rule HOL.refl) +apply (simp (no_asm_simp))+ +done + +lemma MGF_altern: "G,A\{Normal (\ \. p)} t\ {G\} = + G,A\{Normal ((\Y s Z. \w s'. G\s \t\\ (w,s') \ (w,s') = Z) \. p)} + t\ {\Y s Z. (Y,s) = Z}" +apply (unfold MGF_def) +apply (auto del: conjI elim!: conseq12) +apply (case_tac "\w s. G\Norm sa \t\\ (w,s) ") +apply (fast dest: unique_eval) +apply clarsimp +apply (erule thin_rl) +apply (erule thin_rl) +apply (drule split_paired_All [THEN subst]) +apply (clarsimp elim!: state_not_single) +done + + +lemma MGFn_Loop: +"\G,(A::state triple set)\{=:n} In1l expr\ {G\};G,A\{=:n} In1r stmnt\ {G\} \ +\ + G,A\{=:n} In1r (l\ While(expr) stmnt)\ {G\}" +apply (rule MGFn_NormalI, simp) +apply (rule_tac p2 = "\s. card (nyinitcls G s) \ n" in + MGF_altern [unfolded MGF_def, THEN iffD2, THEN conseq1]) +prefer 2 +apply clarsimp +apply (rule_tac P' = +"((\Y s Z. \w s'. G\s \In1r (l\ While(expr) stmnt) \\ (w,s') \ (w,s') = Z) + \. (\s. card (nyinitcls G s) \ n))" in conseq12) +prefer 2 +apply clarsimp +apply (tactic "smp_tac 1 1", erule_tac V = "All ?P" in thin_rl) +apply (rule_tac [2] P' = " (\b s (Y',s') . (\s0. G\s0 \In1l expr\\ (b,s)) \ (if normal s \ the_Bool (the_In1 b) then (\s'' w s0. G\s \stmnt\ s'' \ G\(abupd (absorb (Cont l)) s'') \In1r (l\ While(expr) stmnt) \\ (w,s0) \ (w,s0) = (Y',s')) else (\,s) = (Y',s'))) \. G\init\n" in polymorphic_Loop) +apply (force dest!: eval.Loop split add: split_if_asm) +prefer 2 +apply (erule MGFnD' [THEN conseq12]) +apply clarsimp +apply (erule_tac V = "card (nyinitcls G s') \ n" in thin_rl) +apply (tactic "eval_Force_tac 1") +apply (erule MGFnD' [THEN conseq12] , clarsimp) +apply (rule conjI, erule exI) +apply (tactic "clarsimp_tac eval_css 1") +apply (case_tac "a") +prefer 2 +apply (clarsimp) +apply (clarsimp split add: split_if) +apply (rule conjI, (tactic {* force_tac (claset() addSDs [thm "eval.Loop"], + simpset() addsimps [split_paired_all] addsimprocs [eval_stmt_proc]) 1*})+) +done + +lemma MGFn_lemma [rule_format (no_asm)]: + "\n C sig. G,(A::state triple set)\{=:n} In1l (Methd C sig)\ {G\} \ + \t. G,A\{=:n} t\ {G\}" +apply (rule full_nat_induct) +apply (rule allI) +apply (drule_tac x = n in spec) +apply (drule_tac psi = "All ?P" in asm_rl) +apply (subgoal_tac "\v e c es. G,A\{=:n} In2 v\ {G\} \ G,A\{=:n} In1l e\ {G\} \ G,A\{=:n} In1r c\ {G\} \ G,A\{=:n} In3 es\ {G\}") +apply (tactic "Clarify_tac 2") +apply (induct_tac "t") +apply (induct_tac "a") +apply fast+ +apply (rule var_expr_stmt.induct) +(* 28 subgoals *) +prefer 14 apply fast (* Methd *) +prefer 13 apply (erule (2) MGFn_Call) +apply (erule_tac [!] V = "All ?P" in thin_rl) (* assumptions on Methd *) +apply (erule_tac [24] MGFn_Init) +prefer 19 apply (erule (1) MGFn_Loop) +apply (tactic "ALLGOALS compl_prepare_tac") + +apply (rule ax_derivs.LVar [THEN conseq1], tactic "eval_Force_tac 1") + +apply (rule ax_derivs.FVar) +apply (erule MGFn_InitD) +apply (tactic "forw_hyp_eval_Force_tac 1") + +apply (rule ax_derivs.AVar) +apply (erule MGFnD [THEN ax_NormalD]) +apply (tactic "forw_hyp_eval_Force_tac 1") + +apply (rule ax_derivs.NewC) +apply (erule MGFn_InitD [THEN conseq2]) +apply (tactic "eval_Force_tac 1") + +apply (rule_tac Q = "(\Y' s' s. normal s \ G\s \In1r (init_comp_ty ty) \\ (Y',s')) \. G\init\n" in ax_derivs.NewA) +apply (simp add: init_comp_ty_def split add: split_if) +apply (rule conjI, clarsimp) +apply (erule MGFn_InitD [THEN conseq2]) +apply (tactic "clarsimp_tac eval_css 1") +apply clarsimp +apply (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1") +apply (tactic "forw_hyp_eval_Force_tac 1") + +apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast],tactic"eval_Force_tac 1") + +apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst],tactic"eval_Force_tac 1") +apply (rule ax_derivs.Lit [THEN conseq1], tactic "eval_Force_tac 1") +apply (rule ax_derivs.Super [THEN conseq1], tactic "eval_Force_tac 1") +apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc],tactic"eval_Force_tac 1") + +apply (rule ax_derivs.Ass) +apply (erule MGFnD [THEN ax_NormalD]) +apply (tactic "forw_hyp_eval_Force_tac 1") + +apply (rule ax_derivs.Cond) +apply (erule MGFnD [THEN ax_NormalD]) +apply (rule allI) +apply (rule ax_Normal_cases) +prefer 2 +apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def) +apply (tactic "eval_Force_tac 1") +apply (case_tac "b") +apply (simp, tactic "forw_hyp_eval_Force_tac 1") +apply (simp, tactic "forw_hyp_eval_Force_tac 1") + +apply (rule_tac Q = " (\Y' s' s. normal s \ G\s \Init pid_field_type\ s') \. G\init\n" in ax_derivs.Body) + apply (erule MGFn_InitD [THEN conseq2]) + apply (tactic "eval_Force_tac 1") +apply (tactic "forw_hyp_tac 1") +apply (tactic {* clarsimp_tac (eval_css delsimps2 [split_paired_all]) 1 *}) +apply (erule (1) eval.Body) + +apply (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1") + +apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr],tactic"eval_Force_tac 1") + +apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab]) +apply (tactic "clarsimp_tac eval_css 1") + +apply (rule ax_derivs.Comp) +apply (erule MGFnD [THEN ax_NormalD]) +apply (tactic "forw_hyp_eval_Force_tac 1") + +apply (rule ax_derivs.If) +apply (erule MGFnD [THEN ax_NormalD]) +apply (rule allI) +apply (rule ax_Normal_cases) +prefer 2 +apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def) +apply (tactic "eval_Force_tac 1") +apply (case_tac "b") +apply (simp, tactic "forw_hyp_eval_Force_tac 1") +apply (simp, tactic "forw_hyp_eval_Force_tac 1") + +apply (rule ax_derivs.Do [THEN conseq1]) +apply (tactic {* force_tac (eval_css addsimps2 [thm "abupd_def2"]) 1 *}) + +apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw]) +apply (tactic "clarsimp_tac eval_css 1") + +apply (rule_tac Q = " (\Y' s' s. normal s \ (\s''. G\s \In1r stmt1\\ (Y',s'') \ G\s'' \sxalloc\ s')) \. G\init\n" in ax_derivs.Try) +apply (tactic "eval_Force_tac 3") +apply (tactic "forw_hyp_eval_Force_tac 2") +apply (erule MGFnD [THEN ax_NormalD, THEN conseq2]) +apply (tactic "clarsimp_tac eval_css 1") +apply (force elim: sxalloc_gext [THEN card_nyinitcls_gext]) + +apply (rule_tac Q = " (\Y' s' s. normal s \ G\s \stmt1\ s') \. G\init\n" in ax_derivs.Fin) +apply (tactic "forw_hyp_eval_Force_tac 1") +apply (rule allI) +apply (tactic "forw_hyp_tac 1") +apply (tactic {* pair_tac "sb" 1 *}) +apply (tactic"clarsimp_tac (claset(),simpset() addsimprocs [eval_stmt_proc]) 1") +apply (drule (1) eval.Fin) +apply clarsimp + +apply (rule ax_derivs.Nil [THEN conseq1], tactic "eval_Force_tac 1") + +apply (rule ax_derivs.Cons) +apply (erule MGFnD [THEN ax_NormalD]) +apply (tactic "forw_hyp_eval_Force_tac 1") +done + +lemma MGF_asm: "\C sig. is_methd G C sig \ G,A\{\} In1l (Methd C sig)\ {G\} \ + G,(A::state triple set)\{\} t\ {G\}" +apply (simp (no_asm_use) add: MGF_MGFn_iff) +apply (rule allI) +apply (rule MGFn_lemma) +apply (intro strip) +apply (rule MGFn_free_wt) +apply (force dest: wt_Methd_is_methd) +done + +declare splitI2 [intro!] +ML_setup {* +Addsimprocs [ eval_expr_proc, eval_var_proc, eval_exprs_proc, eval_stmt_proc] +*} + + +section "nested version" + +lemma nesting_lemma' [rule_format (no_asm)]: "[| !!A ts. ts <= A ==> P A ts; + !!A pn. !b:bdy pn. P (insert (mgf_call pn) A) {mgf b} ==> P A {mgf_call pn}; + !!A t. !pn:U. P A {mgf_call pn} ==> P A {mgf t}; + finite U; uA = mgf_call`U |] ==> + !A. A <= uA --> n <= card uA --> card A = card uA - n --> (!t. P A {mgf t})" +proof - + assume ax_derivs_asm: "!!A ts. ts <= A ==> P A ts" + assume MGF_nested_Methd: "!!A pn. !b:bdy pn. P (insert (mgf_call pn) A) + {mgf b} ==> P A {mgf_call pn}" + assume MGF_asm: "!!A t. !pn:U. P A {mgf_call pn} ==> P A {mgf t}" + assume "finite U" "uA = mgf_call`U" + then show ?thesis + apply - + apply (induct_tac "n") + apply (tactic "ALLGOALS Clarsimp_tac") + apply (tactic "dtac (permute_prems 0 1 card_seteq) 1") + apply simp + apply (erule finite_imageI) + apply (simp add: MGF_asm ax_derivs_asm) + apply (rule MGF_asm) + apply (rule ballI) + apply (case_tac "mgf_call pn : A") + apply (fast intro: ax_derivs_asm) + apply (rule MGF_nested_Methd) + apply (rule ballI) + apply (drule spec, erule impE, erule_tac [2] impE, erule_tac [3] impE, + erule_tac [4] spec) + apply fast + apply (erule Suc_leD) + apply (drule finite_subset) + apply (erule finite_imageI) + apply auto + apply arith + done +qed + +lemma nesting_lemma [rule_format (no_asm)]: "[| !!A ts. ts <= A ==> P A ts; + !!A pn. !b:bdy pn. P (insert (mgf (f pn)) A) {mgf b} ==> P A {mgf (f pn)}; + !!A t. !pn:U. P A {mgf (f pn)} ==> P A {mgf t}; + finite U |] ==> P {} {mgf t}" +proof - + assume 2: "!!A pn. !b:bdy pn. P (insert (mgf (f pn)) A) {mgf b} ==> P A {mgf (f pn)}" + assume 3: "!!A t. !pn:U. P A {mgf (f pn)} ==> P A {mgf t}" + assume "!!A ts. ts <= A ==> P A ts" "finite U" + then show ?thesis + apply - + apply (rule_tac mgf = "mgf" in nesting_lemma') + apply (erule_tac [2] 2) + apply (rule_tac [2] 3) + apply (rule_tac [6] le_refl) + apply auto + done +qed + +lemma MGF_nested_Methd: "\ + G,insert ({Normal \} In1l (Methd C sig) \{G\}) A\ + {Normal \} In1l (body G C sig) \{G\} + \ \ G,A\{Normal \} In1l (Methd C sig) \{G\}" +apply (unfold MGF_def) +apply (rule ax_MethdN) +apply (erule conseq2) +apply clarsimp +apply (erule MethdI) +done + +lemma MGF_deriv: "ws_prog G \ G,({}::state triple set)\{\} t\ {G\}" +apply (rule MGFNormalI) +apply (rule_tac mgf = "\t. {Normal \} t\ {G\}" and + bdy = "\ (C,sig) .{In1l (body G C sig) }" and + f = "\ (C,sig) . In1l (Methd C sig) " in nesting_lemma) +apply (erule ax_derivs.asm) +apply (clarsimp simp add: split_tupled_all) +apply (erule MGF_nested_Methd) +apply (erule_tac [2] finite_is_methd) +apply (rule MGF_asm [THEN MGFNormalD]) +apply clarify +apply (rule MGFNormalI) +apply force +done + + +section "simultaneous version" + +lemma MGF_simult_Methd_lemma: "finite ms \ + G,A\ (\(C,sig). {Normal \} In1l (Methd C sig)\ {G\}) ` ms + |\(\(C,sig). {Normal \} In1l (body G C sig)\ {G\}) ` ms \ + G,A|\(\(C,sig). {Normal \} In1l (Methd C sig)\ {G\}) ` ms" +apply (unfold MGF_def) +apply (rule ax_derivs.Methd [unfolded mtriples_def]) +apply (erule ax_finite_pointwise) +prefer 2 +apply (rule ax_derivs.asm) +apply fast +apply clarsimp +apply (rule conseq2) +apply (erule (1) ax_methods_spec) +apply clarsimp +apply (erule eval_Methd) +done + +lemma MGF_simult_Methd: "ws_prog G \ + G,({}::state triple set)|\(\(C,sig). {Normal \} In1l (Methd C sig)\ {G\}) + ` Collect (split (is_methd G)) " +apply (frule finite_is_methd) +apply (rule MGF_simult_Methd_lemma) +apply assumption +apply (erule ax_finite_pointwise) +prefer 2 +apply (rule ax_derivs.asm) +apply blast +apply clarsimp +apply (rule MGF_asm [THEN MGFNormalD]) +apply clarify +apply (rule MGFNormalI) +apply force +done + +lemma MGF_deriv: "ws_prog G \ G,({}::state triple set)\{\} t\ {G\}" +apply (rule MGF_asm) +apply (intro strip) +apply (rule MGFNormalI) +apply (rule ax_derivs.weaken) +apply (erule MGF_simult_Methd) +apply force +done + + +section "corollaries" + +lemma MGF_complete: "G,{}\{P} t\ {Q} \ G,({}::state triple set)\{\} t\ {G\} \ + G,({}::state triple set)\{P::state assn} t\ {Q}" +apply (rule ax_no_hazard) +apply (unfold MGF_def) +apply (erule conseq12) +apply (simp (no_asm_use) add: ax_valids_def triple_valid_def) +apply (fast dest!: eval_evaln) +done + +theorem ax_complete: "ws_prog G \ + G,{}\{P::state assn} t\ {Q} \ G,({}::state triple set)\{P} t\ {Q}" +apply (erule MGF_complete) +apply (erule MGF_deriv) +done + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/AxExample.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/AxExample.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,270 @@ +(* Title: isabelle/Bali/AxExample.thy + ID: $Id$ + Author: David von Oheimb + Copyright 2000 Technische Universitaet Muenchen +*) +header {* Example of a proof based on the Bali axiomatic semantics *} + +theory AxExample = AxSem + Example: + +constdefs + arr_inv :: "st \ bool" + "arr_inv \ \s. \obj a T el. globs s (Stat Base) = Some obj \ + values obj (Inl (arr, Base)) = Some (Addr a) \ + heap s a = Some \tag=Arr T 2,values=el\" + +lemma arr_inv_new_obj: +"\a. \arr_inv s; new_Addr (heap s)=Some a\ \ arr_inv (gupd(Inl a\x) s)" +apply (unfold arr_inv_def) +apply (force dest!: new_AddrD2) +done + +lemma arr_inv_set_locals [simp]: "arr_inv (set_locals l s) = arr_inv s" +apply (unfold arr_inv_def) +apply (simp (no_asm)) +done + +lemma arr_inv_gupd_Stat [simp]: + "Base \ C \ arr_inv (gupd(Stat C\obj) s) = arr_inv s" +apply (unfold arr_inv_def) +apply (simp (no_asm_simp)) +done + +lemma ax_inv_lupd [simp]: "arr_inv (lupd(x\y) s) = arr_inv s" +apply (unfold arr_inv_def) +apply (simp (no_asm)) +done + + +declare split_if_asm [split del] +declare lvar_def [simp] + +ML {* +fun inst1_tac s t = instantiate_tac [(s,t)]; +val ax_tac = REPEAT o rtac allI THEN' + resolve_tac(thm "ax_Skip"::thm "ax_StatRef"::thm "ax_MethdN":: + thm "ax_Alloc"::thm "ax_Alloc_Arr":: + thm "ax_SXAlloc_Normal":: + funpow 7 tl (thms "ax_derivs.intros")) +*} + + +theorem ax_test: "tprg,({}::'a triple set)\ + {Normal (\Y s Z::'a. heap_free four s \ \initd Base s \ \ initd Ext s)} + .test [Class Base]. {\Y s Z. abrupt s = Some (Xcpt (Std IndOutBound))}" +apply (unfold test_def arr_viewed_from_def) +apply (tactic "ax_tac 1" (*;;*)) +defer +apply (tactic "ax_tac 1" (* Try *)) +defer +apply (tactic {* inst1_tac "Q1" + "\Y s Z. arr_inv (snd s) \ tprg,s\catch SXcpt NullPointer" *}) +prefer 2 +apply simp +apply (rule_tac P' = "Normal (\Y s Z. arr_inv (snd s))" in conseq1) +prefer 2 +apply clarsimp +apply (rule_tac Q' = "(\Y s Z. ?Q Y s Z)\=False\=\" in conseq2) +prefer 2 +apply simp +apply (tactic "ax_tac 1" (* While *)) +prefer 2 +apply (rule ax_impossible [THEN conseq1], clarsimp) +apply (rule_tac P' = "Normal ?P" in conseq1) +prefer 2 +apply clarsimp +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1" (* AVar *)) +prefer 2 +apply (rule ax_subst_Val_allI) +apply (tactic {* inst1_tac "P'21" "\u a. Normal (?PP a\?x) u" *}) +apply (simp del: avar_def2 peek_and_def2) +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1") + (* just for clarification: *) +apply (rule_tac Q' = "Normal (\Var:(v, f) u ua. fst (snd (avar tprg (Intg 2) v u)) = Some (Xcpt (Std IndOutBound)))" in conseq2) +prefer 2 +apply (clarsimp simp add: split_beta) +apply (tactic "ax_tac 1" (* FVar *)) +apply (tactic "ax_tac 2" (* StatRef *)) +apply (rule ax_derivs.Done [THEN conseq1]) +apply (clarsimp simp add: arr_inv_def inited_def in_bounds_def) +defer +apply (rule ax_SXAlloc_catch_SXcpt) +apply (rule_tac Q' = "(\Y (x, s) Z. x = Some (Xcpt (Std NullPointer)) \ arr_inv s) \. heap_free two" in conseq2) +prefer 2 +apply (simp add: arr_inv_new_obj) +apply (tactic "ax_tac 1") +apply (rule_tac C = "Ext" in ax_Call_known_DynT) +apply (unfold DynT_prop_def) +apply (simp (no_asm)) +apply (intro strip) +apply (rule_tac P' = "Normal ?P" in conseq1) +apply (tactic "ax_tac 1" (* Methd *)) +apply (rule ax_thin [OF _ empty_subsetI]) +apply (simp (no_asm) add: body_def2) +apply (tactic "ax_tac 1" (* Body *)) +(* apply (rule_tac [2] ax_derivs.Abrupt) *) +defer +apply (simp (no_asm)) +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1") (* Ass *) +prefer 2 +apply (rule ax_subst_Var_allI) +apply (tactic {* inst1_tac "P'27" "\a vs l vf. ?PP a vs l vf\?x \. ?p" *}) +apply (rule allI) +apply (tactic {* simp_tac (simpset() delloop "split_all_tac" delsimps [thm "peek_and_def2"]) 1 *}) +apply (rule ax_derivs.Abrupt) +apply (simp (no_asm)) +apply (tactic "ax_tac 1" (* FVar *)) +apply (tactic "ax_tac 2", tactic "ax_tac 2", tactic "ax_tac 2") +apply (tactic "ax_tac 1") +apply clarsimp +apply (tactic {* inst1_tac "R14" "\a'. Normal ((\Vals:vs (x, s) Z. arr_inv s \ inited Ext (globs s) \ a' \ Null \ hd vs = Null) \. heap_free two)" *}) +prefer 5 +apply (rule ax_derivs.Done [THEN conseq1], force) +apply force +apply (rule ax_subst_Val_allI) +apply (tactic {* inst1_tac "P'33" "\u a. Normal (?PP a\?x) u" *}) +apply (simp (no_asm) del: peek_and_def2) +apply (tactic "ax_tac 1") +prefer 2 +apply (rule ax_subst_Val_allI) +apply (tactic {* inst1_tac "P'4" "\aa v. Normal (?QQ aa v\?y)" *}) +apply (simp del: peek_and_def2) +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1") +(* end method call *) +apply (simp (no_asm)) + (* just for clarification: *) +apply (rule_tac Q' = "Normal ((\Y (x, s) Z. arr_inv s \ (\a. the (locals s (VName e)) = Addr a \ obj_class (the (globs s (Inl a))) = Ext \ + invocation_declclass tprg IntVir s (the (locals s (VName e))) (ClassT Base) + \name = foo, parTs = [Class Base]\ = Ext)) \. initd Ext \. heap_free two)" + in conseq2) +prefer 2 +apply clarsimp +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1") +defer +apply (rule ax_subst_Var_allI) +apply (tactic {* inst1_tac "P'14" "\u vf. Normal (?PP vf \. ?p) u" *}) +apply (simp (no_asm) del: split_paired_All peek_and_def2) +apply (tactic "ax_tac 1" (* NewC *)) +apply (tactic "ax_tac 1" (* ax_Alloc *)) + (* just for clarification: *) +apply (rule_tac Q' = "Normal ((\Y s Z. arr_inv (store s) \ vf=lvar (VName e) (store s)) \. heap_free tree \. initd Ext)" in conseq2) +prefer 2 +apply (simp add: invocation_declclass_def dynmethd_def) +apply (unfold dynlookup_def) +apply (simp add: dynmethd_Ext_foo) +apply (force elim!: arr_inv_new_obj atleast_free_SucD atleast_free_weaken) + (* begin init *) +apply (rule ax_InitS) +apply force +apply (simp (no_asm)) +apply (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *}) +apply (rule ax_Init_Skip_lemma) +apply (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *}) +apply (rule ax_InitS [THEN conseq1] (* init Base *)) +apply force +apply (simp (no_asm)) +apply (unfold arr_viewed_from_def) +apply (rule allI) +apply (rule_tac P' = "Normal ?P" in conseq1) +apply (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *}) +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1") +apply (rule_tac [2] ax_subst_Var_allI) +apply (tactic {* inst1_tac "P'29" "\vf l vfa. Normal (?P vf l vfa)" *}) +apply (tactic {* simp_tac (simpset() delloop "split_all_tac" delsimps [split_paired_All, thm "peek_and_def2"]) 2 *}) +apply (tactic "ax_tac 2" (* NewA *)) +apply (tactic "ax_tac 3" (* ax_Alloc_Arr *)) +apply (tactic "ax_tac 3") +apply (tactic {* inst1_tac "P" "\vf l vfa. Normal (?P vf l vfa\\)" *}) +apply (tactic {* simp_tac (simpset() delloop "split_all_tac") 2 *}) +apply (tactic "ax_tac 2") +apply (tactic "ax_tac 1" (* FVar *)) +apply (tactic "ax_tac 2" (* StatRef *)) +apply (rule ax_derivs.Done [THEN conseq1]) +apply (tactic {* inst1_tac "Q22" "\vf. Normal ((\Y s Z. vf=lvar (VName e) (snd s)) \. heap_free four \. initd Base \. initd Ext)" *}) +apply (clarsimp split del: split_if) +apply (frule atleast_free_weaken [THEN atleast_free_weaken]) +apply (drule initedD) +apply (clarsimp elim!: atleast_free_SucD simp add: arr_inv_def) +apply force +apply (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *}) +apply (rule ax_triv_Init_Object [THEN peek_and_forget2, THEN conseq1]) +apply (rule wf_tprg) +apply clarsimp +apply (tactic {* inst1_tac "P22" "\vf. Normal ((\Y s Z. vf = lvar (VName e) (snd s)) \. heap_free four \. initd Ext)" *}) +apply clarsimp +apply (tactic {* inst1_tac "PP" "\vf. Normal ((\Y s Z. vf = lvar (VName e) (snd s)) \. heap_free four \. Not \ initd Base)" *}) +apply clarsimp + (* end init *) +apply (rule conseq1) +apply (tactic "ax_tac 1") +apply clarsimp +done + +(* +while (true) { + if (i) {throw xcpt;} + else i=j +} +*) +lemma Loop_Xcpt_benchmark: + "Q = (\Y (x,s) Z. x \ None \ the_Bool (the (locals s i))) \ + G,({}::'a triple set)\{Normal (\Y s Z::'a. True)} + .lab1\ While(Lit (Bool True)) (If(Acc (LVar i)) (Throw (Acc (LVar xcpt))) Else + (Expr (Ass (LVar i) (Acc (LVar j))))). {Q}" +apply (rule_tac P' = "Q" and Q' = "Q\=False\=\" in conseq12) +apply safe +apply (tactic "ax_tac 1" (* Loop *)) +apply (rule ax_Normal_cases) +prefer 2 +apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def) +apply (rule conseq1) +apply (tactic "ax_tac 1") +apply clarsimp +prefer 2 +apply clarsimp +apply (tactic "ax_tac 1" (* If *)) +apply (tactic + {* inst1_tac "P'21" "Normal (\s.. (\Y s Z. True)\=Val (the (locals s i)))" *}) +apply (tactic "ax_tac 1") +apply (rule conseq1) +apply (tactic "ax_tac 1") +apply clarsimp +apply (rule allI) +apply (rule ax_escape) +apply auto +apply (rule conseq1) +apply (tactic "ax_tac 1" (* Throw *)) +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1") +apply clarsimp +apply (rule_tac Q' = "Normal (\Y s Z. True)" in conseq2) +prefer 2 +apply clarsimp +apply (rule conseq1) +apply (tactic "ax_tac 1") +apply (tactic "ax_tac 1") +prefer 2 +apply (rule ax_subst_Var_allI) +apply (tactic {* inst1_tac "P'29" "\b Y ba Z vf. \Y (x,s) Z. x=None \ snd vf = snd (lvar i s)" *}) +apply (rule allI) +apply (rule_tac P' = "Normal ?P" in conseq1) +prefer 2 +apply clarsimp +apply (tactic "ax_tac 1") +apply (rule conseq1) +apply (tactic "ax_tac 1") +apply clarsimp +apply (tactic "ax_tac 1") +apply clarsimp +done + +end + diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/AxSem.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/AxSem.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,1191 @@ +(* Title: isabelle/Bali/AxSem.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1998 Technische Universitaet Muenchen +*) + +header {* Axiomatic semantics of Java expressions and statements + (see also Eval.thy) + *} + +theory AxSem = Evaln + TypeSafe: + +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) \ the_In1" + "\Var:v . b" == "(\v. b) \ the_In2" + "\Vals:v. b" == "(\v. b) \ 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 \ 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" ("_\_" [60,61] 60) + "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 + +(*###Do not work for some strange (unification?) reason +lemma subst_res_Val_beta [simp]: "(\Y. P (the_In1 Y))\Val v = (\Y. P v)" +apply (rule ext) +by simp + +lemma subst_res_Var_beta [simp]: "(\Y. P (the_In2 Y))\Var vf = (\Y. P vf)"; +apply (rule ext) +by simp + +lemma subst_res_Vals_beta [simp]: "(\Y. P (the_In3 Y))\Vals vs = (\Y. P vs)"; +apply (rule ext) +by simp +*) + +subsection "subst-Bool" + +constdefs + subst_Bool :: "'a assn \ bool \ 'a assn" ("_\=_" [60,61] 60) + "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" + +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\w = P w\w" +apply (rule ext) +apply (simp (no_asm)) +done + +(* unused *) +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" + +constdefs + ign_res :: " 'a assn \ 'a assn" ("_\" [1000] 1000) + "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" + +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)\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" + +constdefs + ign_res_eq :: "'a assn \ res \ 'a assn" ("_\=_" [60,61] 60) + "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 + +(* unused *) +lemma ign_res_eq_subst_res: "P\=w\w = P\" +apply (rule ext) +apply (rule ext) +apply (rule ext) +apply (simp (no_asm)) +done + +(* unused *) +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" + +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\s \halloc otag\a\ s'\ P (Val (Addr a)) s' Z" + + SXAlloc :: "prog \ 'a assn \ 'a assn" + "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 + +section "validity" + +constdefs + type_ok :: "prog \ term \ state \ bool" + "type_ok G t s \ \L T C. (normal s \ \prg=G,cls=C,lcl=L\\t\T) \ s\\(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_)}/ _\/ {(1_)}" [3,65,3] 75) + 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) + +translations + "{P} e-\ {Q}" == "{P} In1l e\ {Q}" + "{P} e=\ {Q}" == "{P} In2 e\ {Q}" + "{P} e\\ {Q}" == "{P} In3 e\ {Q}" + "{P} .c. {Q}" == "{P} In1r c\ {Q}" + +lemma inj_triple: "inj (\(P,t,Q). {P} t\ {Q})" +apply (rule injI) +apply auto +done + +lemma triple_inj_eq: "({P} t\ {Q} = {P'} t'\ {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_)}/ _-\/ {(1_)} | _}"[3,65,3,65]75) + "{{P} tf-\ {Q} | ms} \ (\(C,sig). {Normal(P C sig)} tf C sig-\ {Q C sig})`ms" + +consts + + triple_valid :: "prog \ nat \ 'a triple \ bool" + ( "_\_:_" [61,0, 58] 57) + ax_valids :: "prog \ 'b triples \ 'a triples \ bool" + ("_,_|\_" [61,58,58] 57) + ax_derivs :: "prog \ ('b triples \ 'a triples) set" + +syntax + + triples_valid:: "prog \ nat \ 'a triples \ bool" + ( "_||=_:_" [61,0, 58] 57) + ax_valid :: "prog \ 'b triples \ 'a triple \ bool" + ( "_,_|=_" [61,58,58] 57) + ax_Derivs:: "prog \ 'b triples \ 'a triples \ bool" + ("_,_||-_" [61,58,58] 57) + ax_Deriv :: "prog \ 'b triples \ 'a triple \ bool" + ( "_,_|-_" [61,58,58] 57) + +syntax (xsymbols) + + triples_valid:: "prog \ nat \ 'a triples \ bool" + ( "_|\_:_" [61,0, 58] 57) + ax_valid :: "prog \ 'b triples \ 'a triple \ bool" + ( "_,_\_" [61,58,58] 57) + ax_Derivs:: "prog \ 'b triples \ 'a triples \ bool" + ("_,_|\_" [61,58,58] 57) + ax_Deriv :: "prog \ 'b triples \ 'a triple \ bool" + ( "_,_\_" [61,58,58] 57) + +defs triple_valid_def: "G\n:t \ case t of {P} t\ {Q} \ + \Y s Z. P Y s Z \ type_ok G t s \ + (\Y' s'. G\s \t\\n\ (Y',s') \ Q Y' s' Z)" +translations "G|\n:ts" == "Ball ts (triple_valid G n)" +defs ax_valids_def:"G,A|\ts \ \n. G|\n:A \ G|\n:ts" +translations "G,A \t" == "G,A|\{t}" + "G,A|\ts" == "(A,ts) \ ax_derivs G" + "G,A \t" == "G,A|\{t}" + +lemma triple_valid_def2: "G\n:{P} t\ {Q} = + (\Y s Z. P Y s Z + \ (\L. (normal s \ (\T C. \prg=G,cls=C,lcl=L\\t\T)) \ s\\(G,L)) \ + (\Y' s'. G\s \t\\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] +ML_setup {* +simpset_ref() := simpset() delloop "split_all_tac"; +claset_ref () := claset () delSWrapper "split_all_tac" +*} + + +inductive "ax_derivs G" intros + + empty: " G,A|\{}" + insert:"\G,A\t; G,A|\ts\ \ + G,A|\insert t ts" + + asm: "ts\A \ G,A|\ts" + +(* could be added for convenience and efficiency, but is not necessary + cut: "\G,A'|\ts; G,A|\A'\ \ + G,A |\ts" +*) + weaken:"\G,A|\ts'; ts \ ts'\ \ G,A|\ts" + + conseq:"\Y s Z . P Y s Z \ (\P' Q'. G,A\{P'} t\ {Q'} \ (\Y' s'. + (\Y Z'. P' Y s Z' \ Q' Y' s' Z') \ + Q Y' s' Z )) + \ G,A\{P } t\ {Q }" + + hazard:"G,A\{P \. Not \ type_ok G t} t\ {Q}" + + Abrupt: "G,A\{P\(arbitrary3 t) \. Not \ normal} t\ {P}" + + (* variables *) + LVar: " G,A\{Normal (\s.. P\Var (lvar vn s))} LVar vn=\ {P}" + + FVar: "\G,A\{Normal P} .Init C. {Q}; + G,A\{Q} e-\ {\Val:a:. fvar C stat fn a ..; R}\ \ + G,A\{Normal P} {C,stat}e..fn=\ {R}" + + AVar: "\G,A\{Normal P} e1-\ {Q}; + \a. G,A\{Q\Val a} e2-\ {\Val:i:. avar G i a ..; R}\ \ + G,A\{Normal P} e1.[e2]=\ {R}" + (* expressions *) + + NewC: "\G,A\{Normal P} .Init C. {Alloc G (CInst C) Q}\ \ + G,A\{Normal P} NewC C-\ {Q}" + + NewA: "\G,A\{Normal P} .init_comp_ty T. {Q}; G,A\{Q} e-\ + {\Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\ \ + G,A\{Normal P} New T[e]-\ {R}" + + Cast: "\G,A\{Normal P} e-\ {\Val:v:. \s.. + abupd (raise_if (\G,s\v fits T) ClassCast) .; Q\Val v}\ \ + G,A\{Normal P} Cast T e-\ {Q}" + + Inst: "\G,A\{Normal P} e-\ {\Val:v:. \s.. + Q\Val (Bool (v\Null \ G,s\v fits RefT T))}\ \ + G,A\{Normal P} e InstOf T-\ {Q}" + + Lit: "G,A\{Normal (P\Val v)} Lit v-\ {P}" + + Super:" G,A\{Normal (\s.. P\Val (val_this s))} Super-\ {P}" + + Acc: "\G,A\{Normal P} va=\ {\Var:(v,f):. Q\Val v}\ \ + G,A\{Normal P} Acc va-\ {Q}" + + Ass: "\G,A\{Normal P} va=\ {Q}; + \vf. G,A\{Q\Var vf} e-\ {\Val:v:. assign (snd vf) v .; R}\ \ + G,A\{Normal P} va:=e-\ {R}" + + Cond: "\G,A \{Normal P} e0-\ {P'}; + \b. G,A\{P'\=b} (if b then e1 else e2)-\ {Q}\ \ + G,A\{Normal P} e0 ? e1 : e2-\ {Q}" + + Call: +"\G,A\{Normal P} e-\ {Q}; \a. G,A\{Q\Val a} args\\ {R a}; + \a vs invC declC l. G,A\{(R a\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\mode\invC\statT)} + Methd declC \name=mn,parTs=pTs\-\ {set_lvars l .; S}\ \ + G,A\{Normal P} {statT,mode}e\mn({pTs}args)-\ {S}" + + Methd:"\G,A\ {{P} Methd-\ {Q} | ms} |\ {{P} body G-\ {Q} | ms}\ \ + G,A|\{{P} Methd-\ {Q} | ms}" + + Body: "\G,A\{Normal P} .Init D. {Q}; + G,A\{Q} .c. {\s.. abupd (absorb Ret) .; R\(In1 (the (locals s Result)))}\ + \ + G,A\{Normal P} Body D c-\ {R}" + + (* expression lists *) + + Nil: "G,A\{Normal (P\Vals [])} []\\ {P}" + + Cons: "\G,A\{Normal P} e-\ {Q}; + \v. G,A\{Q\Val v} es\\ {\Vals:vs:. R\Vals (v#vs)}\ \ + G,A\{Normal P} e#es\\ {R}" + + (* statements *) + + Skip: "G,A\{Normal (P\\)} .Skip. {P}" + + Expr: "\G,A\{Normal P} e-\ {Q\\}\ \ + G,A\{Normal P} .Expr e. {Q}" + + Lab: "\G,A\{Normal P} .c. {abupd (absorb (Break l)) .; Q}\ \ + G,A\{Normal P} .l\ c. {Q}" + + Comp: "\G,A\{Normal P} .c1. {Q}; + G,A\{Q} .c2. {R}\ \ + G,A\{Normal P} .c1;;c2. {R}" + + If: "\G,A \{Normal P} e-\ {P'}; + \b. G,A\{P'\=b} .(if b then c1 else c2). {Q}\ \ + G,A\{Normal P} .If(e) c1 Else c2. {Q}" +(* unfolding variant of Loop, not needed here + LoopU:"\G,A \{Normal P} e-\ {P'}; + \b. G,A\{P'\=b} .(if b then c;;While(e) c else Skip).{Q}\ + \ G,A\{Normal P} .While(e) c. {Q}" +*) + Loop: "\G,A\{P} e-\ {P'}; + G,A\{Normal (P'\=True)} .c. {abupd (absorb (Cont l)) .; P}\ \ + G,A\{P} .l\ While(e) c. {(P'\=False)\=\}" +(** Beware of polymorphic_Loop below: should be identical terms **) + + Do: "G,A\{Normal (abupd (\a. (Some (Jump j))) .; P\\)} .Do j. {P}" + + Throw:"\G,A\{Normal P} e-\ {\Val:a:. abupd (throw a) .; Q\\}\ \ + G,A\{Normal P} .Throw e. {Q}" + + Try: "\G,A\{Normal P} .c1. {SXAlloc G Q}; + G,A\{Q \. (\s. G,s\catch C) ;. new_xcpt_var vn} .c2. {R}; + (Q \. (\s. \G,s\catch C)) \ R\ \ + G,A\{Normal P} .Try c1 Catch(C vn) c2. {R}" + + Fin: "\G,A\{Normal P} .c1. {Q}; + \x. G,A\{Q \. (\s. x = fst s) ;. abupd (\x. None)} + .c2. {abupd (abrupt_if (x\None) x) .; R}\ \ + G,A\{Normal P} .c1 Finally c2. {R}" + + Done: "G,A\{Normal (P\\ \. initd C)} .Init C. {P}" + + Init: "\the (class G C) = c; + G,A\{Normal ((P \. Not \ initd C) ;. supd (init_class_obj G C))} + .(if C = Object then Skip else Init (super c)). {Q}; + \l. G,A\{Q \. (\s. l = locals (store s)) ;. set_lvars empty} + .init c. {set_lvars l .; R}\ \ + G,A\{Normal (P \. Not \ initd C)} .Init C. {R}" + +axioms (** these terms are the same as above, but with generalized typing **) + polymorphic_conseq: + "\Y s Z . P Y s Z \ (\P' Q'. G,A\{P'} t\ {Q'} \ (\Y' s'. + (\Y Z'. P' Y s Z' \ Q' Y' s' Z') \ + Q Y' s' Z )) + \ G,A\{P } t\ {Q }" + + polymorphic_Loop: + "\G,A\{P} e-\ {P'}; + G,A\{Normal (P'\=True)} .c. {abupd (absorb (Cont l)) .; P}\ \ + G,A\{P} .l\ While(e) c. {(P'\=False)\=\}" + +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'|\ts; G,A|\A'\ \ G,A|\ts" +apply (unfold ax_valids_def) +apply fast +done + +(*if cut is available +Goal "\G,A'|\ts; A' \ A; \P Q t. {P} t\ {Q} \ A' \ (\T. (G,L)\t\T) \ \ + G,A|\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)|\(ts::'a triple set) \ \A. A' \ A \ G,A|\ts" +apply (erule ax_derivs.induct) +apply (tactic "ALLGOALS(EVERY'[Clarify_tac,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) +(* 31 subgoals *) +prefer 16 (* Methd *) +apply (rule ax_derivs.Methd, drule spec, erule mp, fast) +apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) (thms "ax_derivs.intros")) + THEN_ALL_NEW Blast_tac) *}) +apply (erule ax_derivs.Call) +apply clarify +apply blast + +apply (rule allI)+ +apply (drule spec)+ +apply blast +done + +lemma ax_thin_insert: "G,(A::'a triple set)\(t::'a triple) \ G,insert x A\t" +apply (erule ax_thin) +apply fast +done + +lemma subset_mtriples_iff: + "ts \ {{P} mb-\ {Q} | ms} = (\ms'. ms'\ms \ ts = {{P} mb-\ {Q} | ms'})" +apply (unfold mtriples_def) +apply (rule subset_image_iff) +done + +lemma weaken: + "G,(A::'a triple set)|\(ts'::'a triple set) \ !ts. ts \ ts' \ G,A|\ts" +apply (erule ax_derivs.induct) +(*36 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 *)) +(*31 subgoals*) +apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) (thms "ax_derivs.intros")) + THEN_ALL_NEW Fast_tac) *}) +(*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" + +lemma conseq12: "\G,A\{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\{P ::'a assn} t\ {Q }" +apply (rule polymorphic_conseq) +apply clarsimp +apply blast +done + +(*unused, but nice variant*) +lemma conseq12': "\G,A\{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\{P } t\ {Q }" +apply (erule conseq12) +apply fast +done + +lemma conseq12_from_conseq12': "\G,A\{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\{P } t\ {Q }" +apply (erule conseq12') +apply blast +done + +lemma conseq1: "\G,A\{P'} t\ {Q}; P \ P'\ \ G,A\{P } t\ {Q}" +apply (erule conseq12) +apply blast +done + +lemma conseq2: "\G,A\{P} t\ {Q'}; Q' \ Q\ \ G,A\{P} t\ {Q}" +apply (erule conseq12) +apply blast +done + +lemma ax_escape: "\\Y s Z. P Y s Z \ G,A\{\Y' s' Z'. (Y',s') = (Y,s)} t\ {\Y s Z'. Q Y s Z}\ \ + G,A\{P} t\ {Q}" +apply (rule polymorphic_conseq) +apply force +done + +(* unused *) +lemma ax_constant: "\ C \ G,A\{P} t\ {Q}\ \ G,A\{\Y s Z. C \ P Y s Z} t\ {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\{\Y s Z. False} t\ {Q}" +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\{\Y s Z. (Y,s)=Z} t\ {\Y s Z. (Y,s)=Z} \ G,A\{P} t\ {P}" +apply (erule conseq12) +apply auto +apply (erule (1) ax_nochange_lemma) +done + +(* unused *) +lemma ax_trivial: "G,A\{P} t\ {\Y s Z. True}" +apply (rule polymorphic_conseq(* unused *)) +apply auto +done + +(* unused *) +lemma ax_disj: "\G,A\{P1} t\ {Q1}; G,A\{P2} t\ {Q2}\ \ + G,A\{\Y s Z. P1 Y s Z \ P2 Y s Z} t\ {\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\{P} .c1. {Q} \ G,A\{Q ;. f} .c2. {R}) = + (\Q'. G,A\{P} .c1. {f .; Q'} \ G,A\{Q'} .c2. {R})" +apply (best elim!: conseq1 conseq2) +done + +lemma ax_cases: "\G,A\{P \. C} t\ {Q}; + G,A\{P \. Not \ C} t\ {Q}\ \ G,A\{P} t\ {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\{P} t\ {Q} \ G,A\{adapt_pre P Q Q'} t\ {Q'}" +apply (unfold adapt_pre_def) +apply (erule conseq12) +apply fast +done + +lemma adapt_pre_adapts: "G,A\{P} t\ {Q} \ G,A\{adapt_pre P Q Q'} t\ {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\{P} t\ {Q} \ G,A\{P'} t\ {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 + +(* +Goal "\(A::'a triple set) t. G,A\{P} t\ {Q} \ G,A\{P'} t\ {Q'} \ + wf_prog G \ G,(A::'a triple set)\{P} t\ {Q::'a assn} \ G,A\{P'} t\ {Q'::'a assn}" +b y fatac ax_sound 1 1; +b y asm_full_simp_tac (simpset() addsimps [ax_valids_def,triple_valid_def2]) 1; +b y rtac ax_no_hazard 1; +b y etac conseq12 1; +b y Clarify_tac 1; +b y case_tac "\Z. \P Y s Z" 1; +b y smp_tac 2 1; +b y etac thin_rl 1; +b y etac thin_rl 1; +b y clarsimp_tac (claset(), simpset() addsimps [type_ok_def]) 1; +b y subgoal_tac "G|\n:A" 1; +b y smp_tac 1 1; +b y smp_tac 3 1; +b y etac impE 1; + back(); + b y Fast_tac 1; +b y +b y rotate_tac 2 1; +b y etac thin_rl 1; +b y etac thin_rl 2; +b y etac thin_rl 2; +b y Clarify_tac 2; +b y dtac spec 2; +b y EVERY'[dtac spec, mp_tac] 2; +b y thin_tac "\n Y s Z. ?PP n Y s Z" 2; +b y thin_tac "P' Y s Z" 2; +b y Blast_tac 2; +b y smp_tac 3 1; +b y case_tac "\Z. \P Y s Z" 1; +b y dres_inst_tac [("x","In1r Skip")] spec 1; +b y Full_simp_tac 1; +*) + +lemma peek_and_forget1_Normal: + "G,A\{Normal P} t\ {Q} \ G,A\{Normal (P \. p)} t\ {Q}" +apply (erule conseq1) +apply (simp (no_asm)) +done + +lemma peek_and_forget1: "G,A\{P} t\ {Q} \ G,A\{P \. p} t\ {Q}" +apply (erule conseq1) +apply (simp (no_asm)) +done + +lemmas ax_NormalD = peek_and_forget1 [of _ _ _ _ _ normal] + +lemma peek_and_forget2: "G,A\{P} t\ {Q \. p} \ G,A\{P} t\ {Q}" +apply (erule conseq2) +apply (simp (no_asm)) +done + +lemma ax_subst_Val_allI: "\v. G,A\{(P' v )\Val v} t\ {Q v} \ + \v. G,A\{(\w:. P' (the_In1 w))\Val v} t\ {Q v}" +apply (force elim!: conseq1) +done + +lemma ax_subst_Var_allI: "\v. G,A\{(P' v )\Var v} t\ {Q v} \ + \v. G,A\{(\w:. P' (the_In2 w))\Var v} t\ {Q v}" +apply (force elim!: conseq1) +done + +lemma ax_subst_Vals_allI: "(\v. G,A\{( P' v )\Vals v} t\ {Q v}) \ + \v. G,A\{(\w:. P' (the_In3 w))\Vals v} t\ {Q v}" +apply (force elim!: conseq1) +done + + +section "alternative axioms" + +lemma ax_Lit2: + "G,(A::'a triple set)\{Normal P::'a assn} Lit v-\ {Normal (P\=Val v)}" +apply (rule ax_derivs.Lit [THEN conseq1]) +apply force +done +lemma ax_Lit2_test_complete: + "G,(A::'a triple set)\{Normal (P\Val v)::'a assn} Lit v-\ {P}" +apply (rule ax_Lit2 [THEN conseq2]) +apply force +done + +lemma ax_LVar2: "G,(A::'a triple set)\{Normal P::'a assn} LVar vn=\ {Normal (\s.. P\=Var (lvar vn s))}" +apply (rule ax_derivs.LVar [THEN conseq1]) +apply force +done + +lemma ax_Super2: "G,(A::'a triple set)\ + {Normal P::'a assn} Super-\ {Normal (\s.. P\=Val (val_this s))}" +apply (rule ax_derivs.Super [THEN conseq1]) +apply force +done + +lemma ax_Nil2: + "G,(A::'a triple set)\{Normal P::'a assn} []\\ {Normal (P\=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)\{Normal (P C sig)::'a assn} mb C sig-\ {Q C sig}\ \ + G,A|\{{P} mb-\ {Q} | F}" +apply (frule (1) finite_subset) +apply (erule make_imp) +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)|\insert (t::'a triple) ts \ G,A\t \ G,A|\ts" +apply (fast intro: ax_derivs.weaken) +done + +lemma ax_methods_spec: +"\G,(A::'a triple set)|\split f ` ms; (C,sig) \ ms\\ G,A\((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)\(f C sig::'a triple)) \ (\(C,sig)\ms. G,A\(g C sig::'a triple))) \ + G,A|\split f ` F \ G,A|\split g ` F" +apply (frule (1) finite_subset) +apply (erule make_imp) +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)\{P \. type_ok G t} t\ {Q::'a assn} \ G,A\{P} t\ {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\\t\T) + \ G,(A::'a triple set)\{Normal P} t\ {Q::'a assn} \ + G,A\{Normal P} t\ {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 {* +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)\{P\\} .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\IntVir\C\statT; + \a vs l. G,A\{(R a\Vals vs \. (\s. l = locals (store 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\{Q\Val a} args\\ + {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)\{Normal P} e-\ {Q::'a assn}\ + \ G,A\{Normal P} {statT,IntVir}e\mn({pTs}args)-\ {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\{R a\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\-\ {set_lvars l .; S}; + G,A\{Normal P} e-\ {Q}; + \ a. G,(A::'a triple set)\{Q\Val a} args\\ {(R::val \ 'a assn) a + \. (\ s. C=invocation_declclass + G Static (store s) a statT \name=mn,parTs=pTs\)} +\ \ G,A\{Normal P} {statT,Static}e\mn({pTs}args)-\ {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) +done + +lemma ax_Methd1: + "\G,A\{{P} Methd-\ {Q} | ms}|\ {{P} body G-\ {Q} | ms}; (C,sig)\ ms\ \ + G,A\{Normal (P C sig)} Methd C sig-\ {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-\ {Q}) A\ + {Normal P} body G C sig-\ {Q} \ + G,A\{Normal P} Methd C sig-\ {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)\{Normal (P\Val Null)} StatRef rt-\ {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\{Q \. (\s. l = locals (store s)) ;. set_lvars empty} + .init c. {set_lvars l .; R}; + G,A\{Normal ((P \. Not \ initd C) ;. supd (init_class_obj G C))} + .Init (super c). {Q}\ \ + G,(A::'a triple set)\{Normal (P \. Not \ 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)\{P\\ \. (\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\\ \ (supd (init_class_obj G C) .; P); + G,A\{Normal (P \. initd C)} .Init (super c). {(P \. initd C)\\}\ \ + G,(A::'a triple set)\{Normal P\\} .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)\ + {Normal ((supd (init_class_obj G Object) .; P\\) \. Not \ 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)\{Normal P\\} .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\{P} .c. {Normal Q} \ G,A\{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\{P} t\ {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\{P} t\ {Alloc G (CInst C) Q}" +apply (erule conseq2) +apply (auto elim!: halloc_elim_cases) +done + +lemma ax_Alloc_Arr: + "G,A\{P} t\ {\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\{P} t\ {\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\{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)) (Heap a) s) Z)) + \. heap_free (Suc (Suc 0))}\ \ + G,A\{P} t\ {SXAlloc G (\Y s Z. Q Y s Z \ G,s\catch SXcpt xn)}" +apply (erule conseq2) +apply (auto elim!: sxalloc_elim_cases halloc_elim_cases) +done + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/AxSound.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/AxSound.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,434 @@ +(* Title: isabelle/Bali/AxSound.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1999 Technische Universitaet Muenchen +*) +header {* Soundness proof for Axiomatic semantics of Java expressions and + statements + *} + + + +theory AxSound = AxSem: + +section "validity" + +consts + + triple_valid2:: "prog \ nat \ 'a triple \ bool" + ( "_\_\_"[61,0, 58] 57) + ax_valids2:: "prog \ 'a triples \ 'a triples \ bool" + ("_,_|\\_" [61,58,58] 57) + +defs triple_valid2_def: "G\n\t \ case t of {P} t\ {Q} \ + \Y s Z. P Y s Z \ (\L. s\\(G,L) + \ (\T C. (normal s \ \prg=G,cls=C,lcl=L\\t\T) \ + (\Y' s'. G\s \t\\n\ (Y',s') \ Q Y' s' Z \ s'\\(G,L))))" + +defs ax_valids2_def: "G,A|\\ts \ \n. (\t\A. G\n\t) \ (\t\ts. G\n\t)" + +lemma triple_valid2_def2: "G\n\{P} t\ {Q} = + (\Y s Z. P Y s Z \ (\Y' s'. G\s \t\\n\ (Y',s')\ + (\L. s\\(G,L) \ (\T C. (normal s \ \prg=G,cls=C,lcl=L\\t\T) \ + Q Y' s' Z \ s'\\(G,L)))))" +apply (unfold triple_valid2_def) +apply (simp (no_asm) add: split_paired_All) +apply blast +done + +lemma triple_valid2_eq [rule_format (no_asm)]: + "wf_prog G ==> triple_valid2 G = triple_valid G" +apply (rule ext) +apply (rule ext) +apply (rule triple.induct) +apply (simp (no_asm) add: triple_valid_def2 triple_valid2_def2) +apply (rule iffI) +apply fast +apply clarify +apply (tactic "smp_tac 3 1") +apply (case_tac "normal s") +apply clarsimp +apply (blast dest: evaln_eval eval_type_sound [THEN conjunct1]) +apply clarsimp +done + +lemma ax_valids2_eq: "wf_prog G \ G,A|\\ts = G,A|\ts" +apply (unfold ax_valids_def ax_valids2_def) +apply (force simp add: triple_valid2_eq) +done + +lemma triple_valid2_Suc [rule_format (no_asm)]: "G\Suc n\t \ G\n\t" +apply (induct_tac "t") +apply (subst triple_valid2_def2) +apply (subst triple_valid2_def2) +apply (fast intro: evaln_nonstrict_Suc) +done + +lemma Methd_triple_valid2_0: "G\0\{Normal P} Methd C sig-\ {Q}" +apply (clarsimp elim!: evaln_elim_cases simp add: triple_valid2_def2) +done + +lemma Methd_triple_valid2_SucI: +"\G\n\{Normal P} body G C sig-\{Q}\ + \ G\Suc n\{Normal P} Methd C sig-\ {Q}" +apply (simp (no_asm_use) add: triple_valid2_def2) +apply (intro strip, tactic "smp_tac 3 1", clarify) +apply (erule wt_elim_cases, erule evaln_elim_cases) +apply (unfold body_def Let_def) +apply clarsimp +apply blast +done + +lemma triples_valid2_Suc: + "Ball ts (triple_valid2 G (Suc n)) \ Ball ts (triple_valid2 G n)" +apply (fast intro: triple_valid2_Suc) +done + +lemma "G|\n:insert t A = (G\n:t \ G|\n:A)"; +oops + + +section "soundness" + +lemma Methd_sound: +"\G,A\ {{P} Methd-\ {Q} | ms}|\\{{P} body G-\ {Q} | ms}\ \ + G,A|\\{{P} Methd-\ {Q} | ms}" +apply (unfold ax_valids2_def mtriples_def) +apply (rule allI) +apply (induct_tac "n") +apply (clarify, tactic {* pair_tac "x" 1 *}, simp (no_asm)) +apply (fast intro: Methd_triple_valid2_0) +apply (clarify, tactic {* pair_tac "xa" 1 *}, simp (no_asm)) +apply (drule triples_valid2_Suc) +apply (erule (1) notE impE) +apply (drule_tac x = na in spec) +apply (tactic {* auto_tac (claset() addSIs [thm "Methd_triple_valid2_SucI"], + simpset() addsimps [ball_Un] addloop ("split_all_tac", split_all_tac)) *}) +done + + +lemma valids2_inductI: "\s t n Y' s'. G\s\t\\n\ (Y',s') \ t = c \ + Ball A (triple_valid2 G n) \ (\Y Z. P Y s Z \ + (\L. s\\(G,L) \ (\T C. (normal s \ \prg=G,cls=C,lcl=L\\t\T) \ + Q Y' s' Z \ s'\\(G, L)))) \ + G,A|\\{ {P} c\ {Q}}" +apply (simp (no_asm) add: ax_valids2_def triple_valid2_def2) +apply clarsimp +done + +ML_setup {* +Delsimprocs [evaln_expr_proc,evaln_var_proc,evaln_exprs_proc,evaln_stmt_proc] +*} + +lemma Loop_sound: "\G,A|\\{ {P} e-\ {P'}}; + G,A|\\{ {Normal (P'\=True)} .c. {abupd (absorb (Cont l)) .; P}}\ \ + G,A|\\{ {P} .l\ While(e) c. {(P'\=False)\=\}}" +apply (rule valids2_inductI) +apply ((rule allI)+, rule impI, tactic {* pair_tac "s" 1*}, tactic {* pair_tac "s'" 1*}) +apply (erule evaln.induct) +apply simp_all (* takes half a minute *) +apply clarify +apply (erule_tac V = "G,A|\\{ {?P'} .c. {?P}}" in thin_rl) +apply (simp_all (no_asm_use) add: ax_valids2_def triple_valid2_def2) +apply (tactic "smp_tac 1 1", tactic "smp_tac 3 1", force) +apply clarify +apply (rule wt_elim_cases, assumption) +apply (tactic "smp_tac 1 1", tactic "smp_tac 1 1", tactic "smp_tac 3 1", + tactic "smp_tac 2 1", tactic "smp_tac 1 1") +apply (erule impE,simp (no_asm),blast) +apply (simp add: imp_conjL split_tupled_all split_paired_All) +apply (case_tac "the_Bool b") +apply clarsimp +apply (case_tac "a") +apply (simp_all) +apply clarsimp +apply (erule_tac V = "c = l\ While(e) c \ ?P" in thin_rl) +apply (blast intro: conforms_absorb) +apply blast+ +done + +declare subst_Bool_def2 [simp del] +lemma all_empty: "(!x. P) = P" +by simp +lemma sound_valid2_lemma: +"\\v n. Ball A (triple_valid2 G n) \ P v n; Ball A (triple_valid2 G n)\ + \P v n" +by blast +ML {* +val fullsimptac = full_simp_tac(simpset() delsimps [thm "all_empty"]); +val sound_prepare_tac = EVERY'[REPEAT o thin_tac "?x \ ax_derivs G", + full_simp_tac (simpset()addsimps[thm "ax_valids2_def",thm "triple_valid2_def2", + thm "imp_conjL"] delsimps[thm "all_empty"]), + Clarify_tac]; +val sound_elim_tac = EVERY'[eresolve_tac (thms "evaln_elim_cases"), + TRY o eresolve_tac (thms "wt_elim_cases"), fullsimptac, Clarify_tac]; +val sound_valid2_tac = REPEAT o FIRST'[smp_tac 1, + datac (thm "sound_valid2_lemma") 1]; +val sound_forw_hyp_tac = + EVERY'[smp_tac 3 + ORELSE' EVERY'[dtac spec,dtac spec, dtac spec,etac impE, Fast_tac] + ORELSE' EVERY'[dtac spec,dtac spec,etac impE, Fast_tac], + fullsimptac, + smp_tac 2,TRY o smp_tac 1, + TRY o EVERY'[etac impE, TRY o rtac impI, + atac ORELSE' (EVERY' [REPEAT o rtac exI,Blast_tac]), + fullsimptac, Clarify_tac, TRY o smp_tac 1]] +*} +(* ### rtac conjI,rtac HOL.refl *) +lemma Call_sound: +"\wf_prog G; G,A|\\{ {Normal P} e-\ {Q}}; \a. G,A|\\{ {Q\Val a} ps\\ {R a}}; + \a vs invC declC l. G,A|\\{ {(R a\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\mode\invC\statT)} + Methd declC \name=mn,parTs=pTs\-\ {set_lvars l .; S}}\ \ + G,A|\\{ {Normal P} {statT,mode}e\mn({pTs}ps)-\ {S}}" +apply (tactic "EVERY'[sound_prepare_tac, sound_elim_tac, sound_valid2_tac] 1") +apply (rename_tac x1 s1 x2 s2 ab bb v vs m pTsa statDeclC) +apply (tactic "smp_tac 6 1") +apply (tactic "sound_forw_hyp_tac 1") +apply (tactic "sound_forw_hyp_tac 1") +apply (drule max_spec2mheads) +apply (drule evaln_eval, drule (3) eval_ts) +apply (drule evaln_eval, frule (3) evals_ts) +apply (drule spec,erule impE,rule exI, blast) +(* apply (drule spec,drule spec,drule spec,erule impE,rule exI,blast) *) +apply (case_tac "if static m then x2 else (np a') x2") +defer 1 +apply (rename_tac x, subgoal_tac "(Some x, s2)\\(G, L)" (* used two times *)) +prefer 2 +apply (force split add: split_if_asm) +apply (simp del: if_raise_if_None) +apply (tactic "smp_tac 2 1") +apply (simp only: init_lvars_def2 invmode_Static_eq) +apply (clarsimp simp del: resTy_mthd) +apply (drule spec,erule swap,erule conforms_set_locals [OF _ lconf_empty]) +apply clarsimp +apply (drule Null_staticD) +apply (drule eval_gext', drule (1) conf_gext, frule (3) DynT_propI) +apply (erule impE) apply blast +apply (subgoal_tac + "G\invmode (mhd (statDeclC,m)) e + \invocation_class (invmode m e) s2 a' statT\statT") +defer apply simp +apply (drule (3) DynT_mheadsD,simp,simp) +apply (clarify, drule wf_mdeclD1, clarify) +apply (frule ty_expr_is_type) apply simp +apply (subgoal_tac "invmode (mhd (statDeclC,m)) e = IntVir \ a' \ Null") +defer apply simp +apply (frule (2) wt_MethdI) +apply clarify +apply (drule (2) conforms_init_lvars) +apply (simp) +apply (assumption)+ +apply simp +apply (assumption)+ +apply (rule impI) apply simp +apply simp +apply simp +apply (rule Ball_weaken) +apply assumption +apply (force simp add: is_acc_type_def) +apply (tactic "smp_tac 2 1") +apply simp +apply (tactic "smp_tac 1 1") +apply (erule_tac V = "?P \ ?Q" in thin_rl) +apply (erule impE) +apply (rule exI)+ +apply (subgoal_tac "is_static dm = (static m)") +prefer 2 apply (simp add: member_is_static_simp) +apply (simp only: ) +apply (simp only: sig.simps) +apply (force dest!: evaln_eval eval_gext' elim: conforms_return + del: impCE simp add: init_lvars_def2) +done + +lemma Init_sound: "\wf_prog G; the (class G C) = c; + G,A|\\{ {Normal ((P \. Not \ initd C) ;. supd (init_class_obj G C))} + .(if C = Object then Skip else Init (super c)). {Q}}; + \l. G,A|\\{ {Q \. (\s. l = locals (store s)) ;. set_lvars empty} + .init c. {set_lvars l .; R}}\ \ + G,A|\\{ {Normal (P \. Not \ initd C)} .Init C. {R}}" +apply (tactic "EVERY'[sound_prepare_tac, sound_elim_tac,sound_valid2_tac] 1") +apply (tactic {* instantiate_tac [("l24","\ n Y Z sa Y' s' L y a b aa ba ab bb. locals b")]*}) +apply (clarsimp simp add: split_paired_Ex) +apply (drule spec, drule spec, drule spec, erule impE) +apply (erule_tac V = "All ?P" in thin_rl, fast) +apply clarsimp +apply (tactic "smp_tac 2 1", drule spec, erule impE, + erule (3) conforms_init_class_obj) +apply (drule (1) wf_prog_cdecl) +apply (erule impE, rule exI,erule_tac V = "All ?P" in thin_rl, + force dest: wf_cdecl_supD split add: split_if simp add: is_acc_class_def) +apply clarify +apply (drule spec) +apply (drule spec) +apply (drule spec) +apply (erule impE) +apply ( fast) +apply (simp (no_asm_use) del: empty_def2) +apply (tactic "smp_tac 2 1") +apply (drule spec, erule impE, erule conforms_set_locals, rule lconf_empty) +apply (erule impE,rule impI,rule exI, erule wf_cdecl_wt_init) +apply clarsimp +apply (erule (1) conforms_return, force dest: evaln_eval eval_gext') +done + +lemma all_conjunct2: "\l. P' l \ P l \ \l. P l" +by fast + +lemma all4_conjunct2: + "\a vs D l. (P' a vs D l \ P a vs D l) \ \a vs D l. P a vs D l" +by fast + +lemmas sound_lemmas = Init_sound Loop_sound Methd_sound + +lemma ax_sound2: "wf_prog G \ G,A|\ts \ G,A|\\ts" +apply (erule ax_derivs.induct) +prefer 20 (* Call *) +apply (erule (1) Call_sound) apply simp apply force apply force + +apply (tactic {* TRYALL (eresolve_tac (thms "sound_lemmas") THEN_ALL_NEW + eresolve_tac [asm_rl, thm "all_conjunct2", thm "all4_conjunct2"]) *}) + +apply(tactic "COND (has_fewer_prems(30+3)) (ALLGOALS sound_prepare_tac) no_tac") + + (*empty*) +apply fast (* insert *) +apply fast (* asm *) +(*apply fast *) (* cut *) +apply fast (* weaken *) +apply (tactic "smp_tac 3 1", clarify, tactic "smp_tac 1 1", frule evaln_eval, +(* conseq *)case_tac"fst s",clarsimp simp add: eval_type_sound [THEN conjunct1], +clarsimp) +apply (simp (no_asm_use) add: type_ok_def, drule evaln_eval,fast) (* hazard *) +apply force (* Abrupt *) + +(* 25 subgoals *) +apply (tactic {* ALLGOALS sound_elim_tac *})(* LVar, Lit, Super, Nil, Skip,Do *) +apply (tactic {* ALLGOALS (asm_simp_tac (noAll_simpset() + delsimps [thm "all_empty"])) *}) (* Done *) +(* for FVar *) +apply (frule wf_ws_prog) +apply (frule ty_expr_is_type [THEN type_is_class, + THEN accfield_declC_is_class]) +apply (simp,simp,simp) +apply (frule_tac [4] wt_init_comp_ty) (* for NewA*) +apply (tactic "ALLGOALS sound_valid2_tac") +apply (tactic "TRYALL sound_forw_hyp_tac") (* Cast, Inst, Acc, Expr *) +apply (tactic {* TRYALL (EVERY'[dtac spec, TRY o EVERY'[rotate_tac ~1, + dtac spec], dtac conjunct2, smp_tac 1, + TRY o dres_inst_tac [("P","P'")] (thm "subst_Bool_the_BoolI")]) *}) +apply (frule_tac [14] x = x1 in conforms_NormI) (* for Fin *) + + +(* 15 subgoals *) +(* FVar *) +apply (tactic "sound_forw_hyp_tac 1") +apply (clarsimp simp add: fvar_def2 Let_def split add: split_if_asm) + +(* AVar *) +(* +apply (drule spec, drule spec, erule impE, fast) +apply (simp) +apply (tactic "smp_tac 2 1") +apply (tactic "smp_tac 1 1") +apply (erule impE) +apply (rule impI) +apply (rule exI)+ +apply blast +apply (clarsimp simp add: avar_def2) +*) +apply (tactic "sound_forw_hyp_tac 1") +apply (clarsimp simp add: avar_def2) + +(* NewC *) +apply (clarsimp simp add: is_acc_class_def) +apply (erule (1) halloc_conforms, simp, simp) + +(* NewA *) +apply (tactic "sound_forw_hyp_tac 1") +apply (rule conjI,blast) +apply (erule (1) halloc_conforms, simp, simp, simp add: is_acc_type_def) + +(* Ass *) +apply (tactic "sound_forw_hyp_tac 1") +apply (case_tac "aa") +prefer 2 +apply clarsimp +apply (drule evaln_eval)+ +apply (frule (3) eval_ts) +apply clarsimp +apply (frule (3) evar_ts [THEN conjunct2]) +apply (frule wf_ws_prog) +apply (drule (2) conf_widen) +apply (drule_tac "s1.0" = b in eval_gext') +apply (clarsimp simp add: assign_conforms_def) + +(* Cond *) + +apply (tactic "smp_tac 3 1") apply (tactic "smp_tac 2 1") +apply (tactic "smp_tac 1 1") apply (erule impE) +apply (rule impI,rule exI) +apply (rule_tac x = "if the_Bool b then T1 else T2" in exI) +apply (force split add: split_if) +apply assumption + +(* Body *) +apply (tactic "sound_forw_hyp_tac 1") +apply (rule conforms_absorb,assumption) + +(* Lab *) +apply (tactic "sound_forw_hyp_tac 1") +apply (rule conforms_absorb,assumption) + +(* If *) +apply (tactic "sound_forw_hyp_tac 1") +apply (tactic "sound_forw_hyp_tac 1") +apply (force split add: split_if) + +(* Throw *) +apply (drule evaln_eval, drule (3) eval_ts) +apply clarsimp +apply (drule (3) Throw_lemma) +apply clarsimp + +(* Try *) +apply (frule (1) sxalloc_type_sound) +apply (erule sxalloc_elim_cases2) +apply (tactic "smp_tac 3 1") +apply (clarsimp split add: option.split_asm) +apply (clarsimp split add: option.split_asm) +apply (tactic "smp_tac 1 1") +apply (simp only: split add: split_if_asm) +prefer 2 +apply (tactic "smp_tac 3 1", erule_tac V = "All ?P" in thin_rl, clarsimp) +apply (drule spec, erule_tac V = "All ?P" in thin_rl, drule spec, drule spec, + erule impE, force) +apply (frule (2) Try_lemma) +apply clarsimp +apply (fast elim!: conforms_deallocL) + +(* Fin *) +apply (tactic "sound_forw_hyp_tac 1") +apply (case_tac "x1", force) +apply clarsimp +apply (drule evaln_eval, drule (4) Fin_lemma) +done + + + +declare subst_Bool_def2 [simp] + +theorem ax_sound: + "wf_prog G \ G,(A::'a triple set)|\(ts::'a triple set) \ G,A|\ts" +apply (subst ax_valids2_eq [symmetric]) +apply assumption +apply (erule (1) ax_sound2) +done + + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Basis.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Basis.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,370 @@ +(* Title: isabelle/Bali/Basis.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen + +*) +header {* Definitions extending HOL as logical basis of Bali *} + +theory Basis = Main: + +ML_setup {* +Unify.search_bound := 40; +Unify.trace_bound := 40; + +quick_and_dirty:=true; + +Pretty.setmargin 77; +goals_limit:=2; +*} +(*print_depth 100;*) +(*Syntax.ambiguity_level := 1;*) + +section "misc" + +declare same_fstI [intro!] (*### TO HOL/Wellfounded_Relations *) + +(* ###TO HOL/???.ML?? *) +ML {* +fun make_simproc name pat pred thm = Simplifier.mk_simproc name + [Thm.read_cterm (Thm.sign_of_thm thm) (pat, HOLogic.typeT)] + (K (K (fn s => if pred s then None else Some (standard (mk_meta_eq thm))))) +*} + +declare split_if_asm [split] option.split [split] option.split_asm [split] +ML {* +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac) +*} +declare if_weak_cong [cong del] option.weak_case_cong [cong del] +declare length_Suc_conv [iff]; + +(*###to be phased out *) +ML {* +bind_thm ("make_imp", rearrange_prems [1,0] mp) +*} + +lemma Collect_split_eq: "{p. P (split f p)} = {(a,b). P (f a b)}" +apply auto +done + +lemma subset_insertD: + "A <= insert x B ==> A <= B & x ~: A | (EX B'. A = insert x B' & B' <= B)" +apply (case_tac "x:A") +apply (rule disjI2) +apply (rule_tac x = "A-{x}" in exI) +apply fast+ +done + +syntax + "3" :: nat ("3") + "4" :: nat ("4") +translations + "3" == "Suc 2" + "4" == "Suc 3" + +(*unused*) +lemma range_bool_domain: "range f = {f True, f False}" +apply auto +apply (case_tac "xa") +apply auto +done + +(* context (theory "Transitive_Closure") *) +lemma irrefl_tranclI': "r^-1 Int r^+ = {} ==> !x. (x, x) ~: r^+" +apply (rule allI) +apply (erule irrefl_tranclI) +done + +lemma trancl_rtrancl_trancl: +"\(x,y)\r^+; (y,z)\r^*\ \ (x,z)\r^+" +by (auto dest: tranclD rtrancl_trans rtrancl_into_trancl2) + +lemma rtrancl_into_trancl3: +"\(a,b)\r^*; a\b\ \ (a,b)\r^+" +apply (drule rtranclD) +apply auto +done + +lemma rtrancl_into_rtrancl2: + "\ (a, b) \ r; (b, c) \ r^* \ \ (a, c) \ r^*" +by (auto intro: r_into_rtrancl rtrancl_trans) + +lemma triangle_lemma: + "\ \ a b c. \(a,b)\r; (a,c)\r\ \ b=c; (a,x)\r\<^sup>*; (a,y)\r\<^sup>*\ + \ (x,y)\r\<^sup>* \ (y,x)\r\<^sup>*" +proof - + note converse_rtrancl_induct = converse_rtrancl_induct [consumes 1] + note converse_rtranclE = converse_rtranclE [consumes 1] + assume unique: "\ a b c. \(a,b)\r; (a,c)\r\ \ b=c" + assume "(a,x)\r\<^sup>*" + then show "(a,y)\r\<^sup>* \ (x,y)\r\<^sup>* \ (y,x)\r\<^sup>*" + proof (induct rule: converse_rtrancl_induct) + assume "(x,y)\r\<^sup>*" + then show ?thesis + by blast + next + fix a v + assume a_v_r: "(a, v) \ r" and + v_x_rt: "(v, x) \ r\<^sup>*" and + a_y_rt: "(a, y) \ r\<^sup>*" and + hyp: "(v, y) \ r\<^sup>* \ (x, y) \ r\<^sup>* \ (y, x) \ r\<^sup>*" + from a_y_rt + show "(x, y) \ r\<^sup>* \ (y, x) \ r\<^sup>*" + proof (cases rule: converse_rtranclE) + assume "a=y" + with a_v_r v_x_rt have "(y,x) \ r\<^sup>*" + by (auto intro: r_into_rtrancl rtrancl_trans) + then show ?thesis + by blast + next + fix w + assume a_w_r: "(a, w) \ r" and + w_y_rt: "(w, y) \ r\<^sup>*" + from a_v_r a_w_r unique + have "v=w" + by auto + with w_y_rt hyp + show ?thesis + by blast + qed + qed +qed + + +lemma rtrancl_cases [consumes 1, case_names Refl Trancl]: + "\(a,b)\r\<^sup>*; a = b \ P; (a,b)\r\<^sup>+ \ P\ \ P" +apply (erule rtranclE) +apply (auto dest: rtrancl_into_trancl1) +done + +(* ### To Transitive_Closure *) +theorems converse_rtrancl_induct + = converse_rtrancl_induct [consumes 1,case_names Id Step] + +theorems converse_trancl_induct + = converse_trancl_induct [consumes 1,case_names Single Step] + +(* context (theory "Set") *) +lemma Ball_weaken:"\Ball s P;\ x. P x\Q x\\Ball s Q" +by auto + +(* context (theory "Finite") *) +lemma finite_SetCompr2: "[| finite (Collect P); !y. P y --> finite (range (f y)) |] ==> + finite {f y x |x y. P y}" +apply (subgoal_tac "{f y x |x y. P y} = UNION (Collect P) (%y. range (f y))") +prefer 2 apply fast +apply (erule ssubst) +apply (erule finite_UN_I) +apply fast +done + + +(* ### TO theory "List" *) +lemma list_all2_trans: "\ a b c. P1 a b \ P2 b c \ P3 a c \ + \xs2 xs3. list_all2 P1 xs1 xs2 \ list_all2 P2 xs2 xs3 \ list_all2 P3 xs1 xs3" +apply (induct_tac "xs1") +apply simp +apply (rule allI) +apply (induct_tac "xs2") +apply simp +apply (rule allI) +apply (induct_tac "xs3") +apply auto +done + + +section "pairs" + +lemma surjective_pairing5: "p = (fst p, fst (snd p), fst (snd (snd p)), fst (snd (snd (snd p))), + snd (snd (snd (snd p))))" +apply auto +done + +lemma fst_splitE [elim!]: +"[| fst s' = x'; !!x s. [| s' = (x,s); x = x' |] ==> Q |] ==> Q" +apply (cut_tac p = "s'" in surjective_pairing) +apply auto +done + +lemma fst_in_set_lemma [rule_format (no_asm)]: "(x, y) : set l --> x : fst ` set l" +apply (induct_tac "l") +apply auto +done + + +section "quantifiers" + +(*###to be phased out *) +ML {* +fun noAll_simpset () = simpset() setmksimps + mksimps (filter (fn (x,_) => x<>"All") mksimps_pairs) +*} + +lemma All_Ex_refl_eq2 [simp]: + "(!x. (? b. x = f b & Q b) \ P x) = (!b. Q b --> P (f b))" +apply auto +done + +lemma ex_ex_miniscope1 [simp]: + "(EX w v. P w v & Q v) = (EX v. (EX w. P w v) & Q v)" +apply auto +done + +lemma ex_miniscope2 [simp]: + "(EX v. P v & Q & R v) = (Q & (EX v. P v & R v))" +apply auto +done + +lemma ex_reorder31: "(\z x y. P x y z) = (\x y z. P x y z)" +apply auto +done + +lemma All_Ex_refl_eq1 [simp]: "(!x. (? b. x = f b) --> P x) = (!b. P (f b))" +apply auto +done + + +section "sums" + +hide const In0 In1 + +syntax + fun_sum :: "('a => 'c) => ('b => 'c) => (('a+'b) => 'c)" (infixr "'(+')"80) +translations + "fun_sum" == "sum_case" + +consts the_Inl :: "'a + 'b \ 'a" + the_Inr :: "'a + 'b \ 'b" +primrec "the_Inl (Inl a) = a" +primrec "the_Inr (Inr b) = b" + +datatype ('a, 'b, 'c) sum3 = In1 'a | In2 'b | In3 'c + +consts the_In1 :: "('a, 'b, 'c) sum3 \ 'a" + the_In2 :: "('a, 'b, 'c) sum3 \ 'b" + the_In3 :: "('a, 'b, 'c) sum3 \ 'c" +primrec "the_In1 (In1 a) = a" +primrec "the_In2 (In2 b) = b" +primrec "the_In3 (In3 c) = c" + +syntax + In1l :: "'al \ ('al + 'ar, 'b, 'c) sum3" + In1r :: "'ar \ ('al + 'ar, 'b, 'c) sum3" +translations + "In1l e" == "In1 (Inl e)" + "In1r c" == "In1 (Inr c)" + +ML {* +fun sum3_instantiate thm = map (fn s => simplify(simpset()delsimps[not_None_eq]) + (read_instantiate [("t","In"^s^" ?x")] thm)) ["1l","2","3","1r"] +*} +(* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *) + +translations + "option"<= (type) "Option.option" + "list" <= (type) "List.list" + "sum3" <= (type) "Basis.sum3" + + +section "quantifiers for option type" + +syntax + Oall :: "[pttrn, 'a option, bool] => bool" ("(3! _:_:/ _)" [0,0,10] 10) + Oex :: "[pttrn, 'a option, bool] => bool" ("(3? _:_:/ _)" [0,0,10] 10) + +syntax (symbols) + Oall :: "[pttrn, 'a option, bool] => bool" ("(3\_\_:/ _)" [0,0,10] 10) + Oex :: "[pttrn, 'a option, bool] => bool" ("(3\_\_:/ _)" [0,0,10] 10) + +translations + "! x:A: P" == "! x:o2s A. P" + "? x:A: P" == "? x:o2s A. P" + + +section "unique association lists" + +constdefs + unique :: "('a \ 'b) list \ bool" + "unique \ nodups \ map fst" + +lemma uniqueD [rule_format (no_asm)]: +"unique l--> (!x y. (x,y):set l --> (!x' y'. (x',y'):set l --> x=x'--> y=y'))" +apply (unfold unique_def o_def) +apply (induct_tac "l") +apply (auto dest: fst_in_set_lemma) +done + +lemma unique_Nil [simp]: "unique []" +apply (unfold unique_def) +apply (simp (no_asm)) +done + +lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))" +apply (unfold unique_def) +apply (auto dest: fst_in_set_lemma) +done + +lemmas unique_ConsI = conjI [THEN unique_Cons [THEN iffD2], standard] + +lemma unique_single [simp]: "!!p. unique [p]" +apply auto +done + +lemma unique_ConsD: "unique (x#xs) ==> unique xs" +apply (simp add: unique_def) +done + +lemma unique_append [rule_format (no_asm)]: "unique l' ==> unique l --> + (!(x,y):set l. !(x',y'):set l'. x' ~= x) --> unique (l @ l')" +apply (induct_tac "l") +apply (auto dest: fst_in_set_lemma) +done + +lemma unique_map_inj [rule_format (no_asm)]: "unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)" +apply (induct_tac "l") +apply (auto dest: fst_in_set_lemma simp add: inj_eq) +done + +lemma map_of_SomeI [rule_format (no_asm)]: "unique l --> (k, x) : set l --> map_of l k = Some x" +apply (induct_tac "l") +apply auto +done + + +section "list patterns" + +consts + lsplit :: "[['a, 'a list] => 'b, 'a list] => 'b" +defs + lsplit_def: "lsplit == %f l. f (hd l) (tl l)" +(* list patterns -- extends pre-defined type "pttrn" used in abstractions *) +syntax + "_lpttrn" :: "[pttrn,pttrn] => pttrn" ("_#/_" [901,900] 900) +translations + "%y#x#xs. b" == "lsplit (%y x#xs. b)" + "%x#xs . b" == "lsplit (%x xs . b)" + +lemma lsplit [simp]: "lsplit c (x#xs) = c x xs" +apply (unfold lsplit_def) +apply (simp (no_asm)) +done + +lemma lsplit2 [simp]: "lsplit P (x#xs) y z = P x xs y z" +apply (unfold lsplit_def) +apply simp +done + + +section "dummy pattern for quantifiers, let, etc." + +syntax + "@dummy_pat" :: pttrn ("'_") + +parse_translation {* +let fun dummy_pat_tr [] = Free ("_",dummyT) + | dummy_pat_tr ts = raise TERM ("dummy_pat_tr", ts); +in [("@dummy_pat", dummy_pat_tr)] +end +*} + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Conform.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Conform.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,447 @@ +(* Title: isabelle/Bali/Conform.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) + +header {* Conformance notions for the type soundness proof for Java *} + +theory Conform = State: + +text {* +design issues: +\begin{itemize} +\item lconf allows for (arbitrary) inaccessible values +\item ''conforms'' does not directly imply that the dynamic types of all + objects on the heap are indeed existing classes. Yet this can be + inferred for all referenced objs. +\end{itemize} +*} + +types env_ = "prog \ (lname, ty) table" (* same as env of WellType.thy *) + + +section "extension of global store" + +constdefs + + gext :: "st \ st \ bool"    ("_\|_"    [71,71] 70) + "s\|s' \ \r. \obj\globs s r: \obj'\globs s' r: tag obj'= tag obj" + +lemma gext_objD: +"\s\|s'; globs s r = Some obj\ +\ \obj'. globs s' r = Some obj' \ tag obj' = tag obj" +apply (simp only: gext_def) +by force + +lemma rev_gext_objD: +"\globs s r = Some obj; s\|s'\ + \ \obj'. globs s' r = Some obj' \ tag obj' = tag obj" +by (auto elim: gext_objD) + +lemma init_class_obj_inited: + "init_class_obj G C s1\|s2 \ inited C (globs s2)" +apply (unfold inited_def init_obj_def) +apply (auto dest!: gext_objD) +done + +lemma gext_refl [intro!, simp]: "s\|s" +apply (unfold gext_def) +apply (fast del: fst_splitE) +done + +lemma gext_gupd [simp, elim!]: "\s. globs s r = None \ s\|gupd(r\x)s" +by (auto simp: gext_def) + +lemma gext_new [simp, elim!]: "\s. globs s r = None \ s\|init_obj G oi r s" +apply (simp only: init_obj_def) +apply (erule_tac gext_gupd) +done + +lemma gext_trans [elim]: "\X. \s\|s'; s'\|s''\ \ s\|s''" +by (force simp: gext_def) + +lemma gext_upd_gobj [intro!]: "s\|upd_gobj r n v s" +apply (simp only: gext_def) +apply auto +apply (case_tac "ra = r") +apply auto +apply (case_tac "globs s r = None") +apply auto +done + +lemma gext_cong1 [simp]: "set_locals l s1\|s2 = s1\|s2" +by (auto simp: gext_def) + +lemma gext_cong2 [simp]: "s1\|set_locals l s2 = s1\|s2" +by (auto simp: gext_def) + + +lemma gext_lupd1 [simp]: "lupd(vn\v)s1\|s2 = s1\|s2" +by (auto simp: gext_def) + +lemma gext_lupd2 [simp]: "s1\|lupd(vn\v)s2 = s1\|s2" +by (auto simp: gext_def) + + +lemma inited_gext: "\inited C (globs s); s\|s'\ \ inited C (globs s')" +apply (unfold inited_def) +apply (auto dest: gext_objD) +done + + +section "value conformance" + +constdefs + + conf :: "prog \ st \ val \ ty \ bool" ("_,_\_\\_" [71,71,71,71] 70) + "G,s\v\\T \ \T'\typeof (\a. option_map obj_ty (heap s a)) v:G\T'\T" + +lemma conf_cong [simp]: "G,set_locals l s\v\\T = G,s\v\\T" +by (auto simp: conf_def) + +lemma conf_lupd [simp]: "G,lupd(vn\va)s\v\\T = G,s\v\\T" +by (auto simp: conf_def) + +lemma conf_PrimT [simp]: "\dt. typeof dt v = Some (PrimT t) \ G,s\v\\PrimT t" +apply (simp add: conf_def) +done + +lemma conf_litval [rule_format (no_asm)]: + "typeof (\a. None) v = Some T \ G,s\v\\T" +apply (unfold conf_def) +apply (rule val.induct) +apply auto +done + +lemma conf_Null [simp]: "G,s\Null\\T = G\NT\T" +by (simp add: conf_def) + +lemma conf_Addr: + "G,s\Addr a\\T = (\obj. heap s a = Some obj \ G\obj_ty obj\T)" +by (auto simp: conf_def) + +lemma conf_AddrI:"\heap s a = Some obj; G\obj_ty obj\T\ \ G,s\Addr a\\T" +apply (rule conf_Addr [THEN iffD2]) +by fast + +lemma defval_conf [rule_format (no_asm), elim]: + "is_type G T \ G,s\default_val T\\T" +apply (unfold conf_def) +apply (induct "T") +apply (auto intro: prim_ty.induct) +done + +lemma conf_widen [rule_format (no_asm), elim]: + "G\T\T' \ G,s\x\\T \ ws_prog G \ G,s\x\\T'" +apply (unfold conf_def) +apply (rule val.induct) +apply (auto elim: ws_widen_trans) +done + +lemma conf_gext [rule_format (no_asm), elim]: + "G,s\v\\T \ s\|s' \ G,s'\v\\T" +apply (unfold gext_def conf_def) +apply (rule val.induct) +apply force+ +done + + +lemma conf_list_widen [rule_format (no_asm)]: +"ws_prog G \ + \Ts Ts'. list_all2 (conf G s) vs Ts + \ G\Ts[\] Ts' \ list_all2 (conf G s) vs Ts'" +apply (unfold widens_def) +apply (rule list_all2_trans) +apply auto +done + +lemma conf_RefTD [rule_format (no_asm)]: + "G,s\a'\\RefT T + \ a' = Null \ (\a obj T'. a' = Addr a \ heap s a = Some obj \ + obj_ty obj = T' \ G\T'\RefT T)" +apply (unfold conf_def) +apply (induct_tac "a'") +apply (auto dest: widen_PrimT) +done + + +section "value list conformance" + +constdefs + + lconf :: "prog \ st \ ('a, val) table \ ('a, ty) table \ bool" + ("_,_\_[\\]_" [71,71,71,71] 70) + "G,s\vs[\\]Ts \ \n. \T\Ts n: \v\vs n: G,s\v\\T" + +lemma lconfD: "\G,s\vs[\\]Ts; Ts n = Some T\ \ G,s\(the (vs n))\\T" +by (force simp: lconf_def) + + +lemma lconf_cong [simp]: "\s. G,set_locals x s\l[\\]L = G,s\l[\\]L" +by (auto simp: lconf_def) + +lemma lconf_lupd [simp]: "G,lupd(vn\v)s\l[\\]L = G,s\l[\\]L" +by (auto simp: lconf_def) + +(* unused *) +lemma lconf_new: "\L vn = None; G,s\l[\\]L\ \ G,s\l(vn\v)[\\]L" +by (auto simp: lconf_def) + +lemma lconf_upd: "\G,s\l[\\]L; G,s\v\\T; L vn = Some T\ \ + G,s\l(vn\v)[\\]L" +by (auto simp: lconf_def) + +lemma lconf_ext: "\G,s\l[\\]L; G,s\v\\T\ \ G,s\l(vn\v)[\\]L(vn\T)" +by (auto simp: lconf_def) + +lemma lconf_map_sum [simp]: + "G,s\l1 (+) l2[\\]L1 (+) L2 = (G,s\l1[\\]L1 \ G,s\l2[\\]L2)" +apply (unfold lconf_def) +apply safe +apply (case_tac [3] "n") +apply (force split add: sum.split)+ +done + +lemma lconf_ext_list [rule_format (no_asm)]: " + \X. \G,s\l[\\]L\ \ + \vs Ts. nodups vns \ length Ts = length vns + \ list_all2 (conf G s) vs Ts \ G,s\l(vns[\]vs)[\\]L(vns[\]Ts)" +apply (unfold lconf_def) +apply (induct_tac "vns") +apply clarsimp +apply clarsimp +apply (frule list_all2_lengthD) +apply clarsimp +done + + +lemma lconf_deallocL: "\G,s\l[\\]L(vn\T); L vn = None\ \ G,s\l[\\]L" +apply (simp only: lconf_def) +apply safe +apply (drule spec) +apply (drule ospec) +apply auto +done + + +lemma lconf_gext [elim]: "\G,s\l[\\]L; s\|s'\ \ G,s'\l[\\]L" +apply (simp only: lconf_def) +apply fast +done + +lemma lconf_empty [simp, intro!]: "G,s\vs[\\]empty" +apply (unfold lconf_def) +apply force +done + +lemma lconf_init_vals [intro!]: + " \n. \T\fs n:is_type G T \ G,s\init_vals fs[\\]fs" +apply (unfold lconf_def) +apply force +done + + +section "object conformance" + +constdefs + + oconf :: "prog \ st \ obj \ oref \ bool" ("_,_\_\\\_" [71,71,71,71] 70) + "G,s\obj\\\r \ G,s\values obj[\\]var_tys G (tag obj) r \ + (case r of + Heap a \ is_type G (obj_ty obj) + | Stat C \ True)" +(* +lemma oconf_def2: "G,s\\tag=oi,values=fs\\\\r = + (G,s\fs[\\]var_tys G oi r \ + (case r of Heap a \ is_type G (obj_ty \tag=oi,values=fs\) | Stat C \ True))" +by (simp add: oconf_def Let_def) +*) +(* +lemma oconf_def2: "G,s\obj\\\r = + (G,s\values obj[\\]var_tys G (tag obj) r \ + (case r of Heap a \ is_type G (obj_ty obj) | Stat C \ True))" +by (simp add: oconf_def Let_def) +*) +lemma oconf_is_type: "G,s\obj\\\Heap a \ is_type G (obj_ty obj)" +by (auto simp: oconf_def Let_def) + +lemma oconf_lconf: "G,s\obj\\\r \ G,s\values obj[\\]var_tys G (tag obj) r" +by (simp add: oconf_def) + +lemma oconf_cong [simp]: "G,set_locals l s\obj\\\r = G,s\obj\\\r" +by (auto simp: oconf_def Let_def) + +lemma oconf_init_obj_lemma: +"\\C c. class G C = Some c \ unique (DeclConcepts.fields G C); + \C c f fld. \class G C = Some c; + table_of (DeclConcepts.fields G C) f = Some fld \ + \ is_type G (type fld); + (case r of + Heap a \ is_type G (obj_ty obj) + | Stat C \ is_class G C) +\ \ G,s\obj \values:=init_vals (var_tys G (tag obj) r)\\\\r" +apply (auto simp add: oconf_def) +apply (drule_tac var_tys_Some_eq [THEN iffD1]) +defer +apply (subst obj_ty_cong) +apply(auto dest!: fields_table_SomeD obj_ty_CInst1 obj_ty_Arr1 + split add: sum.split_asm obj_tag.split_asm) +done + +(* +lemma oconf_init_obj_lemma: +"\\C c. class G C = Some c \ unique (fields G C); + \C c f fld. \class G C = Some c; table_of (fields G C) f = Some fld \ + \ is_type G (type fld); + (case r of + Heap a \ is_type G (obj_ty \tag=oi,values=fs\) + | Stat C \ is_class G C) +\ \ G,s\\tag=oi, values=init_vals (var_tys G oi r)\\\\r" +apply (auto simp add: oconf_def) +apply (drule_tac var_tys_Some_eq [THEN iffD1]) +defer +apply (subst obj_ty_eq) +apply(auto dest!: fields_table_SomeD split add: sum.split_asm obj_tag.split_asm) +done +*) + + +section "state conformance" + +constdefs + + conforms :: "state \ env_ \ bool"  (  "_\\_" [71,71] 70) + "xs\\E \ let (G, L) = E; s = snd xs; l = locals s in + (\r. \obj\globs s r: G,s\obj \\\r) \ + \ G,s\l [\\]L\ \ + (\a. fst xs=Some(Xcpt (Loc a)) \ G,s\Addr a\\ Class (SXcpt Throwable))" + +section "conforms" + +lemma conforms_globsD: +"\(x, s)\\(G, L); globs s r = Some obj\ \ G,s\obj\\\r" +by (auto simp: conforms_def Let_def) + +lemma conforms_localD: "(x, s)\\(G, L) \ G,s\locals s[\\]L" +by (auto simp: conforms_def Let_def) + +lemma conforms_XcptLocD: "\(x, s)\\(G, L); x = Some (Xcpt (Loc a))\ \ + G,s\Addr a\\ Class (SXcpt Throwable)" +by (auto simp: conforms_def Let_def) + +lemma conforms_RefTD: + "\G,s\a'\\RefT t; a' \ Null; (x,s) \\(G, L)\ \ + \a obj. a' = Addr a \ globs s (Inl a) = Some obj \ + G\obj_ty obj\RefT t \ is_type G (obj_ty obj)" +apply (drule_tac conf_RefTD) +apply clarsimp +apply (rule conforms_globsD [THEN oconf_is_type]) +apply auto +done + +lemma conforms_Jump [iff]: + "((Some (Jump j), s)\\(G, L)) = (Norm s\\(G, L))" +by (auto simp: conforms_def) + +lemma conforms_StdXcpt [iff]: + "((Some (Xcpt (Std xn)), s)\\(G, L)) = (Norm s\\(G, L))" +by (auto simp: conforms_def) + +lemma conforms_raise_if [iff]: + "((raise_if c xn x, s)\\(G, L)) = ((x, s)\\(G, L))" +by (auto simp: abrupt_if_def) + + +lemma conforms_NormI: "(x, s)\\(G, L) \ Norm s\\(G, L)" +by (auto simp: conforms_def Let_def) + + +lemma conforms_absorb [rule_format]: + "(a, b)\\(G, L) \ (absorb j a, b)\\(G, L)" +apply (rule impI) +apply ( case_tac a) +apply (case_tac "absorb j a") +apply auto +apply (case_tac "absorb j (Some a)",auto) +apply (erule conforms_NormI) +done + +lemma conformsI: "\\r. \obj\globs s r: G,s\obj\\\r; + G,s\locals s[\\]L; + \a. x = Some (Xcpt (Loc a)) \ G,s\Addr a\\ Class (SXcpt Throwable)\ \ + (x, s)\\(G, L)" +by (auto simp: conforms_def Let_def) + +lemma conforms_xconf: "\(x, s)\\(G,L); + \a. x' = Some (Xcpt (Loc a)) \ G,s\Addr a\\ Class (SXcpt Throwable)\ \ + (x',s)\\(G,L)" +by (fast intro: conformsI elim: conforms_globsD conforms_localD) + +lemma conforms_lupd: + "\(x, s)\\(G, L); L vn = Some T; G,s\v\\T\ \ (x, lupd(vn\v)s)\\(G, L)" +by (force intro: conformsI lconf_upd dest: conforms_globsD conforms_localD + conforms_XcptLocD simp: oconf_def) + + +lemmas conforms_allocL_aux = conforms_localD [THEN lconf_ext] + +lemma conforms_allocL: + "\(x, s)\\(G, L); G,s\v\\T\ \ (x, lupd(vn\v)s)\\(G, L(vn\T))" +by (force intro: conformsI dest: conforms_globsD + elim: conforms_XcptLocD conforms_allocL_aux simp: oconf_def) + +lemmas conforms_deallocL_aux = conforms_localD [THEN lconf_deallocL] + +lemma conforms_deallocL: "\s.\s\\(G, L(vn\T)); L vn = None\ \ s\\(G,L)" +by (fast intro: conformsI dest: conforms_globsD + elim: conforms_XcptLocD conforms_deallocL_aux) + +lemma conforms_gext: "\(x, s)\\(G,L); s\|s'; + \r. \obj\globs s' r: G,s'\obj\\\r; + locals s'=locals s\ \ (x,s')\\(G,L)" +by (force intro!: conformsI dest: conforms_localD conforms_XcptLocD) + + +lemma conforms_xgext: + "\(x ,s)\\(G,L); (x', s')\\(G, L); s'\|s\ \ (x',s)\\(G,L)" +apply (erule_tac conforms_xconf) +apply (fast dest: conforms_XcptLocD) +done + +lemma conforms_gupd: "\obj. \(x, s)\\(G, L); G,s\obj\\\r; s\|gupd(r\obj)s\ +\ (x, gupd(r\obj)s)\\(G, L)" +apply (rule conforms_gext) +apply auto +apply (force dest: conforms_globsD simp add: oconf_def)+ +done + +lemma conforms_upd_gobj: "\(x,s)\\(G, L); globs s r = Some obj; + var_tys G (tag obj) r n = Some T; G,s\v\\T\ \ (x,upd_gobj r n v s)\\(G,L)" +apply (rule conforms_gext) +apply auto +apply (drule (1) conforms_globsD) +apply (simp add: oconf_def) +apply safe +apply (rule lconf_upd) +apply auto +apply (simp only: obj_ty_cong) +apply (force dest: conforms_globsD intro!: lconf_upd + simp add: oconf_def cong del: sum.weak_case_cong) +done + +lemma conforms_set_locals: + "\(x,s)\\(G, L'); G,s\l[\\]L\ \ (x,set_locals l s)\\(G,L)" +apply (auto intro!: conformsI dest: conforms_globsD + elim!: conforms_XcptLocD simp add: oconf_def) +done + +lemma conforms_return: "\s'. \(x,s)\\(G, L); (x',s')\\(G, L'); s\|s'\ \ + (x',set_locals (locals s) s')\\(G, L)" +apply (rule conforms_xconf) +prefer 2 apply (force dest: conforms_XcptLocD) +apply (erule conforms_gext) +apply (force dest: conforms_globsD)+ +done + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Decl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Decl.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,867 @@ +(* Title: isabelle/Bali/Decl.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* Field, method, interface, and class declarations, whole Java programs +*} + +(** order is significant, because of clash for "var" **) +theory Decl = Term + Table: + +text {* +improvements: +\begin{itemize} +\item clarification and correction of some aspects of the package/access concept + (Also submitted as bug report to the Java Bug Database: + Bug Id: 4485402 and Bug Id: 4493343 + http://developer.java.sun.com/developer/bugParade/index.jshtml + ) +\end{itemize} +simplifications: +\begin{itemize} +\item the only field and method modifiers are static and the access modifiers +\item no constructors, which may be simulated by new + suitable methods +\item there is just one global initializer per class, which can simulate all + others + +\item no throws clause +\item a void method is replaced by one that returns Unit (of dummy type Void) + +\item no interface fields + +\item every class has an explicit superclass (unused for Object) +\item the (standard) methods of Object and of standard exceptions are not + specified + +\item no main method +\end{itemize} +*} + +subsection {* Modifier*} + +subsubsection {* Access modifier *} + +datatype acc_modi (* access modifier *) + = Private | Package | Protected | Public + +text {* +We can define a linear order for the access modifiers. With Private yielding the +most restrictive access and public the most liberal access policy: + Private < Package < Protected < Public +*} + +instance acc_modi:: ord +by intro_classes + +defs (overloaded) +less_acc_def: + "a < (b::acc_modi) + \ (case a of + Private \ (b=Package \ b=Protected \ b=Public) + | Package \ (b=Protected \ b=Public) + | Protected \ (b=Public) + | Public \ False)" +le_acc_def: + "a \ (b::acc_modi) \ (a = b) \ (a < b)" + +instance acc_modi:: order +proof (intro_classes) + fix x y z::acc_modi + { + show "x \ x" \\ -- reflexivity + by (auto simp add: le_acc_def) + next + assume "x \ y" "y \ z" -- transitivity + thus "x \ z" + by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split) + next + assume "x \ y" "y \ x" -- antisymmetry + thus "x = y" + proof - + have "\ x y. x < (y::acc_modi) \ y < x \ False" + by (auto simp add: less_acc_def split add: acc_modi.split) + with prems show ?thesis + by (auto simp add: le_acc_def) + qed + next + show "(x < y) = (x \ y \ x \ y)" + by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split) + } +qed + +instance acc_modi:: linorder +proof intro_classes + fix x y:: acc_modi + show "x \ y \ y \ x" + by (auto simp add: less_acc_def le_acc_def split add: acc_modi.split) +qed + +lemma acc_modi_top [simp]: "Public \ a \ a = Public" +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) + +lemma acc_modi_top1 [simp, intro!]: "a \ Public" +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) + +lemma acc_modi_le_Public: +"a \ Public \ a=Private \ a = Package \ a=Protected \ a=Public" +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) + +lemma acc_modi_bottom: "a \ Private \ a = Private" +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) + +lemma acc_modi_Private_le: +"Private \ a \ a=Private \ a = Package \ a=Protected \ a=Public" +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) + +lemma acc_modi_Package_le: + "Package \ a \ a = Package \ a=Protected \ a=Public" +by (auto simp add: le_acc_def less_acc_def split: acc_modi.split) + +lemma acc_modi_le_Package: + "a \ Package \ a=Private \ a = Package" +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) + +lemma acc_modi_Protected_le: + "Protected \ a \ a=Protected \ a=Public" +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) + +lemma acc_modi_le_Protected: + "a \ Protected \ a=Private \ a = Package \ a = Protected" +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) + + +lemmas acc_modi_le_Dests = acc_modi_top acc_modi_le_Public + acc_modi_Private_le acc_modi_bottom + acc_modi_Package_le acc_modi_le_Package + acc_modi_Protected_le acc_modi_le_Protected + +lemma acc_modi_Package_le_cases + [consumes 1,case_names Package Protected Public]: + "Package \ m \ ( m = Package \ P m) \ (m=Protected \ P m) \ + (m=Public \ P m) \ P m" +by (auto dest: acc_modi_Package_le) + + +subsubsection {* Static Modifier *} +types stat_modi = bool (* modifier: static *) + +subsection {* Declaration (base "class" for member,interface and class + declarations *} + +record decl = + access :: acc_modi + +translations + "decl" <= (type) "\access::acc_modi\" + "decl" <= (type) "\access::acc_modi,\::'a\" + +subsection {* Member (field or method)*} +record member = decl + + static :: stat_modi + +translations + "member" <= (type) "\access::acc_modi,static::bool\" + "member" <= (type) "\access::acc_modi,static::bool,\::'a\" + +subsection {* Field *} + +record field = member + + type :: ty +translations + "field" <= (type) "\access::acc_modi, static::bool, type::ty\" + "field" <= (type) "\access::acc_modi, static::bool, type::ty,\::'a\" + +types + fdecl (* field declaration, cf. 8.3 *) + = "vname \ field" + + +translations + "fdecl" <= (type) "vname \ field" + +subsection {* Method *} + +record mhead = member + (* method head (excluding signature) *) + pars ::"vname list" (* parameter names *) + resT ::ty (* result type *) + +record mbody = (* method body *) + lcls:: "(vname \ ty) list" (* local variables *) + stmt:: stmt (* the body statement *) + +record methd = mhead + (* method in a class *) + mbody::mbody + +types mdecl = "sig \ methd" (* method declaration in a class *) + + +translations + "mhead" <= (type) "\access::acc_modi, static::bool, + pars::vname list, resT::ty\" + "mhead" <= (type) "\access::acc_modi, static::bool, + pars::vname list, resT::ty,\::'a\" + "mbody" <= (type) "\lcls::(vname \ ty) list,stmt::stmt\" + "mbody" <= (type) "\lcls::(vname \ ty) list,stmt::stmt,\::'a\" + "methd" <= (type) "\access::acc_modi, static::bool, + pars::vname list, resT::ty,mbody::mbody\" + "methd" <= (type) "\access::acc_modi, static::bool, + pars::vname list, resT::ty,mbody::mbody,\::'a\" + "mdecl" <= (type) "sig \ methd" + + +constdefs + mhead::"methd \ mhead" + "mhead m \ \access=access m, static=static m, pars=pars m, resT=resT m\" + +lemma access_mhead [simp]:"access (mhead m) = access m" +by (simp add: mhead_def) + +lemma static_mhead [simp]:"static (mhead m) = static m" +by (simp add: mhead_def) + +lemma pars_mhead [simp]:"pars (mhead m) = pars m" +by (simp add: mhead_def) + +lemma resT_mhead [simp]:"resT (mhead m) = resT m" +by (simp add: mhead_def) + +text {* To be able to talk uniformaly about field and method declarations we +introduce the notion of a member declaration (e.g. useful to define +accessiblity ) *} + +datatype memberdecl = fdecl fdecl | mdecl mdecl + +datatype memberid = fid vname | mid sig + +axclass has_memberid < "type" +consts + memberid :: "'a::has_memberid \ memberid" + +instance memberdecl::has_memberid +by (intro_classes) + +defs (overloaded) +memberdecl_memberid_def: + "memberid m \ (case m of + fdecl (vn,f) \ fid vn + | mdecl (sig,m) \ mid sig)" + +lemma memberid_fdecl_simp[simp]: "memberid (fdecl (vn,f)) = fid vn" +by (simp add: memberdecl_memberid_def) + +lemma memberid_fdecl_simp1: "memberid (fdecl f) = fid (fst f)" +by (cases f) (simp add: memberdecl_memberid_def) + +lemma memberid_mdecl_simp[simp]: "memberid (mdecl (sig,m)) = mid sig" +by (simp add: memberdecl_memberid_def) + +lemma memberid_mdecl_simp1: "memberid (mdecl m) = mid (fst m)" +by (cases m) (simp add: memberdecl_memberid_def) + +instance * :: ("type","has_memberid") has_memberid +by (intro_classes) + +defs (overloaded) +pair_memberid_def: + "memberid p \ memberid (snd p)" + +lemma memberid_pair_simp[simp]: "memberid (c,m) = memberid m" +by (simp add: pair_memberid_def) + +lemma memberid_pair_simp1: "memberid p = memberid (snd p)" +by (simp add: pair_memberid_def) + +constdefs is_field :: "qtname \ memberdecl \ bool" +"is_field m \ \ declC f. m=(declC,fdecl f)" + +lemma is_fieldD: "is_field m \ \ declC f. m=(declC,fdecl f)" +by (simp add: is_field_def) + +lemma is_fieldI: "is_field (C,fdecl f)" +by (simp add: is_field_def) + +constdefs is_method :: "qtname \ memberdecl \ bool" +"is_method membr \ \ declC m. membr=(declC,mdecl m)" + +lemma is_methodD: "is_method membr \ \ declC m. membr=(declC,mdecl m)" +by (simp add: is_method_def) + +lemma is_methodI: "is_method (C,mdecl m)" +by (simp add: is_method_def) + + +subsection {* Interface *} + + +record ibody = decl + (* interface body *) + imethods :: "(sig \ mhead) list" (* method heads *) + +record iface = ibody + (* interface *) + isuperIfs:: "qtname list" (* superinterface list *) +types + idecl (* interface declaration, cf. 9.1 *) + = "qtname \ iface" + +translations + "ibody" <= (type) "\access::acc_modi,imethods::(sig \ mhead) list\" + "ibody" <= (type) "\access::acc_modi,imethods::(sig \ mhead) list,\::'a\" + "iface" <= (type) "\access::acc_modi,imethods::(sig \ mhead) list, + isuperIfs::qtname list\" + "iface" <= (type) "\access::acc_modi,imethods::(sig \ mhead) list, + isuperIfs::qtname list,\::'a\" + "idecl" <= (type) "qtname \ iface" + +constdefs + ibody :: "iface \ ibody" + "ibody i \ \access=access i,imethods=imethods i\" + +lemma access_ibody [simp]: "(access (ibody i)) = access i" +by (simp add: ibody_def) + +lemma imethods_ibody [simp]: "(imethods (ibody i)) = imethods i" +by (simp add: ibody_def) + +subsection {* Class *} +record cbody = decl + (* class body *) + cfields:: "fdecl list" + methods:: "mdecl list" + init :: "stmt" (* initializer *) + +record class = cbody + (* class *) + super :: "qtname" (* superclass *) + superIfs:: "qtname list" (* implemented interfaces *) +types + cdecl (* class declaration, cf. 8.1 *) + = "qtname \ class" + +translations + "cbody" <= (type) "\access::acc_modi,cfields::fdecl list, + methods::mdecl list,init::stmt\" + "cbody" <= (type) "\access::acc_modi,cfields::fdecl list, + methods::mdecl list,init::stmt,\::'a\" + "class" <= (type) "\access::acc_modi,cfields::fdecl list, + methods::mdecl list,init::stmt, + super::qtname,superIfs::qtname list\" + "class" <= (type) "\access::acc_modi,cfields::fdecl list, + methods::mdecl list,init::stmt, + super::qtname,superIfs::qtname list,\::'a\" + "cdecl" <= (type) "qtname \ class" + +constdefs + cbody :: "class \ cbody" + "cbody c \ \access=access c, cfields=cfields c,methods=methods c,init=init c\" + +lemma access_cbody [simp]:"access (cbody c) = access c" +by (simp add: cbody_def) + +lemma cfields_cbody [simp]:"cfields (cbody c) = cfields c" +by (simp add: cbody_def) + +lemma methods_cbody [simp]:"methods (cbody c) = methods c" +by (simp add: cbody_def) + +lemma init_cbody [simp]:"init (cbody c) = init c" +by (simp add: cbody_def) + + +section "standard classes" + +consts + + Object_mdecls ::  "mdecl list" (* methods of Object *) + SXcpt_mdecls ::  "mdecl list" (* methods of SXcpts *) + ObjectC ::  "cdecl" (* declaration of root class *) + SXcptC ::"xname \ cdecl" (* declarations of throwable classes *) + +defs + + +ObjectC_def:"ObjectC \ (Object,\access=Public,cfields=[],methods=Object_mdecls, + init=Skip,super=arbitrary,superIfs=[]\)" +SXcptC_def:"SXcptC xn\ (SXcpt xn,\access=Public,cfields=[],methods=SXcpt_mdecls, + init=Skip, + super=if xn = Throwable then Object + else SXcpt Throwable, + superIfs=[]\)" + +lemma ObjectC_neq_SXcptC [simp]: "ObjectC \ SXcptC xn" +by (simp add: ObjectC_def SXcptC_def Object_def SXcpt_def) + +lemma SXcptC_inject [simp]: "(SXcptC xn = SXcptC xm) = (xn = xm)" +apply (simp add: SXcptC_def) +apply auto +done + +constdefs standard_classes :: "cdecl list" + "standard_classes \ [ObjectC, SXcptC Throwable, + SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast, + SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]" + + +section "programs" + +record prog = + ifaces ::"idecl list" + "classes"::"cdecl list" + +translations + "prog"<= (type) "\ifaces::idecl list,classes::cdecl list\" + "prog"<= (type) "\ifaces::idecl list,classes::cdecl list,\::'a\" + +syntax + iface :: "prog \ (qtname, iface) table" + class :: "prog \ (qtname, class) table" + is_iface :: "prog \ qtname \ bool" + is_class :: "prog \ qtname \ bool" + +translations + "iface G I" == "table_of (ifaces G) I" + "class G C" == "table_of (classes G) C" + "is_iface G I" == "iface G I \ None" + "is_class G C" == "class G C \ None" + + +section "is type" + +consts + is_type :: "prog \ ty \ bool" + isrtype :: "prog \ ref_ty \ bool" + +primrec "is_type G (PrimT pt) = True" + "is_type G (RefT rt) = isrtype G rt" + "isrtype G (NullT ) = True" + "isrtype G (IfaceT tn) = is_iface G tn" + "isrtype G (ClassT tn) = is_class G tn" + "isrtype G (ArrayT T ) = is_type G T" + +lemma type_is_iface: "is_type G (Iface I) \ is_iface G I" +by auto + +lemma type_is_class: "is_type G (Class C) \ is_class G C" +by auto + + +section "subinterface and subclass relation, in anticipation of TypeRel.thy" + +consts + subint1 :: "prog \ (qtname \ qtname) set" + subcls1 :: "prog \ (qtname \ qtname) set" + +defs + subint1_def: "subint1 G \ {(I,J). \i\iface G I: J\set (isuperIfs i)}" + subcls1_def: "subcls1 G \ {(C,D). C\Object \ (\c\class G C: super c = D)}" + +syntax + "@subcls1" :: "prog => [qtname, qtname] => bool" ("_|-_<:C1_" [71,71,71] 70) + "@subclseq":: "prog => [qtname, qtname] => bool" ("_|-_<=:C _"[71,71,71] 70) + "@subcls" :: "prog => [qtname, qtname] => bool" ("_|-_<:C _"[71,71,71] 70) + +syntax (xsymbols) + "@subcls1" :: "prog \ [qtname, qtname] \ bool" ("_\_\\<^sub>C\<^sub>1_" [71,71,71] 70) + "@subclseq":: "prog \ [qtname, qtname] \ bool" ("_\_\\<^sub>C _" [71,71,71] 70) + "@subcls" :: "prog \ [qtname, qtname] \ bool" ("_\_\\<^sub>C _" [71,71,71] 70) + +translations + "G\C \\<^sub>C\<^sub>1 D" == "(C,D) \ subcls1 G" + "G\C \\<^sub>C D" == "(C,D) \(subcls1 G)^*" (* cf. 8.1.3 *) + "G\C \\<^sub>C D" == "(C,D) \(subcls1 G)^+" + + +lemma subint1I: "\iface G I = Some i; J \ set (isuperIfs i)\ + \ (I,J) \ subint1 G" +apply (simp add: subint1_def) +done + +lemma subcls1I:"\class G C = Some c; C \ Object\ \ (C,(super c)) \ subcls1 G" +apply (simp add: subcls1_def) +done + + +lemma subint1D: "(I,J)\subint1 G\ \i\iface G I: J\set (isuperIfs i)" +by (simp add: subint1_def) + +lemma subcls1D: + "(C,D)\subcls1 G \ C\Object \ (\c. class G C = Some c \ (super c = D))" +apply (simp add: subcls1_def) +apply auto +done + +lemma subint1_def2: + "subint1 G = (\ I\{I. is_iface G I}. set (isuperIfs (the (iface G I))))" +apply (unfold subint1_def) +apply auto +done + +lemma subcls1_def2: + "subcls1 G = (\C\{C. is_class G C}. {D. C\Object \ super (the(class G C))=D})" +apply (unfold subcls1_def) +apply auto +done + +lemma subcls_is_class: +"\G\C \\<^sub>C D\ \ \ c. class G C = Some c" +by (auto simp add: subcls1_def dest: tranclD) + +lemma no_subcls1_Object:"G\Object\\<^sub>C\<^sub>1 D \ P" +by (auto simp add: subcls1_def) + +lemma no_subcls_Object: "G\Object\\<^sub>C D \ P" +apply (erule trancl_induct) +apply (auto intro: no_subcls1_Object) +done + +section "well-structured programs" + +constdefs + ws_idecl :: "prog \ qtname \ qtname list \ bool" + "ws_idecl G I si \ \J\set si. is_iface G J \ (J,I)\(subint1 G)^+" + + ws_cdecl :: "prog \ qtname \ qtname \ bool" + "ws_cdecl G C sc \ C\Object \ is_class G sc \ (sc,C)\(subcls1 G)^+" + + ws_prog :: "prog \ bool" + "ws_prog G \ (\(I,i)\set (ifaces G). ws_idecl G I (isuperIfs i)) \ + (\(C,c)\set (classes G). ws_cdecl G C (super c))" + + +lemma ws_progI: +"\\(I,i)\set (ifaces G). \J\set (isuperIfs i). is_iface G J \ + (J,I) \ (subint1 G)^+; + \(C,c)\set (classes G). C\Object \ is_class G (super c) \ + ((super c),C) \ (subcls1 G)^+ + \ \ ws_prog G" +apply (unfold ws_prog_def ws_idecl_def ws_cdecl_def) +apply (erule_tac conjI) +apply blast +done + +lemma ws_prog_ideclD: +"\iface G I = Some i; J\set (isuperIfs i); ws_prog G\ \ + is_iface G J \ (J,I)\(subint1 G)^+" +apply (unfold ws_prog_def ws_idecl_def) +apply clarify +apply (drule_tac map_of_SomeD) +apply auto +done + +lemma ws_prog_cdeclD: +"\class G C = Some c; C\Object; ws_prog G\ \ + is_class G (super c) \ (super c,C)\(subcls1 G)^+" +apply (unfold ws_prog_def ws_cdecl_def) +apply clarify +apply (drule_tac map_of_SomeD) +apply auto +done + + +section "well-foundedness" + +lemma finite_is_iface: "finite {I. is_iface G I}" +apply (fold dom_def) +apply (rule_tac finite_dom_map_of) +done + +lemma finite_is_class: "finite {C. is_class G C}" +apply (fold dom_def) +apply (rule_tac finite_dom_map_of) +done + +lemma finite_subint1: "finite (subint1 G)" +apply (subst subint1_def2) +apply (rule finite_SigmaI) +apply (rule finite_is_iface) +apply (simp (no_asm)) +done + +lemma finite_subcls1: "finite (subcls1 G)" +apply (subst subcls1_def2) +apply (rule finite_SigmaI) +apply (rule finite_is_class) +apply (rule_tac B = "{super (the (class G C))}" in finite_subset) +apply auto +done + +lemma subint1_irrefl_lemma1: + "ws_prog G \ (subint1 G)^-1 \ (subint1 G)^+ = {}" +apply (force dest: subint1D ws_prog_ideclD conjunct2) +done + +lemma subcls1_irrefl_lemma1: + "ws_prog G \ (subcls1 G)^-1 \ (subcls1 G)^+ = {}" +apply (force dest: subcls1D ws_prog_cdeclD conjunct2) +done + +lemmas subint1_irrefl_lemma2 = subint1_irrefl_lemma1 [THEN irrefl_tranclI'] +lemmas subcls1_irrefl_lemma2 = subcls1_irrefl_lemma1 [THEN irrefl_tranclI'] + +lemma subint1_irrefl: "\(x, y) \ subint1 G; ws_prog G\ \ x \ y" +apply (rule irrefl_trancl_rD) +apply (rule subint1_irrefl_lemma2) +apply auto +done + +lemma subcls1_irrefl: "\(x, y) \ subcls1 G; ws_prog G\ \ x \ y" +apply (rule irrefl_trancl_rD) +apply (rule subcls1_irrefl_lemma2) +apply auto +done + +lemmas subint1_acyclic = subint1_irrefl_lemma2 [THEN acyclicI, standard] +lemmas subcls1_acyclic = subcls1_irrefl_lemma2 [THEN acyclicI, standard] + + +lemma wf_subint1: "ws_prog G \ wf ((subint1 G)\)" +by (auto intro: finite_acyclic_wf_converse finite_subint1 subint1_acyclic) + +lemma wf_subcls1: "ws_prog G \ wf ((subcls1 G)\)" +by (auto intro: finite_acyclic_wf_converse finite_subcls1 subcls1_acyclic) + + +lemma subint1_induct: + "\ws_prog G; \x. \y. (x, y) \ subint1 G \ P y \ P x\ \ P a" +apply (frule wf_subint1) +apply (erule wf_induct) +apply (simp (no_asm_use) only: converse_iff) +apply blast +done + +lemma subcls1_induct [consumes 1]: + "\ws_prog G; \x. \y. (x, y) \ subcls1 G \ P y \ P x\ \ P a" +apply (frule wf_subcls1) +apply (erule wf_induct) +apply (simp (no_asm_use) only: converse_iff) +apply blast +done + +lemma ws_subint1_induct: + "\is_iface G I; ws_prog G; \I i. \iface G I = Some i \ + (\J \ set (isuperIfs i). (I,J)\subint1 G \ P J \ is_iface G J)\ \ P I + \ \ P I" +apply (erule make_imp) +apply (rule subint1_induct) +apply assumption +apply safe +apply (fast dest: subint1I ws_prog_ideclD) +done + + +lemma ws_subcls1_induct: "\is_class G C; ws_prog G; + \C c. \class G C = Some c; + (C \ Object \ (C,(super c))\subcls1 G \ + P (super c) \ is_class G (super c))\ \ P C + \ \ P C" +apply (erule make_imp) +apply (rule subcls1_induct) +apply assumption +apply safe +apply (fast dest: subcls1I ws_prog_cdeclD) +done + +lemma ws_class_induct [consumes 2, case_names Object Subcls]: +"\class G C = Some c; ws_prog G; + \ co. class G Object = Some co \ P Object; + \ C c. \class G C = Some c; C \ Object; P (super c)\ \ P C + \ \ P C" +proof - + assume clsC: "class G C = Some c" + and init: "\ co. class G Object = Some co \ P Object" + and step: "\ C c. \class G C = Some c; C \ Object; P (super c)\ \ P C" + assume ws: "ws_prog G" + then have "is_class G C \ P C" + proof (induct rule: subcls1_induct) + fix C + assume hyp:"\ S. G\C \\<^sub>C\<^sub>1 S \ is_class G S \ P S" + and iscls:"is_class G C" + show "P C" + proof (cases "C=Object") + case True with iscls init show "P C" by auto + next + case False with ws step hyp iscls + show "P C" by (auto dest: subcls1I ws_prog_cdeclD) + qed + qed + with clsC show ?thesis by simp +qed + +lemma ws_class_induct' [consumes 2, case_names Object Subcls]: +"\is_class G C; ws_prog G; + \ co. class G Object = Some co \ P Object; + \ C c. \class G C = Some c; C \ Object; P (super c)\ \ P C + \ \ P C" +by (blast intro: ws_class_induct) + +lemma ws_class_induct'' [consumes 2, case_names Object Subcls]: +"\class G C = Some c; ws_prog G; + \ co. class G Object = Some co \ P Object co; + \ C c sc. \class G C = Some c; class G (super c) = Some sc; + C \ Object; P (super c) sc\ \ P C c + \ \ P C c" +proof - + assume clsC: "class G C = Some c" + and init: "\ co. class G Object = Some co \ P Object co" + and step: "\ C c sc . \class G C = Some c; class G (super c) = Some sc; + C \ Object; P (super c) sc\ \ P C c" + assume ws: "ws_prog G" + then have "\ c. class G C = Some c\ P C c" + proof (induct rule: subcls1_induct) + fix C c + assume hyp:"\ S. G\C \\<^sub>C\<^sub>1 S \ (\ s. class G S = Some s \ P S s)" + and iscls:"class G C = Some c" + show "P C c" + proof (cases "C=Object") + case True with iscls init show "P C c" by auto + next + case False + with ws iscls obtain sc where + sc: "class G (super c) = Some sc" + by (auto dest: ws_prog_cdeclD) + from iscls False have "G\C \\<^sub>C\<^sub>1 (super c)" by (rule subcls1I) + with False ws step hyp iscls sc + show "P C c" + by (auto) + qed + qed + with clsC show "P C c" by auto +qed + +lemma ws_interface_induct [consumes 2, case_names Step]: + (assumes is_if_I: "is_iface G I" and + ws: "ws_prog G" and + hyp_sub: "\I i. \iface G I = Some i; + \ J \ set (isuperIfs i). + (I,J)\subint1 G \ P J \ is_iface G J\ \ P I" + ) "P I" +proof - + from is_if_I ws + show "P I" + proof (rule ws_subint1_induct) + fix I i + assume hyp: "iface G I = Some i \ + (\J\set (isuperIfs i). (I,J) \subint1 G \ P J \ is_iface G J)" + then have if_I: "iface G I = Some i" + by blast + show "P I" + proof (cases "isuperIfs i") + case Nil + with if_I hyp_sub + show "P I" + by auto + next + case (Cons hd tl) + with hyp if_I hyp_sub + show "P I" + by auto + qed + qed +qed + +section "general recursion operators for the interface and class hiearchies" + +consts + iface_rec :: "prog \ qtname \ \ (qtname \ iface \ 'a set \ 'a) \ 'a" + class_rec :: "prog \ qtname \ 'a \ (qtname \ class \ 'a \ 'a) \ 'a" + +recdef iface_rec "same_fst ws_prog (\G. (subint1 G)^-1)" +"iface_rec (G,I) = + (\f. case iface G I of + None \ arbitrary + | Some i \ if ws_prog G + then f I i + ((\J. iface_rec (G,J) f)`set (isuperIfs i)) + else arbitrary)" +(hints recdef_wf: wf_subint1 intro: subint1I) +declare iface_rec.simps [simp del] + +lemma iface_rec: +"\iface G I = Some i; ws_prog G\ \ + iface_rec (G,I) f = f I i ((\J. iface_rec (G,J) f)`set (isuperIfs i))" +apply (subst iface_rec.simps) +apply simp +done + +recdef class_rec "same_fst ws_prog (\G. (subcls1 G)^-1)" +"class_rec(G,C) = + (\t f. case class G C of + None \ arbitrary + | Some c \ if ws_prog G + then f C c + (if C = Object then t + else class_rec (G,super c) t f) + else arbitrary)" +(hints recdef_wf: wf_subcls1 intro: subcls1I) +declare class_rec.simps [simp del] + +lemma class_rec: "\class G C = Some c; ws_prog G\ \ + class_rec (G,C) t f = + f C c (if C = Object then t else class_rec (G,super c) t f)" +apply (rule class_rec.simps [THEN trans [THEN fun_cong [THEN fun_cong]]]) +apply simp +done +(* +lemma bar: + "[| P; !!x. P ==> Q x |] ==> Q x" +by simp + +lemma metaMP: "[| A ==> B; A |] ==> B" +by blast + +lemma True +proof- + presume t: "C ==> E" + thm metaMP [OF t] + + presume r1: "\ B. P \ B" + presume r2: "\ C. C \ P" + thm r1 [OF r2] + + thm metaMP [OF t] + +lemma ws_subcls1_induct4: "\is_class G C; ws_prog G; + \C c. \C \ Object\ P (super c)\ \ P C + \ \ P C" +proof - + assume cls_C: "is_class G C" + and ws: "ws_prog G" + and hyp: "\C c. \C \ Object\ P (super c)\ \ P C" + thm ws_subcls1_induct [OF cls_C ws hyp] + +show +(\C c. class G C = Some c \ + (C \ Object \ G\C\\<^sub>C\<^sub>1super c \ ?P (super c) \ is_class G (super c)) \ + ?P C) \ +?P C + show ?thesis + thm "thm ws_subcls1_induct [OF cls_C ws hyp]" + apply (rule ws_subcls1_induct) + proof (rule ws_subcls1_induct) + fix C c + assume "class G C = Some c \ + (C \ Object \ + G\C\\<^sub>C\<^sub>1super c \ P (super c) \ is_class G (super c))" + show "C \ Object \ P (super (?c C c))" +apply (erule ws_subcls1_induct) +apply assumption +apply (erule conjE) +apply (case_tac "C=Object") +apply blast +apply (erule impE) +apply assumption +apply (erule conjE)+ +apply (rotate_tac 2) +sorry + +*) + + +constdefs +imethds:: "prog \ qtname \ (sig,qtname \ mhead) tables" + (* methods of an interface, with overriding and inheritance, cf. 9.2 *) +"imethds G I + \ iface_rec (G,I) + (\I i ts. (Un_tables ts) \\ + (o2s \ table_of (map (\(s,m). (s,I,m)) (imethods i))))" + + + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/DeclConcepts.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/DeclConcepts.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,2540 @@ +header {* Advanced concepts on Java declarations like overriding, inheritance, +dynamic method lookup*} + +theory DeclConcepts = TypeRel: + +section "access control (cf. 6.6), overriding and hiding (cf. 8.4.6.1)" + +constdefs +is_public :: "prog \ qtname \ bool" +"is_public G qn \ (case class G qn of + None \ (case iface G qn of + None \ False + | Some iface \ access iface = Public) + | Some class \ access class = Public)" + +subsection "accessibility of types (cf. 6.6.1)" +text {* +Primitive types are always accessible, interfaces and classes are accessible +in their package or if they are defined public, an array type is accessible if +its element type is accessible *} + +consts accessible_in :: "prog \ ty \ pname \ bool" + ("_ \ _ accessible'_in _" [61,61,61] 60) + rt_accessible_in:: "prog \ ref_ty \ pname \ bool" + ("_ \ _ accessible'_in' _" [61,61,61] 60) +primrec +"G\(PrimT p) accessible_in pack = True" +accessible_in_RefT_simp: +"G\(RefT r) accessible_in pack = G\r accessible_in' pack" + +"G\(NullT) accessible_in' pack = True" +"G\(IfaceT I) accessible_in' pack = ((pid I = pack) \ is_public G I)" +"G\(ClassT C) accessible_in' pack = ((pid C = pack) \ is_public G C)" +"G\(ArrayT ty) accessible_in' pack = G\ty accessible_in pack" + +declare accessible_in_RefT_simp [simp del] + +constdefs + is_acc_class :: "prog \ pname \ qtname \ bool" + "is_acc_class G P C \ is_class G C \ G\(Class C) accessible_in P" + is_acc_iface :: "prog \ pname \ qtname \ bool" + "is_acc_iface G P I \ is_iface G I \ G\(Iface I) accessible_in P" + is_acc_type :: "prog \ pname \ ty \ bool" + "is_acc_type G P T \ is_type G T \ G\T accessible_in P" + is_acc_reftype :: "prog \ pname \ ref_ty \ bool" + "is_acc_reftype G P T \ isrtype G T \ G\T accessible_in' P" + +lemma is_acc_classD: + "is_acc_class G P C \ is_class G C \ G\(Class C) accessible_in P" +by (simp add: is_acc_class_def) + +lemma is_acc_class_is_class: "is_acc_class G P C \ is_class G C" +by (auto simp add: is_acc_class_def) + +lemma is_acc_ifaceD: + "is_acc_iface G P I \ is_iface G I \ G\(Iface I) accessible_in P" +by (simp add: is_acc_iface_def) + +lemma is_acc_typeD: + "is_acc_type G P T \ is_type G T \ G\T accessible_in P" +by (simp add: is_acc_type_def) + +lemma is_acc_reftypeD: +"is_acc_reftype G P T \ isrtype G T \ G\T accessible_in' P" +by (simp add: is_acc_reftype_def) + +subsection "accessibility of members" +text {* +The accessibility of members is more involved as the accessibility of types. +We have to distinguish several cases to model the different effects of +accessibility during inheritance, overriding and ordinary member access +*} + +subsubsection {* Various technical conversion and selection functions *} + +text {* overloaded selector @{text accmodi} to select the access modifier +out of various HOL types *} + +axclass has_accmodi < "type" +consts accmodi:: "'a::has_accmodi \ acc_modi" + +instance acc_modi::has_accmodi +by (intro_classes) + +defs (overloaded) +acc_modi_accmodi_def: "accmodi (a::acc_modi) \ a" + +lemma acc_modi_accmodi_simp[simp]: "accmodi (a::acc_modi) = a" +by (simp add: acc_modi_accmodi_def) + +instance access_field_type:: ("type","type") has_accmodi +by (intro_classes) + +defs (overloaded) +decl_acc_modi_def: "accmodi (d::('a:: type) decl_scheme) \ access d" + + +lemma decl_acc_modi_simp[simp]: "accmodi (d::('a::type) decl_scheme) = access d" +by (simp add: decl_acc_modi_def) + +instance * :: ("type",has_accmodi) has_accmodi +by (intro_classes) + +defs (overloaded) +pair_acc_modi_def: "accmodi p \ (accmodi (snd p))" + +lemma pair_acc_modi_simp[simp]: "accmodi (x,a) = (accmodi a)" +by (simp add: pair_acc_modi_def) + +instance memberdecl :: has_accmodi +by (intro_classes) + +defs (overloaded) +memberdecl_acc_modi_def: "accmodi m \ (case m of + fdecl f \ accmodi f + | mdecl m \ accmodi m)" + +lemma memberdecl_fdecl_acc_modi_simp[simp]: + "accmodi (fdecl m) = accmodi m" +by (simp add: memberdecl_acc_modi_def) + +lemma memberdecl_mdecl_acc_modi_simp[simp]: + "accmodi (mdecl m) = accmodi m" +by (simp add: memberdecl_acc_modi_def) + +text {* overloaded selector @{text declclass} to select the declaring class +out of various HOL types *} + +axclass has_declclass < "type" +consts declclass:: "'a::has_declclass \ qtname" + +instance pid_field_type::("type","type") has_declclass +by (intro_classes) + +defs (overloaded) +qtname_declclass_def: "declclass (q::qtname) \ q" + +lemma qtname_declclass_simp[simp]: "declclass (q::qtname) = q" +by (simp add: qtname_declclass_def) + +instance * :: ("has_declclass","type") has_declclass +by (intro_classes) + +defs (overloaded) +pair_declclass_def: "declclass p \ declclass (fst p)" + +lemma pair_declclass_simp[simp]: "declclass (c,x) = declclass c" +by (simp add: pair_declclass_def) + +text {* overloaded selector @{text is_static} to select the static modifier +out of various HOL types *} + + +axclass has_static < "type" +consts is_static :: "'a::has_static \ bool" + +(* +consts is_static :: "'a \ bool" +*) + +instance access_field_type :: ("type","has_static") has_static +by (intro_classes) + +defs (overloaded) +decl_is_static_def: + "is_static (m::('a::has_static) decl_scheme) \ is_static (Decl.decl.more m)" + +instance static_field_type :: ("type","type") has_static +by (intro_classes) + +defs (overloaded) +static_field_type_is_static_def: + "is_static (m::(bool,'b::type) static_field_type) \ static_val m" + +lemma member_is_static_simp: "is_static (m::'a member_scheme) = static m" +apply (cases m) +apply (simp add: static_field_type_is_static_def + decl_is_static_def Decl.member.dest_convs) +done + +instance * :: ("type","has_static") has_static +by (intro_classes) + +defs (overloaded) +pair_is_static_def: "is_static p \ is_static (snd p)" + +lemma pair_is_static_simp [simp]: "is_static (x,s) = is_static s" +by (simp add: pair_is_static_def) + +lemma pair_is_static_simp1: "is_static p = is_static (snd p)" +by (simp add: pair_is_static_def) + +instance memberdecl:: has_static +by (intro_classes) + +defs (overloaded) +memberdecl_is_static_def: + "is_static m \ (case m of + fdecl f \ is_static f + | mdecl m \ is_static m)" + +lemma memberdecl_is_static_fdecl_simp[simp]: + "is_static (fdecl f) = is_static f" +by (simp add: memberdecl_is_static_def) + +lemma memberdecl_is_static_mdecl_simp[simp]: + "is_static (mdecl m) = is_static m" +by (simp add: memberdecl_is_static_def) + +lemma mhead_static_simp [simp]: "is_static (mhead m) = is_static m" +by (cases m) (simp add: mhead_def member_is_static_simp) + +constdefs (* some mnemotic selectors for (qtname \ ('a::more) decl_scheme) + * the first component is a class or interface name + * the second component is a method, field or method head *) +(* "declclass":: "(qtname \ ('a::more) decl_scheme) \ qtname"*) +(* "declclass \ fst" *) (* get the class component *) + + "decliface":: "(qtname \ ('a::type) decl_scheme) \ qtname" + "decliface \ fst" (* get the interface component *) + +(* + "member":: "(qtname \ ('a::type) decl_scheme) \ ('a::type) decl_scheme" +*) + "mbr":: "(qtname \ memberdecl) \ memberdecl" + "mbr \ snd" (* get the memberdecl component *) + + "mthd":: "('b \ 'a) \ 'a" + (* also used for mdecl,mhead *) + "mthd \ snd" (* get the method component *) + + "fld":: "('b \ ('a::type) decl_scheme) \ ('a::type) decl_scheme" + (* also used for ((vname \ qtname)\ field) *) + "fld \ snd" (* get the field component *) + +(* "accmodi" :: "('b \ ('a::type) decl_scheme) \ acc_modi"*) + (* also used for mdecl *) +(* "accmodi \ access \ snd"*) (* get the access modifier *) +(* + "is_static" ::"('b \ ('a::type) member_scheme) \ bool" *) + (* also defined for emhead cf. WellType *) + (*"is_static \ static \ snd"*) (* get the static modifier *) + +constdefs (* some mnemotic selectors for (vname \ qtname) *) + fname:: "(vname \ 'a) \ vname" (* also used for fdecl *) + "fname \ fst" + + declclassf:: "(vname \ qtname) \ qtname" + "declclassf \ snd" + +(* +lemma declclass_simp[simp]: "declclass (C,m) = C" +by (simp add: declclass_def) +*) + +lemma decliface_simp[simp]: "decliface (I,m) = I" +by (simp add: decliface_def) + +lemma mbr_simp[simp]: "mbr (C,m) = m" +by (simp add: mbr_def) + +lemma access_mbr_simp [simp]: "(accmodi (mbr m)) = accmodi m" +by (cases m) (simp add: mbr_def) + +lemma mthd_simp[simp]: "mthd (C,m) = m" +by (simp add: mthd_def) + +lemma fld_simp[simp]: "fld (C,f) = f" +by (simp add: fld_def) + +lemma accmodi_simp[simp]: "accmodi (C,m) = access m" +by (simp ) + +lemma access_mthd_simp [simp]: "(access (mthd m)) = accmodi m" +by (cases m) (simp add: mthd_def) + +lemma access_fld_simp [simp]: "(access (fld f)) = accmodi f" +by (cases f) (simp add: fld_def) + +(* +lemma is_static_simp[simp]: "is_static (C,m) = static m" +by (simp add: is_static_def) +*) + +lemma static_mthd_simp[simp]: "static (mthd m) = is_static m" +by (cases m) (simp add: mthd_def member_is_static_simp) + +lemma mthd_is_static_simp [simp]: "is_static (mthd m) = is_static m" +by (cases m) simp + +lemma static_fld_simp[simp]: "static (fld f) = is_static f" +by (cases f) (simp add: fld_def member_is_static_simp) + +lemma ext_field_simp [simp]: "(declclass f,fld f) = f" +by (cases f) (simp add: fld_def) + +lemma ext_method_simp [simp]: "(declclass m,mthd m) = m" +by (cases m) (simp add: mthd_def) + +lemma ext_mbr_simp [simp]: "(declclass m,mbr m) = m" +by (cases m) (simp add: mbr_def) + +lemma fname_simp[simp]:"fname (n,c) = n" +by (simp add: fname_def) + +lemma declclassf_simp[simp]:"declclassf (n,c) = c" +by (simp add: declclassf_def) + +constdefs (* some mnemotic selectors for (vname \ qtname) *) + "fldname" :: "(vname \ qtname) \ vname" + "fldname \ fst" + + "fldclass" :: "(vname \ qtname) \ qtname" + "fldclass \ snd" + +lemma fldname_simp[simp]: "fldname (n,c) = n" +by (simp add: fldname_def) + +lemma fldclass_simp[simp]: "fldclass (n,c) = c" +by (simp add: fldclass_def) + +lemma ext_fieldname_simp[simp]: "(fldname f,fldclass f) = f" +by (simp add: fldname_def fldclass_def) + +text {* Convert a qualified method declaration (qualified with its declaring +class) to a qualified member declaration: @{text methdMembr} *} + +constdefs +methdMembr :: "(qtname \ mdecl) \ (qtname \ memberdecl)" + "methdMembr m \ (fst m,mdecl (snd m))" + +lemma methdMembr_simp[simp]: "methdMembr (c,m) = (c,mdecl m)" +by (simp add: methdMembr_def) + +lemma accmodi_methdMembr_simp[simp]: "accmodi (methdMembr m) = accmodi m" +by (cases m) (simp add: methdMembr_def) + +lemma is_static_methdMembr_simp[simp]: "is_static (methdMembr m) = is_static m" +by (cases m) (simp add: methdMembr_def) + +lemma declclass_methdMembr_simp[simp]: "declclass (methdMembr m) = declclass m" +by (cases m) (simp add: methdMembr_def) + +text {* Convert a qualified method (qualified with its declaring +class) to a qualified member declaration: @{text method} *} + +constdefs +method :: "sig \ (qtname \ methd) \ (qtname \ memberdecl)" +"method sig m \ (declclass m, mdecl (sig, mthd m))" + +lemma method_simp[simp]: "method sig (C,m) = (C,mdecl (sig,m))" +by (simp add: method_def) + +lemma accmodi_method_simp[simp]: "accmodi (method sig m) = accmodi m" +by (simp add: method_def) + +lemma declclass_method_simp[simp]: "declclass (method sig m) = declclass m" +by (simp add: method_def) + +lemma is_static_method_simp[simp]: "is_static (method sig m) = is_static m" +by (cases m) (simp add: method_def) + +lemma mbr_method_simp[simp]: "mbr (method sig m) = mdecl (sig,mthd m)" +by (simp add: mbr_def method_def) + +lemma memberid_method_simp[simp]: "memberid (method sig m) = mid sig" + by (simp add: method_def) + +constdefs +fieldm :: "vname \ (qtname \ field) \ (qtname \ memberdecl)" +"fieldm n f \ (declclass f, fdecl (n, fld f))" + +lemma fieldm_simp[simp]: "fieldm n (C,f) = (C,fdecl (n,f))" +by (simp add: fieldm_def) + +lemma accmodi_fieldm_simp[simp]: "accmodi (fieldm n f) = accmodi f" +by (simp add: fieldm_def) + +lemma declclass_fieldm_simp[simp]: "declclass (fieldm n f) = declclass f" +by (simp add: fieldm_def) + +lemma is_static_fieldm_simp[simp]: "is_static (fieldm n f) = is_static f" +by (cases f) (simp add: fieldm_def) + +lemma mbr_fieldm_simp[simp]: "mbr (fieldm n f) = fdecl (n,fld f)" +by (simp add: mbr_def fieldm_def) + +lemma memberid_fieldm_simp[simp]: "memberid (fieldm n f) = fid n" +by (simp add: fieldm_def) + +text {* Select the signature out of a qualified method declaration: + @{text msig} *} + +constdefs msig:: "(qtname \ mdecl) \ sig" +"msig m \ fst (snd m)" + +lemma msig_simp[simp]: "msig (c,(s,m)) = s" +by (simp add: msig_def) + +text {* Convert a qualified method (qualified with its declaring +class) to a qualified method declaration: @{text qmdecl} *} + +constdefs qmdecl :: "sig \ (qtname \ methd) \ (qtname \ mdecl)" +"qmdecl sig m \ (declclass m, (sig,mthd m))" + +lemma qmdecl_simp[simp]: "qmdecl sig (C,m) = (C,(sig,m))" +by (simp add: qmdecl_def) + +lemma declclass_qmdecl_simp[simp]: "declclass (qmdecl sig m) = declclass m" +by (simp add: qmdecl_def) + +lemma accmodi_qmdecl_simp[simp]: "accmodi (qmdecl sig m) = accmodi m" +by (simp add: qmdecl_def) + +lemma is_static_qmdecl_simp[simp]: "is_static (qmdecl sig m) = is_static m" +by (cases m) (simp add: qmdecl_def) + +lemma msig_qmdecl_simp[simp]: "msig (qmdecl sig m) = sig" +by (simp add: qmdecl_def) + +lemma mdecl_qmdecl_simp[simp]: + "mdecl (mthd (qmdecl sig new)) = mdecl (sig, mthd new)" +by (simp add: qmdecl_def) + +lemma methdMembr_qmdecl_simp [simp]: + "methdMembr (qmdecl sig old) = method sig old" +by (simp add: methdMembr_def qmdecl_def method_def) + +text {* overloaded selector @{text resTy} to select the result type +out of various HOL types *} + +axclass has_resTy < "type" +consts resTy:: "'a::has_resTy \ ty" + +instance access_field_type :: ("type","has_resTy") has_resTy +by (intro_classes) + +defs (overloaded) +decl_resTy_def: + "resTy (m::('a::has_resTy) decl_scheme) \ resTy (Decl.decl.more m)" + +instance static_field_type :: ("type","has_resTy") has_resTy +by (intro_classes) + +defs (overloaded) +static_field_type_resTy_def: + "resTy (m::(bool,'b::has_resTy) static_field_type) + \ resTy (static_more m)" + +instance pars_field_type :: ("type","has_resTy") has_resTy +by (intro_classes) + +defs (overloaded) +pars_field_type_resTy_def: + "resTy (m::(vname list,'b::has_resTy) pars_field_type) + \ resTy (pars_more m)" + +instance resT_field_type :: ("type","type") has_resTy +by (intro_classes) + +defs (overloaded) +resT_field_type_resTy_def: + "resTy (m::(ty,'b::type) resT_field_type) + \ resT_val m" + +lemma mhead_resTy_simp: "resTy (m::'a mhead_scheme) = resT m" +apply (cases m) +apply (simp add: decl_resTy_def static_field_type_resTy_def + pars_field_type_resTy_def resT_field_type_resTy_def + member.dest_convs mhead.dest_convs) +done + +lemma resTy_mhead [simp]:"resTy (mhead m) = resTy m" +by (simp add: mhead_def mhead_resTy_simp) + +instance * :: ("type","has_resTy") has_resTy +by (intro_classes) + +defs (overloaded) +pair_resTy_def: "resTy p \ resTy (snd p)" + +lemma pair_resTy_simp[simp]: "resTy (x,m) = resTy m" +by (simp add: pair_resTy_def) + +lemma qmdecl_resTy_simp [simp]: "resTy (qmdecl sig m) = resTy m" +by (cases m) (simp) + +lemma resTy_mthd [simp]:"resTy (mthd m) = resTy m" +by (cases m) (simp add: mthd_def ) + +subsubsection "inheritable-in" +text {* +@{text "G\m inheritable_in P"}: m can be inherited by +classes in package P if: +\begin{itemize} +\item the declaration class of m is accessible in P and +\item the member m is declared with protected or public access or if it is + declared with default (package) access, the package of the declaration + class of m is also P. If the member m is declared with private access + it is not accessible for inheritance at all. +\end{itemize} +*} +constdefs +inheritable_in:: + "prog \ (qtname \ memberdecl) \ pname \ bool" + ("_ \ _ inheritable'_in _" [61,61,61] 60) +"G\membr inheritable_in pack + \ (case (accmodi membr) of + Private \ False + | Package \ (pid (declclass membr)) = pack + | Protected \ True + | Public \ True)" + +syntax +Method_inheritable_in:: + "prog \ (qtname \ mdecl) \ pname \ bool" + ("_ \Method _ inheritable'_in _ " [61,61,61] 60) + +translations +"G\Method m inheritable_in p" == "G\methdMembr m inheritable_in p" + +syntax +Methd_inheritable_in:: + "prog \ sig \ (qtname \ methd) \ pname \ bool" + ("_ \Methd _ _ inheritable'_in _ " [61,61,61,61] 60) + +translations +"G\Methd s m inheritable_in p" == "G\(method s m) inheritable_in p" + +subsubsection "declared-in/undeclared-in" + +constdefs cdeclaredmethd:: "prog \ qtname \ (sig,methd) table" +"cdeclaredmethd G C + \ (case class G C of + None \ \ sig. None + | Some c \ table_of (methods c) + )" + +constdefs +cdeclaredfield:: "prog \ qtname \ (vname,field) table" +"cdeclaredfield G C + \ (case class G C of + None \ \ sig. None + | Some c \ table_of (cfields c) + )" + + +constdefs +declared_in:: "prog \ memberdecl \ qtname \ bool" + ("_\ _ declared'_in _" [61,61,61] 60) +"G\m declared_in C \ (case m of + fdecl (fn,f ) \ cdeclaredfield G C fn = Some f + | mdecl (sig,m) \ cdeclaredmethd G C sig = Some m)" + +syntax +method_declared_in:: "prog \ (qtname \ mdecl) \ qtname \ bool" + ("_\Method _ declared'_in _" [61,61,61] 60) +translations +"G\Method m declared_in C" == "G\mdecl (mthd m) declared_in C" + +syntax +methd_declared_in:: "prog \ sig \(qtname \ methd) \ qtname \ bool" + ("_\Methd _ _ declared'_in _" [61,61,61,61] 60) +translations +"G\Methd s m declared_in C" == "G\mdecl (s,mthd m) declared_in C" + +lemma declared_in_classD: + "G\m declared_in C \ is_class G C" +by (cases m) + (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def) + +constdefs +undeclared_in:: "prog \ memberid \ qtname \ bool" + ("_\ _ undeclared'_in _" [61,61,61] 60) + +"G\m undeclared_in C \ (case m of + fid fn \ cdeclaredfield G C fn = None + | mid sig \ cdeclaredmethd G C sig = None)" + +lemma method_declared_inI: + "\class G C = Some c; table_of (methods c) sig = Some m\ + \ G\mdecl (sig,m) declared_in C" +by (auto simp add: declared_in_def cdeclaredmethd_def) + + +subsubsection "members" + +consts +members:: "prog \ (qtname \ (qtname \ memberdecl)) set" +(* Can't just take a function: prog \ qtname \ memberdecl set because + the class qtname changes to the superclass in the inductive definition + below +*) + +syntax +member_of:: "prog \ (qtname \ memberdecl) \ qtname \ bool" + ("_ \ _ member'_of _" [61,61,61] 60) + +translations + "G\m member_of C" \ "(C,m) \ members G" + +inductive "members G" intros + +Immediate: "\G\mbr m declared_in C;declclass m = C\ \ G\m member_of C" +Inherited: "\G\m inheritable_in (pid C); G\memberid m undeclared_in C; + G\C \\<^sub>C\<^sub>1 S; G\(Class S) accessible_in (pid C);G\m member_of S + \ \ G\m member_of C" +text {* Note that in the case of an inherited member only the members of the +direct superclass are concerned. If a member of a superclass of the direct +superclass isn't inherited in the direct superclass (not member of the +direct superclass) than it can't be a member of the class. E.g. If a +member of a class A is defined with package access it isn't member of a +subclass S if S isn't in the same package as A. Any further subclasses +of S will not inherit the member, regardless if they are in the same +package as A or not.*} + +syntax +method_member_of:: "prog \ (qtname \ mdecl) \ qtname \ bool" + ("_ \Method _ member'_of _" [61,61,61] 60) + +translations + "G\Method m member_of C" \ "G\(methdMembr m) member_of C" + +syntax +methd_member_of:: "prog \ sig \ (qtname \ methd) \ qtname \ bool" + ("_ \Methd _ _ member'_of _" [61,61,61,61] 60) + +translations + "G\Methd s m member_of C" \ "G\(method s m) member_of C" + +syntax +fieldm_member_of:: "prog \ vname \ (qtname \ field) \ qtname \ bool" + ("_ \Field _ _ member'_of _" [61,61,61] 60) + +translations + "G\Field n f member_of C" \ "G\fieldm n f member_of C" + +constdefs +inherits:: "prog \ qtname \ (qtname \ memberdecl) \ bool" + ("_ \ _ inherits _" [61,61,61] 60) +"G\C inherits m + \ G\m inheritable_in (pid C) \ G\memberid m undeclared_in C \ + (\ S. G\C \\<^sub>C\<^sub>1 S \ G\(Class S) accessible_in (pid C) \ G\m member_of S)" + +lemma inherits_member: "G\C inherits m \ G\m member_of C" +by (auto simp add: inherits_def intro: members.Inherited) + + +constdefs member_in::"prog \ (qtname \ memberdecl) \ qtname \ bool" + ("_ \ _ member'_in _" [61,61,61] 60) +"G\m member_in C \ \ provC. G\ C \\<^sub>C provC \ G \ m member_of provC" +text {* A member is in a class if it is member of the class or a superclass. +If a member is in a class we can select this member. This additional notion +is necessary since not all members are inherited to subclasses. So such +members are not member-of the subclass but member-in the subclass.*} + +syntax +method_member_in:: "prog \ (qtname \ mdecl) \ qtname \ bool" + ("_ \Method _ member'_in _" [61,61,61] 60) + +translations + "G\Method m member_in C" \ "G\(methdMembr m) member_in C" + +syntax +methd_member_in:: "prog \ sig \ (qtname \ methd) \ qtname \ bool" + ("_ \Methd _ _ member'_in _" [61,61,61,61] 60) + +translations + "G\Methd s m member_in C" \ "G\(method s m) member_in C" + +consts stat_overridesR:: + "prog \ ((qtname \ mdecl) \ (qtname \ mdecl)) set" + +lemma member_inD: "G\m member_in C + \ \ provC. G\ C \\<^sub>C provC \ G \ m member_of provC" +by (auto simp add: member_in_def) + +lemma member_inI: "\G \ m member_of provC;G\ C \\<^sub>C provC\ \ G\m member_in C" +by (auto simp add: member_in_def) + +lemma member_of_to_member_in: "G \ m member_of C \ G \m member_in C" +by (auto intro: member_inI) + + +subsubsection "overriding" + +text {* Unfortunately the static notion of overriding (used during the +typecheck of the compiler) and the dynamic notion of overriding (used during +execution in the JVM) are not exactly the same. +*} + +text {* Static overriding (used during the typecheck of the compiler) *} +syntax +stat_overrides:: "prog \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" + ("_ \ _ overrides\<^sub>S _" [61,61,61] 60) +translations + "G\new overrides\<^sub>S old" == "(new,old) \ stat_overridesR G " + +inductive "stat_overridesR G" intros + +Direct: "\\ is_static new; msig new = msig old; + G\(declclass new) \\<^sub>C (declclass old); + G\Method new declared_in (declclass new); + G\Method old declared_in (declclass old); + G\Method old inheritable_in pid (declclass new); + G\(declclass new) \\<^sub>C\<^sub>1 superNew; + G \Method old member_of superNew + \ \ G\new overrides\<^sub>S old" + +Indirect: "\G\new overrides\<^sub>S inter; G\inter overrides\<^sub>S old\ + \ G\new overrides\<^sub>S old" + +text {* Dynamic overriding (used during the typecheck of the compiler) *} +consts overridesR:: + "prog \ ((qtname \ mdecl) \ (qtname \ mdecl)) set" + + +overrides:: "prog \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" + ("_ \ _ overrides _" [61,61,61] 60) +translations + "G\new overrides old" == "(new,old) \ overridesR G " + +inductive "overridesR G" intros + +Direct: "\\ is_static new; \ is_static old; accmodi new \ Private; + msig new = msig old; + G\(declclass new) \\<^sub>C (declclass old); + G\Method new declared_in (declclass new); + G\Method old declared_in (declclass old); + G\Method old inheritable_in pid (declclass new); + G\resTy new \ resTy old + \ \ G\new overrides old" + +Indirect: "\G\new overrides inter; G\inter overrides old\ + \ G\new overrides old" + +syntax +sig_stat_overrides:: + "prog \ sig \ (qtname \ methd) \ (qtname \ methd) \ bool" + ("_,_\ _ overrides\<^sub>S _" [61,61,61,61] 60) +translations + "G,s\new overrides\<^sub>S old" \ "G\(qmdecl s new) overrides\<^sub>S (qmdecl s old)" + +syntax +sig_overrides:: "prog \ sig \ (qtname \ methd) \ (qtname \ methd) \ bool" + ("_,_\ _ overrides _" [61,61,61,61] 60) +translations + "G,s\new overrides old" \ "G\(qmdecl s new) overrides (qmdecl s old)" + +subsubsection "Hiding" + +constdefs hides:: +"prog \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" + ("_\ _ hides _" [61,61,61] 60) +"G\new hides old + \ is_static new \ msig new = msig old \ + G\(declclass new) \\<^sub>C (declclass old) \ + G\Method new declared_in (declclass new) \ + G\Method old declared_in (declclass old) \ + G\Method old inheritable_in pid (declclass new)" + +syntax +sig_hides:: "prog \ sig \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" + ("_,_\ _ hides _" [61,61,61,61] 60) +translations + "G,s\new hides old" \ "G\(qmdecl s new) hides (qmdecl s old)" + +lemma hidesI: +"\is_static new; msig new = msig old; + G\(declclass new) \\<^sub>C (declclass old); + G\Method new declared_in (declclass new); + G\Method old declared_in (declclass old); + G\Method old inheritable_in pid (declclass new) + \ \ G\new hides old" +by (auto simp add: hides_def) + +lemma hidesD: +"\G\new hides old\ \ + declclass new \ Object \ is_static new \ msig new = msig old \ + G\(declclass new) \\<^sub>C (declclass old) \ + G\Method new declared_in (declclass new) \ + G\Method old declared_in (declclass old)" +by (auto simp add: hides_def) + +lemma overrides_commonD: +"\G\new overrides old\ \ + declclass new \ Object \ \ is_static new \ \ is_static old \ + accmodi new \ Private \ + msig new = msig old \ + G\(declclass new) \\<^sub>C (declclass old) \ + G\Method new declared_in (declclass new) \ + G\Method old declared_in (declclass old)" +by (induct set: overridesR) (auto intro: trancl_trans) + +lemma ws_overrides_commonD: +"\G\new overrides old;ws_prog G\ \ + declclass new \ Object \ \ is_static new \ \ is_static old \ + accmodi new \ Private \ G\resTy new \ resTy old \ + msig new = msig old \ + G\(declclass new) \\<^sub>C (declclass old) \ + G\Method new declared_in (declclass new) \ + G\Method old declared_in (declclass old)" +by (induct set: overridesR) (auto intro: trancl_trans ws_widen_trans) + +lemma stat_overrides_commonD: +"\G\new overrides\<^sub>S old\ \ + declclass new \ Object \ \ is_static new \ msig new = msig old \ + G\(declclass new) \\<^sub>C (declclass old) \ + G\Method new declared_in (declclass new) \ + G\Method old declared_in (declclass old)" +by (induct set: stat_overridesR) (auto intro: trancl_trans) + +lemma overrides_eq_sigD: + "\G\new overrides old\ \ msig old=msig new" +by (auto dest: overrides_commonD) + +lemma hides_eq_sigD: + "\G\new hides old\ \ msig old=msig new" +by (auto simp add: hides_def) + +subsubsection "permits access" +constdefs +permits_acc:: + "prog \ (qtname \ memberdecl) \ qtname \ qtname \ bool" + ("_ \ _ in _ permits'_acc'_to _" [61,61,61,61] 60) + +"G\membr in class permits_acc_to accclass + \ (case (accmodi membr) of + Private \ (declclass membr = accclass) + | Package \ (pid (declclass membr) = pid accclass) + | Protected \ (pid (declclass membr) = pid accclass) + \ + (G\accclass \\<^sub>C declclass membr \ G\class \\<^sub>C accclass) + | Public \ True)" +text {* +The subcondition of the @{term "Protected"} case: +@{term "G\accclass \\<^sub>C declclass membr"} could also be relaxed to: +@{term "G\accclass \\<^sub>C declclass membr"} since in case both classes are the +same the other condition @{term "(pid (declclass membr) = pid accclass)"} +holds anyway. +*} + +text {* Like in case of overriding, the static and dynamic accessibility +of members is not uniform. +\begin{itemize} +\item Statically the class/interface of the member must be accessible for the + member to be accessible. During runtime this is not necessary. For + Example, if a class is accessible and we are allowed to access a member + of this class (statically) we expect that we can access this member in + an arbitrary subclass (during runtime). It's not intended to restrict + the access to accessible subclasses during runtime. +\item Statically the member we want to access must be "member of" the class. + Dynamically it must only be "member in" the class. +\end{itemize} +*} + + +consts +accessible_fromR:: + "prog \ qtname \ ((qtname \ memberdecl) \ qtname) set" + +syntax +accessible_from:: + "prog \ (qtname \ memberdecl) \ qtname \ qtname \ bool" + ("_ \ _ of _ accessible'_from _" [61,61,61,61] 60) + +translations +"G\membr of cls accessible_from accclass" + \ "(membr,cls) \ accessible_fromR G accclass" + +syntax +method_accessible_from:: + "prog \ (qtname \ mdecl) \ qtname \ qtname \ bool" + ("_ \Method _ of _ accessible'_from _" [61,61,61,61] 60) + +translations +"G\Method m of cls accessible_from accclass" + \ "G\methdMembr m of cls accessible_from accclass" + +syntax +methd_accessible_from:: + "prog \ sig \ (qtname \ methd) \ qtname \ qtname \ bool" + ("_ \Methd _ _ of _ accessible'_from _" [61,61,61,61,61] 60) + +translations +"G\Methd s m of cls accessible_from accclass" + \ "G\(method s m) of cls accessible_from accclass" + +syntax +field_accessible_from:: + "prog \ vname \ (qtname \ field) \ qtname \ qtname \ bool" + ("_ \Field _ _ of _ accessible'_from _" [61,61,61,61,61] 60) + +translations +"G\Field fn f of C accessible_from accclass" + \ "G\(fieldm fn f) of C accessible_from accclass" + +inductive "accessible_fromR G accclass" intros +immediate: "\G\membr member_of class; + G\(Class class) accessible_in (pid accclass); + G\membr in class permits_acc_to accclass + \ \ G\membr of class accessible_from accclass" + +overriding: "\G\membr member_of class; + G\(Class class) accessible_in (pid accclass); + membr=(C,mdecl new); + G\(C,new) overrides\<^sub>S old; + G\class \\<^sub>C sup; + G\Method old of sup accessible_from accclass + \\ G\membr of class accessible_from accclass" + +consts +dyn_accessible_fromR:: + "prog \ qtname \ ((qtname \ memberdecl) \ qtname) set" + +syntax +dyn_accessible_from:: + "prog \ (qtname \ memberdecl) \ qtname \ qtname \ bool" + ("_ \ _ in _ dyn'_accessible'_from _" [61,61,61,61] 60) + +translations +"G\membr in C dyn_accessible_from accC" + \ "(membr,C) \ dyn_accessible_fromR G accC" + +syntax +method_dyn_accessible_from:: + "prog \ (qtname \ mdecl) \ qtname \ qtname \ bool" + ("_ \Method _ in _ dyn'_accessible'_from _" [61,61,61,61] 60) + +translations +"G\Method m in C dyn_accessible_from accC" + \ "G\methdMembr m in C dyn_accessible_from accC" + +syntax +methd_dyn_accessible_from:: + "prog \ sig \ (qtname \ methd) \ qtname \ qtname \ bool" + ("_ \Methd _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60) + +translations +"G\Methd s m in C dyn_accessible_from accC" + \ "G\(method s m) in C dyn_accessible_from accC" + +syntax +field_dyn_accessible_from:: + "prog \ vname \ (qtname \ field) \ qtname \ qtname \ bool" + ("_ \Field _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60) + +translations +"G\Field fn f in dynC dyn_accessible_from accC" + \ "G\(fieldm fn f) in dynC dyn_accessible_from accC" + +(* #### Testet JVM noch über den Bytecodeverifier hinaus ob der + statische Typ accessible ist bevor es den Zugriff erlaubt + \ Test mit Reflektion\ +*) +inductive "dyn_accessible_fromR G accclass" intros +immediate: "\G\membr member_in class; + G\membr in class permits_acc_to accclass + \ \ G\membr in class dyn_accessible_from accclass" + +overriding: "\G\membr member_in class; + membr=(C,mdecl new); + G\(C,new) overrides old; + G\class \\<^sub>C sup; + G\Method old in sup dyn_accessible_from accclass + \\ G\membr in class dyn_accessible_from accclass" + + +lemma accessible_from_commonD: "G\m of C accessible_from S + \ G\m member_of C \ G\(Class C) accessible_in (pid S)" +by (auto elim: accessible_fromR.induct) + +lemma declared_not_undeclared: + "G\m declared_in C \ \ G\ memberid m undeclared_in C" +by (cases m) (auto simp add: declared_in_def undeclared_in_def) + +lemma not_undeclared_declared: + "\ G\ membr_id undeclared_in C \ (\ m. G\m declared_in C \ + membr_id = memberid m)" +proof - + assume not_undecl:"\ G\ membr_id undeclared_in C" + show ?thesis (is "?P membr_id") + proof (cases membr_id) + case (fid vname) + with not_undecl + obtain fld where + "G\fdecl (vname,fld) declared_in C" + by (auto simp add: undeclared_in_def declared_in_def + cdeclaredfield_def) + with fid show ?thesis + by auto + next + case (mid sig) + with not_undecl + obtain mthd where + "G\mdecl (sig,mthd) declared_in C" + by (auto simp add: undeclared_in_def declared_in_def + cdeclaredmethd_def) + with mid show ?thesis + by auto + qed +qed + +lemma unique_declared_in: + "\G\m declared_in C; G\n declared_in C; memberid m = memberid n\ + \ m = n" +by (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def + split: memberdecl.splits) + +lemma unique_member_of: + (assumes n: "G\n member_of C" and + m: "G\m member_of C" and + eqid: "memberid n = memberid m" + ) "n=m" +proof - + from n m eqid + show "n=m" + proof (induct) + case (Immediate C n) + assume member_n: "G\ mbr n declared_in C" "declclass n = C" + assume eqid: "memberid n = memberid m" + assume "G \ m member_of C" + then show "n=m" + proof (cases) + case (Immediate _ m') + with eqid + have "m=m'" + "memberid n = memberid m" + "G\ mbr m declared_in C" + "declclass m = C" + by auto + with member_n + show ?thesis + by (cases n, cases m) + (auto simp add: declared_in_def + cdeclaredmethd_def cdeclaredfield_def + split: memberdecl.splits) + next + case (Inherited _ _ m') + then have "G\ memberid m undeclared_in C" + by simp + with eqid member_n + show ?thesis + by (cases n) (auto dest: declared_not_undeclared) + qed + next + case (Inherited C S n) + assume undecl: "G\ memberid n undeclared_in C" + assume super: "G\C\\<^sub>C\<^sub>1S" + assume hyp: "\G \ m member_of S; memberid n = memberid m\ \ n = m" + assume eqid: "memberid n = memberid m" + assume "G \ m member_of C" + then show "n=m" + proof (cases) + case Immediate + then have "G\ mbr m declared_in C" by simp + with eqid undecl + show ?thesis + by (cases m) (auto dest: declared_not_undeclared) + next + case Inherited + with super have "G \ m member_of S" + by (auto dest!: subcls1D) + with eqid hyp + show ?thesis + by blast + qed + qed +qed + +lemma member_of_is_classD: "G\m member_of C \ is_class G C" +proof (induct set: members) + case (Immediate C m) + assume "G\ mbr m declared_in C" + then show "is_class G C" + by (cases "mbr m") + (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def) +next + case (Inherited C S m) + assume "G\C\\<^sub>C\<^sub>1S" and "is_class G S" + then show "is_class G C" + by - (rule subcls_is_class2,auto) +qed + +lemma member_of_declC: + "G\m member_of C + \ G\mbr m declared_in (declclass m)" +by (induct set: members) auto + +lemma member_of_member_of_declC: + "G\m member_of C + \ G\m member_of (declclass m)" +by (auto dest: member_of_declC intro: members.Immediate) + +lemma member_of_class_relation: + "G\m member_of C \ G\C \\<^sub>C declclass m" +proof (induct set: members) + case (Immediate C m) + then show "G\C \\<^sub>C declclass m" by simp +next + case (Inherited C S m) + then show "G\C \\<^sub>C declclass m" + by (auto dest: r_into_rtrancl intro: rtrancl_trans) +qed + +lemma member_in_class_relation: + "G\m member_in C \ G\C \\<^sub>C declclass m" +by (auto dest: member_inD member_of_class_relation + intro: rtrancl_trans) + +lemma member_of_Package: + "\G\m member_of C; accmodi m = Package\ + \ pid (declclass m) = pid C" +proof - + assume member: "G\m member_of C" + then show " accmodi m = Package \ ?thesis" (is "PROP ?P m C") + proof (induct rule: members.induct) + fix C m + assume C: "declclass m = C" + then show "pid (declclass m) = pid C" + by simp + next + fix C S m + assume inheritable: "G \ m inheritable_in pid C" + assume hyp: "PROP ?P m S" and + package_acc: "accmodi m = Package" + with inheritable package_acc hyp + show "pid (declclass m) = pid C" + by (auto simp add: inheritable_in_def) + qed +qed + +lemma dyn_accessible_from_commonD: "G\m in C dyn_accessible_from S + \ G\m member_in C" +by (auto elim: dyn_accessible_fromR.induct) + +(* ### Gilt nicht für wf_progs!dynmaisches Override, + da die accmodi Bedingung nur für stat override gilt! *) +(* +lemma override_Package: + "\G\new overrides old; + \ new old. G\new overrides old \ accmodi old \ accmodi new; + accmodi old = Package; accmodi new = Package\ + \ pid (declclass old) = pid (declclass new)" +proof - + assume wf: "\ new old. G\new overrides old \ accmodi old \ accmodi new" + assume ovverride: "G\new overrides old" + then show "\accmodi old = Package;accmodi new = Package\ \ ?thesis" + (is "?Pack old \ ?Pack new \ ?EqPid old new") + proof (induct rule: overridesR.induct) + case Direct + fix new old + assume "accmodi old = Package" + "G \ methdMembr old inheritable_in pid (declclass new)" + then show "pid (declclass old) = pid (declclass new)" + by (auto simp add: inheritable_in_def) + next + case (Indirect inter new old) + assume accmodi_old: "accmodi old = Package" and + accmodi_new: "accmodi new = Package" + assume "G \ new overrides inter" + with wf have le_inter_new: "accmodi inter \ accmodi new" + by blast + assume "G \ inter overrides old" + with wf have le_old_inter: "accmodi old \ accmodi inter" + by blast + from accmodi_old accmodi_new le_inter_new le_old_inter + have "accmodi inter = Package" + by(auto simp add: le_acc_def less_acc_def) + with Indirect accmodi_old accmodi_new + show "?EqPid old new" + by auto + qed +qed + +lemma stat_override_Package: + "\G\new overrides\<^sub>S old; + \ new old. G\new overrides\<^sub>S old \ accmodi old \ accmodi new; + accmodi old = Package; accmodi new = Package\ + \ pid (declclass old) = pid (declclass new)" +proof - + assume wf: "\ new old. G\new overrides\<^sub>S old \ accmodi old \ accmodi new" + assume ovverride: "G\new overrides\<^sub>S old" + then show "\accmodi old = Package;accmodi new = Package\ \ ?thesis" + (is "?Pack old \ ?Pack new \ ?EqPid old new") + proof (induct rule: stat_overridesR.induct) + case Direct + fix new old + assume "accmodi old = Package" + "G \ methdMembr old inheritable_in pid (declclass new)" + then show "pid (declclass old) = pid (declclass new)" + by (auto simp add: inheritable_in_def) + next + case (Indirect inter new old) + assume accmodi_old: "accmodi old = Package" and + accmodi_new: "accmodi new = Package" + assume "G \ new overrides\<^sub>S inter" + with wf have le_inter_new: "accmodi inter \ accmodi new" + by blast + assume "G \ inter overrides\<^sub>S old" + with wf have le_old_inter: "accmodi old \ accmodi inter" + by blast + from accmodi_old accmodi_new le_inter_new le_old_inter + have "accmodi inter = Package" + by(auto simp add: le_acc_def less_acc_def) + with Indirect accmodi_old accmodi_new + show "?EqPid old new" + by auto + qed +qed + +*) +lemma no_Private_stat_override: + "\G\new overrides\<^sub>S old\ \ accmodi old \ Private" +by (induct set: stat_overridesR) (auto simp add: inheritable_in_def) + +lemma no_Private_override: "\G\new overrides old\ \ accmodi old \ Private" +by (induct set: overridesR) (auto simp add: inheritable_in_def) + +lemma permits_acc_inheritance: + "\G\m in statC permits_acc_to accC; G\dynC \\<^sub>C statC + \ \ G\m in dynC permits_acc_to accC" +by (cases "accmodi m") + (auto simp add: permits_acc_def + intro: subclseq_trans) + +lemma field_accessible_fromD: + "\G\membr of C accessible_from accC;is_field membr\ + \ G\membr member_of C \ + G\(Class C) accessible_in (pid accC) \ + G\membr in C permits_acc_to accC" +by (cases set: accessible_fromR) + (auto simp add: is_field_def split: memberdecl.splits) + +lemma field_accessible_from_permits_acc_inheritance: +"\G\membr of statC accessible_from accC; is_field membr; G \ dynC \\<^sub>C statC\ +\ G\membr in dynC permits_acc_to accC" +by (auto dest: field_accessible_fromD intro: permits_acc_inheritance) + + +(* +lemma accessible_Package: + "\G \ m of C accessible_from S;accmodi m = Package; + \ new old. G\new overrides\<^sub>S old \ accmodi old \ accmodi new\ + \ pid S = pid C \ pid C = pid (declclass m)" +proof - + assume wf: "\ new old. G\new overrides\<^sub>S old \ accmodi old \ accmodi new" + assume "G \ m of C accessible_from S" + then show "accmodi m = Package \ pid S = pid C \ pid C = pid (declclass m)" + (is "?Pack m \ ?P C m") + proof (induct rule: accessible_fromR.induct) + fix C m + assume "G\m member_of C" + "G \ m in C permits_acc_to S" + "accmodi m = Package" + then show "?P C m" + by (auto dest: member_of_Package simp add: permits_acc_def) + next + fix declC C new newm old Sup + assume member_new: "G \ new member_of C" and + acc_C: "G \ Class C accessible_in pid S" and + new: "new = (declC, mdecl newm)" and + override: "G \ (declC, newm) overrides\<^sub>S old" and + subcls_C_Sup: "G\C \\<^sub>C Sup" and + acc_old: "G \ methdMembr old of Sup accessible_from S" and + hyp: "?Pack (methdMembr old) \ ?P Sup (methdMembr old)" and + accmodi_new: "accmodi new = Package" + from override wf + have accmodi_weaken: "accmodi old \ accmodi newm" + by (cases old,cases newm) auto + from override new + have "accmodi old \ Private" + by (simp add: no_Private_stat_override) + with accmodi_weaken accmodi_new new + have accmodi_old: "accmodi old = Package" + by (cases "accmodi old") (auto simp add: le_acc_def less_acc_def) + with hyp + have P_sup: "?P Sup (methdMembr old)" + by (simp) + from wf override new accmodi_old accmodi_new + have eq_pid_new_old: "pid (declclass new) = pid (declclass old)" + by (auto dest: stat_override_Package) + from member_new accmodi_new + have "pid (declclass new) = pid C" + by (auto dest: member_of_Package) + with eq_pid_new_old P_sup show "?P C new" + by auto + qed +qed +*) +lemma accessible_fieldD: + "\G\membr of C accessible_from accC; is_field membr\ + \ G\membr member_of C \ + G\(Class C) accessible_in (pid accC) \ + G\membr in C permits_acc_to accC" +by (induct rule: accessible_fromR.induct) (auto dest: is_fieldD) + +(* lemmata: + Wegen G\Super accessible_in (pid C) folgt: + G\m declared_in C; G\m member_of D; accmodi m = Package (G\D \\<^sub>C C) + \ pid C = pid D + + C package + m public in C + für alle anderen D: G\m undeclared_in C + m wird in alle subklassen vererbt, auch aus dem Package heraus! + + G\m member_of C \ \ D. G\C \\<^sub>C D \ G\m declared_in D +*) + +(* Begriff (C,m) overrides (D,m) + 3 Fälle: Direkt, + Indirekt über eine Zwischenklasse (ohne weiteres override) + Indirekt über override +*) + +(* +"G\m member_of C \ +constdefs declares_method:: "prog \ sig \ qtname \ methd \ bool" + ("_,_\ _ declares'_method _" [61,61,61,61] 60) +"G,sig\C declares_method m \ cdeclaredmethd G C sig = Some m" + +constdefs is_declared:: "prog \ sig \ (qtname \ methd) \ bool" +"is_declared G sig em \ G,sig\declclass em declares_method mthd em" +*) + +lemma member_of_Private: +"\G\m member_of C; accmodi m = Private\ \ declclass m = C" +by (induct set: members) (auto simp add: inheritable_in_def) + +lemma member_of_subclseq_declC: + "G\m member_of C \ G\C \\<^sub>C declclass m" +by (induct set: members) (auto dest: r_into_rtrancl intro: rtrancl_trans) + +lemma member_of_inheritance: + (assumes m: "G\m member_of D" and + subclseq_D_C: "G\D \\<^sub>C C" and + subclseq_C_m: "G\C \\<^sub>C declclass m" and + ws: "ws_prog G" + ) "G\m member_of C" +proof - + from m subclseq_D_C subclseq_C_m + show ?thesis + proof (induct) + case (Immediate D m) + assume "declclass m = D" and + "G\D\\<^sub>C C" and "G\C\\<^sub>C declclass m" + with ws have "D=C" + by (auto intro: subclseq_acyclic) + with Immediate + show "G\m member_of C" + by (auto intro: members.Immediate) + next + case (Inherited D S m) + assume member_of_D_props: + "G \ m inheritable_in pid D" + "G\ memberid m undeclared_in D" + "G \ Class S accessible_in pid D" + "G \ m member_of S" + assume super: "G\D\\<^sub>C\<^sub>1S" + assume hyp: "\G\S\\<^sub>C C; G\C\\<^sub>C declclass m\ \ G \ m member_of C" + assume subclseq_C_m: "G\C\\<^sub>C declclass m" + assume "G\D\\<^sub>C C" + then show "G\m member_of C" + proof (cases rule: subclseq_cases) + case Eq + assume "D=C" + with super member_of_D_props + show ?thesis + by (auto intro: members.Inherited) + next + case Subcls + assume "G\D\\<^sub>C C" + with super + have "G\S\\<^sub>C C" + by (auto dest: subcls1D subcls_superD) + with subclseq_C_m hyp show ?thesis + by blast + qed + qed +qed + +lemma member_of_subcls: + (assumes old: "G\old member_of C" and + new: "G\new member_of D" and + eqid: "memberid new = memberid old" and + subclseq_D_C: "G\D \\<^sub>C C" and + subcls_new_old: "G\declclass new \\<^sub>C declclass old" and + ws: "ws_prog G" + ) "G\D \\<^sub>C C" +proof - + from old + have subclseq_C_old: "G\C \\<^sub>C declclass old" + by (auto dest: member_of_subclseq_declC) + from new + have subclseq_D_new: "G\D \\<^sub>C declclass new" + by (auto dest: member_of_subclseq_declC) + from subcls_new_old ws + have neq_new_old: "new\old" + by (cases new,cases old) (auto dest: subcls_irrefl) + from subclseq_D_new subclseq_D_C + have "G\(declclass new) \\<^sub>C C \ G\C \\<^sub>C (declclass new)" + by (rule subcls_compareable) + then have "G\(declclass new) \\<^sub>C C" + proof + assume "G\declclass new\\<^sub>C C" then show ?thesis . + next + assume "G\C \\<^sub>C (declclass new)" + with new subclseq_D_C ws + have "G\new member_of C" + by (blast intro: member_of_inheritance) + with eqid old + have "new=old" + by (blast intro: unique_member_of) + with neq_new_old + show ?thesis + by contradiction + qed + then show ?thesis + proof (cases rule: subclseq_cases) + case Eq + assume "declclass new = C" + with new have "G\new member_of C" + by (auto dest: member_of_member_of_declC) + with eqid old + have "new=old" + by (blast intro: unique_member_of) + with neq_new_old + show ?thesis + by contradiction + next + case Subcls + assume "G\declclass new\\<^sub>C C" + with subclseq_D_new + show "G\D\\<^sub>C C" + by (rule rtrancl_trancl_trancl) + qed +qed + +corollary member_of_overrides_subcls: + "\G\Methd sig old member_of C; G\Methd sig new member_of D;G\D \\<^sub>C C; + G,sig\new overrides old; ws_prog G\ + \ G\D \\<^sub>C C" +by (drule overrides_commonD) (auto intro: member_of_subcls) + +corollary member_of_stat_overrides_subcls: + "\G\Methd sig old member_of C; G\Methd sig new member_of D;G\D \\<^sub>C C; + G,sig\new overrides\<^sub>S old; ws_prog G\ + \ G\D \\<^sub>C C" +by (drule stat_overrides_commonD) (auto intro: member_of_subcls) + + + +lemma inherited_field_access: + (assumes stat_acc: "G\membr of statC accessible_from accC" and + is_field: "is_field membr" and + subclseq: "G \ dynC \\<^sub>C statC" + ) "G\membr in dynC dyn_accessible_from accC" +proof - + from stat_acc is_field subclseq + show ?thesis + by (auto dest: accessible_fieldD + intro: dyn_accessible_fromR.immediate + member_inI + permits_acc_inheritance) +qed + +lemma accessible_inheritance: + (assumes stat_acc: "G\m of statC accessible_from accC" and + subclseq: "G\dynC \\<^sub>C statC" and + member_dynC: "G\m member_of dynC" and + dynC_acc: "G\(Class dynC) accessible_in (pid accC)" + ) "G\m of dynC accessible_from accC" +proof - + from stat_acc + have member_statC: "G\m member_of statC" + by (auto dest: accessible_from_commonD) + from stat_acc + show ?thesis + proof (cases) + case immediate + with member_dynC member_statC subclseq dynC_acc + show ?thesis + by (auto intro: accessible_fromR.immediate permits_acc_inheritance) + next + case overriding + with member_dynC subclseq dynC_acc + show ?thesis + by (auto intro: accessible_fromR.overriding rtrancl_trancl_trancl) + qed +qed + +section "fields and methods" + + +types + fspec = "vname \ qtname" + +translations + "fspec" <= (type) "vname \ qtname" + +constdefs +imethds:: "prog \ qtname \ (sig,qtname \ mhead) tables" +"imethds G I + \ iface_rec (G,I) + (\I i ts. (Un_tables ts) \\ + (o2s \ table_of (map (\(s,m). (s,I,m)) (imethods i))))" +text {* methods of an interface, with overriding and inheritance, cf. 9.2 *} + +constdefs +accimethds:: "prog \ pname \ qtname \ (sig,qtname \ mhead) tables" +"accimethds G pack I + \ if G\Iface I accessible_in pack + then imethds G I + else \ k. {}" +text {* only returns imethds if the interface is accessible *} + +constdefs +methd:: "prog \ qtname \ (sig,qtname \ methd) table" + +"methd G C + \ class_rec (G,C) empty + (\C c subcls_mthds. + filter_tab (\sig m. G\C inherits method sig m) + subcls_mthds + ++ + table_of (map (\(s,m). (s,C,m)) (methods c)))" +text {* @{term "methd G C"}: methods of a class C (statically visible from C), + with inheritance and hiding cf. 8.4.6; + Overriding is captured by @{text dynmethd}. + Every new method with the same signature coalesces the + method of a superclass. *} + +constdefs +accmethd:: "prog \ qtname \ qtname \ (sig,qtname \ methd) table" +"accmethd G S C + \ filter_tab (\sig m. G\method sig m of C accessible_from S) + (methd G C)" +text {* @{term "accmethd G S C"}: only those methods of @{term "methd G C"}, + accessible from S *} + +text {* Note the class component in the accessibility filter. The class where + method @{term m} is declared (@{term declC}) isn't necessarily accessible + from the current scope @{term S}. The method can be made accessible + through inheritance, too. + So we must test accessibility of method @{term m} of class @{term C} + (not @{term "declclass m"}) *} + +constdefs +dynmethd:: "prog \ qtname \ qtname \ (sig,qtname \ methd) table" +"dynmethd G statC dynC + \ \ sig. + (if G\dynC \\<^sub>C statC + then (case methd G statC sig of + None \ None + | Some statM + \ (class_rec (G,dynC) empty + (\C c subcls_mthds. + subcls_mthds + ++ + (filter_tab + (\ _ dynM. G,sig\dynM overrides statM \ dynM=statM) + (methd G C) )) + ) sig + ) + else None)" +(* +"dynmethd G statC dynC + \ \ sig. + (if G\dynC \\<^sub>C statC + then (case methd G statC sig of + None \ None + | Some statM + \ (class_rec (G,statC) empty + (\C c subcls_mthds. + subcls_mthds + ++ + (filter_tab + (\ _ dynM. G,sig\dynM overrides statM) + (table_of (map (\(s,m). (s,C,m)) (methods c))))) + ) sig + ) + else None)"*) +text {* @{term "dynmethd G statC dynC"}: dynamic method lookup of a reference + with dynamic class @{term dynC} and static class @{term statC} *} +text {* Note some kind of duality between @{term methd} and @{term dynmethd} + in the @{term class_rec} arguments. Whereas @{term methd} filters the + subclass methods (to get only the inherited ones), @{term dynmethd} + filters the new methods (to get only those methods which actually + override the methods of the static class) *} + +constdefs +dynimethd:: "prog \ qtname \ qtname \ (sig,qtname \ methd) table" +"dynimethd G I dynC + \ \ sig. if imethds G I sig \ {} + then methd G dynC sig + else dynmethd G Object dynC sig" +text {* @{term "dynimethd G I dynC"}: dynamic method lookup of a reference with + dynamic class dynC and static interface type I *} +text {* + When calling an interface method, we must distinguish if the method signature + was defined in the interface or if it must be an Object method in the other + case. If it was an interface method we search the class hierarchy + starting at the dynamic class of the object up to Object to find the + first matching method (@{term methd}). Since all interface methods have + public access the method can't be coalesced due to some odd visibility + effects like in case of dynmethd. The method will be inherited or + overridden in all classes from the first class implementing the interface + down to the actual dynamic class. + *} + +constdefs +dynlookup::"prog \ ref_ty \ qtname \ (sig,qtname \ methd) table" +"dynlookup G statT dynC + \ (case statT of + NullT \ empty + | IfaceT I \ dynimethd G I dynC + | ClassT statC \ dynmethd G statC dynC + | ArrayT ty \ dynmethd G Object dynC)" +text {* @{term "dynlookup G statT dynC"}: dynamic lookup of a method within the + static reference type statT and the dynamic class dynC. + In a wellformd context statT will not be NullT and in case + statT is an array type, dynC=Object *} + +constdefs +fields:: "prog \ qtname \ ((vname \ qtname) \ field) list" +"fields G C + \ class_rec (G,C) [] (\C c ts. map (\(n,t). ((n,C),t)) (cfields c) @ ts)" +text {* @{term "fields G C"} + list of fields of a class, including all the fields of the superclasses + (private, inherited and hidden ones) not only the accessible ones + (an instance of a object allocates all these fields *} + +constdefs +accfield:: "prog \ qtname \ qtname \ (vname, qtname \ field) table" +"accfield G S C + \ let field_tab = table_of((map (\((n,d),f).(n,(d,f)))) (fields G C)) + in filter_tab (\n (declC,f). G\ (declC,fdecl (n,f)) of C accessible_from S) + field_tab" +text {* @{term "accfield G C S"}: fields of a class @{term C} which are + accessible from scope of class + @{term S} with inheritance and hiding, cf. 8.3 *} +text {* note the class component in the accessibility filter (see also + @{term methd}). + The class declaring field @{term f} (@{term declC}) isn't necessarily + accessible from scope @{term S}. The field can be made visible through + inheritance, too. So we must test accessibility of field @{term f} of class + @{term C} (not @{term "declclass f"}) *} + +constdefs + + is_methd :: "prog \ qtname \ sig \ bool" + "is_methd G \ \C sig. is_class G C \ methd G C sig \ None" + +constdefs efname:: "((vname \ qtname) \ field) \ (vname \ qtname)" +"efname \ fst" + +lemma efname_simp[simp]:"efname (n,f) = n" +by (simp add: efname_def) + + +subsection "imethds" + +lemma imethds_rec: "\iface G I = Some i; ws_prog G\ \ + imethds G I = Un_tables ((\J. imethds G J)`set (isuperIfs i)) \\ + (o2s \ table_of (map (\(s,mh). (s,I,mh)) (imethods i)))" +apply (unfold imethds_def) +apply (rule iface_rec [THEN trans]) +apply auto +done + + +(* local lemma *) +lemma imethds_norec: + "\iface G md = Some i; ws_prog G; table_of (imethods i) sig = Some mh\ \ + (md, mh) \ imethds G md sig" +apply (subst imethds_rec) +apply assumption+ +apply (rule iffD2) +apply (rule overrides_t_Some_iff) +apply (rule disjI1) +apply (auto elim: table_of_map_SomeI) +done + +lemma imethds_declI: "\m \ imethds G I sig; ws_prog G; is_iface G I\ \ + (\i. iface G (decliface m) = Some i \ + table_of (imethods i) sig = Some (mthd m)) \ + (I,decliface m) \ (subint1 G)^* \ m \ imethds G (decliface m) sig" +apply (erule make_imp) +apply (rule ws_subint1_induct, assumption, assumption) +apply (subst imethds_rec, erule conjunct1, assumption) +apply (force elim: imethds_norec intro: rtrancl_into_rtrancl2) +done + +lemma imethds_cases [consumes 3, case_names NewMethod InheritedMethod]: + (assumes im: "im \ imethds G I sig" and + ifI: "iface G I = Some i" and + ws: "ws_prog G" and + hyp_new: "table_of (map (\(s, mh). (s, I, mh)) (imethods i)) sig + = Some im \ P" and + hyp_inh: "\ J. \J \ set (isuperIfs i); im \ imethds G J sig\ \ P" + ) "P" +proof - + from ifI ws im hyp_new hyp_inh + show "P" + by (auto simp add: imethds_rec) +qed + +subsection "accimethd" + +lemma accimethds_simp [simp]: +"G\Iface I accessible_in pack \ accimethds G pack I = imethds G I" +by (simp add: accimethds_def) + +lemma accimethdsD: + "im \ accimethds G pack I sig + \ im \ imethds G I sig \ G\Iface I accessible_in pack" +by (auto simp add: accimethds_def) + +lemma accimethdsI: +"\im \ imethds G I sig;G\Iface I accessible_in pack\ + \ im \ accimethds G pack I sig" +by simp + +subsection "methd" + +lemma methd_rec: "\class G C = Some c; ws_prog G\ \ + methd G C + = (if C = Object + then empty + else filter_tab (\sig m. G\C inherits method sig m) + (methd G (super c))) + ++ table_of (map (\(s,m). (s,C,m)) (methods c))" +apply (unfold methd_def) +apply (erule class_rec [THEN trans], assumption) +apply (simp) +done + +(* local lemma *) +lemma methd_norec: + "\class G declC = Some c; ws_prog G;table_of (methods c) sig = Some m\ + \ methd G declC sig = Some (declC, m)" +apply (simp only: methd_rec) +apply (rule disjI1 [THEN override_Some_iff [THEN iffD2]]) +apply (auto elim: table_of_map_SomeI) +done + + +lemma methd_declC: +"\methd G C sig = Some m; ws_prog G;is_class G C\ \ + (\d. class G (declclass m)=Some d \ table_of (methods d) sig=Some (mthd m)) \ + G\C \\<^sub>C (declclass m) \ methd G (declclass m) sig = Some m" +apply (erule make_imp) +apply (rule ws_subcls1_induct, assumption, assumption) +apply (subst methd_rec, assumption) +apply (case_tac "Ca=Object") +apply (force elim: methd_norec ) + +apply simp +apply (case_tac "table_of (map (\(s, m). (s, Ca, m)) (methods c)) sig") +apply (force intro: rtrancl_into_rtrancl2) + +apply (auto intro: methd_norec) +done + +lemma methd_inheritedD: + "\class G C = Some c; ws_prog G;methd G C sig = Some m\ + \ (declclass m \ C \ G \C inherits method sig m)" +by (auto simp add: methd_rec) + +lemma methd_diff_cls: +"\ws_prog G; is_class G C; is_class G D; + methd G C sig = m; methd G D sig = n; m\n +\ \ C\D" +by (auto simp add: methd_rec) + +lemma method_declared_inI: + "\table_of (methods c) sig = Some m; class G C = Some c\ + \ G\mdecl (sig,m) declared_in C" +by (auto simp add: cdeclaredmethd_def declared_in_def) + +lemma methd_declared_in_declclass: + "\methd G C sig = Some m; ws_prog G;is_class G C\ + \ G\Methd sig m declared_in (declclass m)" +by (auto dest: methd_declC method_declared_inI) + +lemma member_methd: + (assumes member_of: "G\Methd sig m member_of C" and + ws: "ws_prog G" + ) "methd G C sig = Some m" +proof - + from member_of + have iscls_C: "is_class G C" + by (rule member_of_is_classD) + from iscls_C ws member_of + show ?thesis (is "?Methd C") + proof (induct rule: ws_class_induct') + case (Object co) + assume "G \Methd sig m member_of Object" + then have "G\Methd sig m declared_in Object \ declclass m = Object" + by (cases set: members) (cases m, auto dest: subcls1D) + with ws Object + show "?Methd Object" + by (cases m) + (auto simp add: declared_in_def cdeclaredmethd_def methd_rec + intro: table_of_mapconst_SomeI) + next + case (Subcls C c) + assume clsC: "class G C = Some c" and + neq_C_Obj: "C \ Object" and + hyp: "G \Methd sig m member_of super c \ ?Methd (super c)" and + member_of: "G \Methd sig m member_of C" + from member_of + show "?Methd C" + proof (cases) + case (Immediate Ca membr) + then have "Ca=C" "membr = method sig m" and + "G\Methd sig m declared_in C" "declclass m = C" + by (cases m,auto) + with clsC + have "table_of (map (\(s, m). (s, C, m)) (methods c)) sig = Some m" + by (cases m) + (auto simp add: declared_in_def cdeclaredmethd_def + intro: table_of_mapconst_SomeI) + with clsC neq_C_Obj ws + show ?thesis + by (simp add: methd_rec) + next + case (Inherited Ca S membr) + with clsC + have eq_Ca_C: "Ca=C" and + undecl: "G\mid sig undeclared_in C" and + super: "G \Methd sig m member_of (super c)" + by (auto dest: subcls1D) + from eq_Ca_C clsC undecl + have "table_of (map (\(s, m). (s, C, m)) (methods c)) sig = None" + by (auto simp add: undeclared_in_def cdeclaredmethd_def + intro: table_of_mapconst_NoneI) + moreover + from Inherited have "G \ C inherits (method sig m)" + by (auto simp add: inherits_def) + moreover + note clsC neq_C_Obj ws super hyp + ultimately + show ?thesis + by (auto simp add: methd_rec intro: filter_tab_SomeI) + qed + qed +qed + +(*unused*) +lemma finite_methd:"ws_prog G \ finite {methd G C sig |sig C. is_class G C}" +apply (rule finite_is_class [THEN finite_SetCompr2]) +apply (intro strip) +apply (erule_tac ws_subcls1_induct, assumption) +apply (subst methd_rec) +apply (assumption) +apply (auto intro!: finite_range_map_of finite_range_filter_tab finite_range_map_of_override) +done + +lemma finite_dom_methd: + "\ws_prog G; is_class G C\ \ finite (dom (methd G C))" +apply (erule_tac ws_subcls1_induct) +apply assumption +apply (subst methd_rec) +apply (assumption) +apply (auto intro!: finite_dom_map_of finite_dom_filter_tab) +done + + +subsection "accmethd" + +lemma accmethd_SomeD: +"accmethd G S C sig = Some m + \ methd G C sig = Some m \ G\method sig m of C accessible_from S" +by (auto simp add: accmethd_def dest: filter_tab_SomeD) + +lemma accmethd_SomeI: +"\methd G C sig = Some m; G\method sig m of C accessible_from S\ + \ accmethd G S C sig = Some m" +by (auto simp add: accmethd_def intro: filter_tab_SomeI) + +lemma accmethd_declC: +"\accmethd G S C sig = Some m; ws_prog G; is_class G C\ \ + (\d. class G (declclass m)=Some d \ + table_of (methods d) sig=Some (mthd m)) \ + G\C \\<^sub>C (declclass m) \ methd G (declclass m) sig = Some m \ + G\method sig m of C accessible_from S" +by (auto dest: accmethd_SomeD methd_declC accmethd_SomeI) + + +lemma finite_dom_accmethd: + "\ws_prog G; is_class G C\ \ finite (dom (accmethd G S C))" +by (auto simp add: accmethd_def intro: finite_dom_filter_tab finite_dom_methd) + + +subsection "dynmethd" + +lemma dynmethd_rec: +"\class G dynC = Some c; ws_prog G\ \ + dynmethd G statC dynC sig + = (if G\dynC \\<^sub>C statC + then (case methd G statC sig of + None \ None + | Some statM + \ (case methd G dynC sig of + None \ dynmethd G statC (super c) sig + | Some dynM \ + (if G,sig\ dynM overrides statM \ dynM = statM + then Some dynM + else (dynmethd G statC (super c) sig) + ))) + else None)" + (is "_ \ _ \ ?Dynmethd_def dynC sig = ?Dynmethd_rec dynC c sig") +proof - + assume clsDynC: "class G dynC = Some c" and + ws: "ws_prog G" + then show "?Dynmethd_def dynC sig = ?Dynmethd_rec dynC c sig" + proof (induct rule: ws_class_induct'') + case (Object co) + show "?Dynmethd_def Object sig = ?Dynmethd_rec Object co sig" + proof (cases "G\Object \\<^sub>C statC") + case False + then show ?thesis by (simp add: dynmethd_def) + next + case True + then have eq_statC_Obj: "statC = Object" .. + show ?thesis + proof (cases "methd G statC sig") + case None then show ?thesis by (simp add: dynmethd_def) + next + case Some + with True Object ws eq_statC_Obj + show ?thesis + by (auto simp add: dynmethd_def class_rec + intro: filter_tab_SomeI) + qed + qed + next + case (Subcls dynC c sc) + show "?Dynmethd_def dynC sig = ?Dynmethd_rec dynC c sig" + proof (cases "G\dynC \\<^sub>C statC") + case False + then show ?thesis by (simp add: dynmethd_def) + next + case True + note subclseq_dynC_statC = True + show ?thesis + proof (cases "methd G statC sig") + case None then show ?thesis by (simp add: dynmethd_def) + next + case (Some statM) + note statM = Some + let "?filter C" = + "filter_tab + (\_ dynM. G,sig \ dynM overrides statM \ dynM = statM) + (methd G C)" + let "?class_rec C" = + "(class_rec (G, C) empty + (\C c subcls_mthds. subcls_mthds ++ (?filter C)))" + from statM Subcls ws subclseq_dynC_statC + have dynmethd_dynC_def: + "?Dynmethd_def dynC sig = + ((?class_rec (super c)) + ++ + (?filter dynC)) sig" + by (simp (no_asm_simp) only: dynmethd_def class_rec) + auto + show ?thesis + proof (cases "dynC = statC") + case True + with subclseq_dynC_statC statM dynmethd_dynC_def + have "?Dynmethd_def dynC sig = Some statM" + by (auto intro: override_find_right filter_tab_SomeI) + with subclseq_dynC_statC True Some + show ?thesis + by auto + next + case False + with subclseq_dynC_statC Subcls + have subclseq_super_statC: "G\(super c) \\<^sub>C statC" + by (blast dest: subclseq_superD) + show ?thesis + proof (cases "methd G dynC sig") + case None + then have "?filter dynC sig = None" + by (rule filter_tab_None) + then have "?Dynmethd_def dynC sig=?class_rec (super c) sig" + by (simp add: dynmethd_dynC_def) + with subclseq_super_statC statM None + have "?Dynmethd_def dynC sig = ?Dynmethd_def (super c) sig" + by (auto simp add: empty_def dynmethd_def) + with None subclseq_dynC_statC statM + show ?thesis + by simp + next + case (Some dynM) + note dynM = Some + let ?Termination = "G \ qmdecl sig dynM overrides qmdecl sig statM \ + dynM = statM" + show ?thesis + proof (cases "?filter dynC sig") + case None + with dynM + have no_termination: "\ ?Termination" + by (simp add: filter_tab_def) + from None + have "?Dynmethd_def dynC sig=?class_rec (super c) sig" + by (simp add: dynmethd_dynC_def) + with subclseq_super_statC statM dynM no_termination + show ?thesis + by (auto simp add: empty_def dynmethd_def) + next + case Some + with dynM + have termination: "?Termination" + by (auto) + with Some dynM + have "?Dynmethd_def dynC sig=Some dynM" + by (auto simp add: dynmethd_dynC_def) + with subclseq_super_statC statM dynM termination + show ?thesis + by (auto simp add: dynmethd_def) + qed + qed + qed + qed + qed + qed +qed + +lemma dynmethd_C_C:"\is_class G C; ws_prog G\ +\ dynmethd G C C sig = methd G C sig" +apply (auto simp add: dynmethd_rec) +done + +lemma dynmethdSomeD: + "\dynmethd G statC dynC sig = Some dynM; is_class G dynC; ws_prog G\ + \ G\dynC \\<^sub>C statC \ (\ statM. methd G statC sig = Some statM)" +apply clarify +apply rotate_tac +by (auto simp add: dynmethd_rec) + +lemma dynmethd_Some_cases [consumes 3, case_names Static Overrides]: + (assumes dynM: "dynmethd G statC dynC sig = Some dynM" and + is_cls_dynC: "is_class G dynC" and + ws: "ws_prog G" and + hyp_static: "methd G statC sig = Some dynM \ P" and + hyp_override: "\ statM. \methd G statC sig = Some statM;dynM\statM; + G,sig\dynM overrides statM\ \ P" + ) "P" +proof - + from is_cls_dynC obtain dc where clsDynC: "class G dynC = Some dc" by blast + from clsDynC ws dynM hyp_static hyp_override + show "P" + proof (induct rule: ws_class_induct) + case (Object co) + with ws have "statC = Object" + by (auto simp add: dynmethd_rec) + with ws Object show ?thesis by (auto simp add: dynmethd_C_C) + next + case (Subcls C c) + with ws show ?thesis + by (auto simp add: dynmethd_rec) + qed +qed + +lemma no_override_in_Object: + (assumes dynM: "dynmethd G statC dynC sig = Some dynM" and + is_cls_dynC: "is_class G dynC" and + ws: "ws_prog G" and + statM: "methd G statC sig = Some statM" and + neq_dynM_statM: "dynM\statM" + ) + "dynC \ Object" +proof - + from is_cls_dynC obtain dc where clsDynC: "class G dynC = Some dc" by blast + from clsDynC ws dynM statM neq_dynM_statM + show ?thesis (is "?P dynC") + proof (induct rule: ws_class_induct) + case (Object co) + with ws have "statC = Object" + by (auto simp add: dynmethd_rec) + with ws Object show "?P Object" by (auto simp add: dynmethd_C_C) + next + case (Subcls dynC c) + with ws show "?P dynC" + by (auto simp add: dynmethd_rec) + qed +qed + + +lemma dynmethd_Some_rec_cases [consumes 3, + case_names Static Override Recursion]: +(assumes dynM: "dynmethd G statC dynC sig = Some dynM" and + clsDynC: "class G dynC = Some c" and + ws: "ws_prog G" and + hyp_static: "methd G statC sig = Some dynM \ P" and + hyp_override: "\ statM. \methd G statC sig = Some statM; + methd G dynC sig = Some dynM; statM\dynM; + G,sig\ dynM overrides statM\ \ P" and + + hyp_recursion: "\dynC\Object; + dynmethd G statC (super c) sig = Some dynM\ \ P" + ) "P" +proof - + from clsDynC have "is_class G dynC" by simp + note no_override_in_Object' = no_override_in_Object [OF dynM this ws] + from ws clsDynC dynM hyp_static hyp_override hyp_recursion + show ?thesis + by (auto simp add: dynmethd_rec dest: no_override_in_Object') +qed + +lemma dynmethd_declC: +"\dynmethd G statC dynC sig = Some m; + is_class G statC;ws_prog G + \ \ + (\d. class G (declclass m)=Some d \ table_of (methods d) sig=Some (mthd m)) \ + G\dynC \\<^sub>C (declclass m) \ methd G (declclass m) sig = Some m" +proof - + assume is_cls_statC: "is_class G statC" + assume ws: "ws_prog G" + assume m: "dynmethd G statC dynC sig = Some m" + from m + have "G\dynC \\<^sub>C statC" by (auto simp add: dynmethd_def) + from this is_cls_statC + have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2) + from is_cls_dynC ws m + show ?thesis (is "?P dynC") + proof (induct rule: ws_class_induct') + case (Object co) + with ws have "statC=Object" by (auto simp add: dynmethd_rec) + with ws Object + show "?P Object" + by (auto simp add: dynmethd_C_C dest: methd_declC) + next + case (Subcls dynC c) + assume hyp: "dynmethd G statC (super c) sig = Some m \ ?P (super c)" and + clsDynC: "class G dynC = Some c" and + m': "dynmethd G statC dynC sig = Some m" and + neq_dynC_Obj: "dynC \ Object" + from ws this obtain statM where + subclseq_dynC_statC: "G\dynC \\<^sub>C statC" and + statM: "methd G statC sig = Some statM" + by (blast dest: dynmethdSomeD) + from clsDynC neq_dynC_Obj + have subclseq_dynC_super: "G\dynC \\<^sub>C (super c)" + by (auto intro: subcls1I) + from m' clsDynC ws + show "?P dynC" + proof (cases rule: dynmethd_Some_rec_cases) + case Static + with is_cls_statC ws subclseq_dynC_statC + show ?thesis + by (auto intro: rtrancl_trans dest: methd_declC) + next + case Override + with clsDynC ws + show ?thesis + by (auto dest: methd_declC) + next + case Recursion + with hyp subclseq_dynC_super + show ?thesis + by (auto intro: rtrancl_trans) + qed + qed +qed + +lemma methd_Some_dynmethd_Some: + (assumes statM: "methd G statC sig = Some statM" and + subclseq: "G\dynC \\<^sub>C statC" and + is_cls_statC: "is_class G statC" and + ws: "ws_prog G" + ) "\ dynM. dynmethd G statC dynC sig = Some dynM" + (is "?P dynC") +proof - + from subclseq is_cls_statC + have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2) + then obtain dc where + clsDynC: "class G dynC = Some dc" by blast + from clsDynC ws subclseq + show "?thesis" + proof (induct rule: ws_class_induct) + case (Object co) + with ws have "statC = Object" + by (auto) + with ws Object statM + show "?P Object" + by (auto simp add: dynmethd_C_C) + next + case (Subcls dynC dc) + assume clsDynC': "class G dynC = Some dc" + assume neq_dynC_Obj: "dynC \ Object" + assume hyp: "G\super dc\\<^sub>C statC \ ?P (super dc)" + assume subclseq': "G\dynC\\<^sub>C statC" + then + show "?P dynC" + proof (cases rule: subclseq_cases) + case Eq + with ws statM clsDynC' + show ?thesis + by (auto simp add: dynmethd_rec) + next + case Subcls + assume "G\dynC\\<^sub>C statC" + from this clsDynC' + have "G\super dc\\<^sub>C statC" by (rule subcls_superD) + with hyp ws clsDynC' subclseq' statM + show ?thesis + by (auto simp add: dynmethd_rec) + qed + qed +qed + +lemma dynmethd_cases [consumes 4, case_names Static Overrides]: + (assumes statM: "methd G statC sig = Some statM" and + subclseq: "G\dynC \\<^sub>C statC" and + is_cls_statC: "is_class G statC" and + ws: "ws_prog G" and + hyp_static: "dynmethd G statC dynC sig = Some statM \ P" and + hyp_override: "\ dynM. \dynmethd G statC dynC sig = Some dynM; + dynM\statM; + G,sig\dynM overrides statM\ \ P" + ) "P" +proof - + from subclseq is_cls_statC + have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2) + then obtain dc where + clsDynC: "class G dynC = Some dc" by blast + from statM subclseq is_cls_statC ws + obtain dynM + where dynM: "dynmethd G statC dynC sig = Some dynM" + by (blast dest: methd_Some_dynmethd_Some) + from dynM is_cls_dynC ws + show ?thesis + proof (cases rule: dynmethd_Some_cases) + case Static + with hyp_static dynM statM show ?thesis by simp + next + case Overrides + with hyp_override dynM statM show ?thesis by simp + qed +qed + +lemma ws_dynmethd: + (assumes statM: "methd G statC sig = Some statM" and + subclseq: "G\dynC \\<^sub>C statC" and + is_cls_statC: "is_class G statC" and + ws: "ws_prog G" + ) + "\ dynM. dynmethd G statC dynC sig = Some dynM \ + is_static dynM = is_static statM \ G\resTy dynM\resTy statM" +proof - + from statM subclseq is_cls_statC ws + show ?thesis + proof (cases rule: dynmethd_cases) + case Static + with statM + show ?thesis + by simp + next + case Overrides + with ws + show ?thesis + by (auto dest: ws_overrides_commonD) + qed +qed + +(* +lemma dom_dynmethd: + "dom (dynmethd G statC dynC) \ dom (methd G statC) \ dom (methd G dynC)" +by (auto simp add: dynmethd_def dom_def) + +lemma finite_dom_dynmethd: + "\ws_prog G; is_class G statC; is_class G dynC\ + \ finite (dom (dynmethd G statC dynC))" +apply (rule_tac B="dom (methd G statC) \ dom (methd G dynC)" in finite_subset) +apply (rule dom_dynmethd) +apply (rule finite_UnI) +apply (drule (2) finite_dom_methd)+ +done +*) +(* +lemma dynmethd_SomeD: +"\ws_prog G; is_class G statC; is_class G dynC; + methd G statC sig = Some sm; dynmethd G statC dynC sig = Some dm; sm \ dm + \ \ G\dynC \\<^sub>C statC \ + (declclass dm \ dynC \ G \ dm accessible_through_inheritance_in dynC)" +by (auto simp add: dynmethd_def + dest: methd_inheritedD methd_diff_cls + intro: rtrancl_into_trancl3) +*) + +subsection "dynlookup" + +lemma dynlookup_cases [consumes 1, case_names NullT IfaceT ClassT ArrayT]: +"\dynlookup G statT dynC sig = x; + \statT = NullT ; empty sig = x \ \ P; + \ I. \statT = IfaceT I ; dynimethd G I dynC sig = x\ \ P; + \ statC.\statT = ClassT statC; dynmethd G statC dynC sig = x\ \ P; + \ ty. \statT = ArrayT ty ; dynmethd G Object dynC sig = x\ \ P + \ \ P" +by (cases statT) (auto simp add: dynlookup_def) + +subsection "fields" + +lemma fields_rec: "\class G C = Some c; ws_prog G\ \ + fields G C = map (\(fn,ft). ((fn,C),ft)) (cfields c) @ + (if C = Object then [] else fields G (super c))" +apply (simp only: fields_def) +apply (erule class_rec [THEN trans]) +apply assumption +apply clarsimp +done + +(* local lemma *) +lemma fields_norec: +"\class G fd = Some c; ws_prog G; table_of (cfields c) fn = Some f\ + \ table_of (fields G fd) (fn,fd) = Some f" +apply (subst fields_rec) +apply assumption+ +apply (subst map_of_override [symmetric]) +apply (rule disjI1 [THEN override_Some_iff [THEN iffD2]]) +apply (auto elim: table_of_map2_SomeI) +done + +(* local lemma *) +lemma table_of_fieldsD: +"table_of (map (\(fn,ft). ((fn,C),ft)) (cfields c)) efn = Some f +\ (declclassf efn) = C \ table_of (cfields c) (fname efn) = Some f" +apply (case_tac "efn") +by auto + +lemma fields_declC: + "\table_of (fields G C) efn = Some f; ws_prog G; is_class G C\ \ + (\d. class G (declclassf efn) = Some d \ + table_of (cfields d) (fname efn)=Some f) \ + G\C \\<^sub>C (declclassf efn) \ table_of (fields G (declclassf efn)) efn = Some f" +apply (erule make_imp) +apply (rule ws_subcls1_induct, assumption, assumption) +apply (subst fields_rec, assumption) +apply clarify +apply (simp only: map_of_override [symmetric] del: map_of_override) +apply (case_tac "table_of (map (split (\fn. Pair (fn, Ca))) (cfields c)) efn") +apply (force intro:rtrancl_into_rtrancl2 simp add: override_def) + +apply (frule_tac fd="Ca" in fields_norec) +apply assumption +apply blast +apply (frule table_of_fieldsD) +apply (frule_tac n="table_of (map (split (\fn. Pair (fn, Ca))) (cfields c))" + and m="table_of (if Ca = Object then [] else fields G (super c))" + in override_find_right) +apply (case_tac "efn") +apply (simp) +done + +lemma fields_emptyI: "\y. \ws_prog G; class G C = Some c;cfields c = []; + C \ Object \ class G (super c) = Some y \ fields G (super c) = []\ \ + fields G C = []" +apply (subst fields_rec) +apply assumption +apply auto +done + +(* easier than with table_of *) +lemma fields_mono_lemma: +"\x \ set (fields G C); G\D \\<^sub>C C; ws_prog G\ + \ x \ set (fields G D)" +apply (erule make_imp) +apply (erule converse_rtrancl_induct) +apply fast +apply (drule subcls1D) +apply clarsimp +apply (subst fields_rec) +apply auto +done + + +lemma ws_unique_fields_lemma: + "\(efn,fd) \ set (fields G (super c)); fc \ set (cfields c); ws_prog G; + fname efn = fname fc; declclassf efn = C; + class G C = Some c; C \ Object; class G (super c) = Some d\ \ R" +apply (frule_tac ws_prog_cdeclD [THEN conjunct2], assumption, assumption) +apply (drule_tac weak_map_of_SomeI) +apply (frule_tac subcls1I [THEN subcls1_irrefl], assumption, assumption) +apply (auto dest: fields_declC [THEN conjunct2 [THEN conjunct1[THEN rtranclD]]]) +done + +lemma ws_unique_fields: "\is_class G C; ws_prog G; + \C c. \class G C = Some c\ \ unique (cfields c) \ \ + unique (fields G C)" +apply (rule ws_subcls1_induct, assumption, assumption) +apply (subst fields_rec, assumption) +apply (auto intro!: unique_map_inj injI + elim!: unique_append ws_unique_fields_lemma fields_norec + ) +done + + +subsection "accfield" + +lemma accfield_fields: + "accfield G S C fn = Some f + \ table_of (fields G C) (fn, declclass f) = Some (fld f)" +apply (simp only: accfield_def Let_def) +apply (rule table_of_remap_SomeD) +apply (auto dest: filter_tab_SomeD) +done + + +lemma accfield_declC_is_class: + "\is_class G C; accfield G S C en = Some (fd, f); ws_prog G\ \ + is_class G fd" +apply (drule accfield_fields) +apply (drule fields_declC [THEN conjunct1], assumption) +apply auto +done + +lemma accfield_accessibleD: + "accfield G S C fn = Some f \ G\Field fn f of C accessible_from S" +by (auto simp add: accfield_def Let_def) + +subsection "is methd" + +lemma is_methdI: +"\class G C = Some y; methd G C sig = Some b\ \ is_methd G C sig" +apply (unfold is_methd_def) +apply auto +done + +lemma is_methdD: +"is_methd G C sig \ class G C \ None \ methd G C sig \ None" +apply (unfold is_methd_def) +apply auto +done + +lemma finite_is_methd: + "ws_prog G \ finite (Collect (split (is_methd G)))" +apply (unfold is_methd_def) +apply (subst SetCompr_Sigma_eq) +apply (rule finite_is_class [THEN finite_SigmaI]) +apply (simp only: mem_Collect_eq) +apply (fold dom_def) +apply (erule finite_dom_methd) +apply assumption +done + +section "calculation of the superclasses of a class" + +constdefs + superclasses:: "prog \ qtname \ qtname set" + "superclasses G C \ class_rec (G,C) {} + (\ C c superclss. (if C=Object + then {} + else insert (super c) superclss))" + +lemma superclasses_rec: "\class G C = Some c; ws_prog G\ \ + superclasses G C + = (if (C=Object) + then {} + else insert (super c) (superclasses G (super c)))" +apply (unfold superclasses_def) +apply (erule class_rec [THEN trans], assumption) +apply (simp) +done + +lemma superclasses_mono: +"\G\C \\<^sub>C D;ws_prog G; class G C = Some c; + \ C c. \class G C = Some c;C\Object\ \ \ sc. class G (super c) = Some sc; + x\superclasses G D +\ \ x\superclasses G C" +proof - + + assume ws: "ws_prog G" and + cls_C: "class G C = Some c" and + wf: "\C c. \class G C = Some c; C \ Object\ + \ \sc. class G (super c) = Some sc" + assume clsrel: "G\C\\<^sub>C D" + thus "\ c. \class G C = Some c; x\superclasses G D\\ + x\superclasses G C" (is "PROP ?P C" + is "\ c. ?CLS C c \ ?SUP D \ ?SUP C") + proof (induct ?P C rule: converse_trancl_induct) + fix C c + assume "G\C\\<^sub>C\<^sub>1D" "class G C = Some c" "x \ superclasses G D" + with wf ws show "?SUP C" + by (auto intro: no_subcls1_Object + simp add: superclasses_rec subcls1_def) + next + fix C S c + assume clsrel': "G\C \\<^sub>C\<^sub>1 S" "G\S \\<^sub>C D" + and hyp : "\ s. \class G S = Some s; x \ superclasses G D\ + \ x \ superclasses G S" + and cls_C': "class G C = Some c" + and x: "x \ superclasses G D" + moreover note wf ws + moreover from calculation + have "?SUP S" + by (force intro: no_subcls1_Object simp add: subcls1_def) + moreover from calculation + have "super c = S" + by (auto intro: no_subcls1_Object simp add: subcls1_def) + ultimately show "?SUP C" + by (auto intro: no_subcls1_Object simp add: superclasses_rec) + qed +qed + +lemma subclsEval: +"\G\C \\<^sub>C D;ws_prog G; class G C = Some c; + \ C c. \class G C = Some c;C\Object\ \ \ sc. class G (super c) = Some sc + \ \ D\superclasses G C" +proof - + note converse_trancl_induct + = converse_trancl_induct [consumes 1,case_names Single Step] + assume + ws: "ws_prog G" and + cls_C: "class G C = Some c" and + wf: "\C c. \class G C = Some c; C \ Object\ + \ \sc. class G (super c) = Some sc" + assume clsrel: "G\C\\<^sub>C D" + thus "\ c. class G C = Some c\ D\superclasses G C" + (is "PROP ?P C" is "\ c. ?CLS C c \ ?SUP C") + proof (induct ?P C rule: converse_trancl_induct) + fix C c + assume "G\C \\<^sub>C\<^sub>1 D" "class G C = Some c" + with ws wf show "?SUP C" + by (auto intro: no_subcls1_Object simp add: superclasses_rec subcls1_def) + next + fix C S c + assume "G\C \\<^sub>C\<^sub>1 S" "G\S\\<^sub>C D" + "\s. class G S = Some s \ D \ superclasses G S" + "class G C = Some c" + with ws wf show "?SUP C" + by - (rule superclasses_mono, + auto dest: no_subcls1_Object simp add: subcls1_def ) + qed +qed + +end \ No newline at end of file diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Eval.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Eval.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,993 @@ +(* Title: isabelle/Bali/Eval.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* Operational evaluation (big-step) semantics of Java expressions and + statements +*} + +theory Eval = State + DeclConcepts: + +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" + +syntax (xsymbols) + dummy_res :: "vals" ("\") +translations + "\" == "In1 Unit" + +constdefs + arbitrary3 :: "('al + 'ar, 'b, 'c) sum3 \ vals" + "arbitrary3 \ sum3_case (In1 \ 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" ("_,_\_ fits _"[61,61,61,61]60) + "G,s\a' fits T \ (\rt. T=RefT rt) \ a'=Null \ G\obj_ty(lookup_obj s a')\T" + +lemma fits_Null [simp]: "G,s\Null fits T" +by (simp add: fits_def) + + +lemma fits_Addr_RefT [simp]: + "G,s\Addr a fits RefT t = G\obj_ty (the (heap s a))\RefT t" +by (simp add: fits_def) + +lemma fitsD: "\X. G,s\a' fits T \ (\pt. T = PrimT pt) \ + (\t. T = RefT t) \ a' = Null \ + (\t. T = RefT t) \ a' \ Null \ G\obj_ty (lookup_obj s a')\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 +done + +constdefs + catch ::"prog \ state \ qtname \ bool" ("_,_\catch _"[61,61,61]60) + "G,s\catch C\\xc. abrupt s=Some (Xcpt xc) \ + G,store s\Addr (the_Loc xc) fits Class C" + +lemma catch_Norm [simp]: "\G,Norm s\catch tn" +apply (unfold catch_def) +apply (simp (no_asm)) +done + +lemma catch_XcptLoc [simp]: + "G,(Some (Xcpt (Loc a)),s)\catch C = G,s\Addr a fits Class C" +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\Addr (the_Loc (the_Xcpt (the x)))) s)" + +lemma new_xcpt_var_def2 [simp]: + "new_xcpt_var vn (x,s) = + Norm (lupd(VName vn\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_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 + +(* + target :: "inv_mode \ st \ val \ ref_ty \ qtname" + "target m s a' t + \ if m = IntVir + then obj_class (lookup_obj s a') + else the_Class (RefT t)" +*) + + 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_notIntVir [simp]: + "m \ IntVir \ invocation_class m 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 \ (init_vals (table_of (lcls (mbody m))) + ((pars m)[\]pvs)) v + | Res \ Some (default_val (resTy m))) + | 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: "init_lvars G C sig mode a' pvs (x,s) = + set_lvars + (\ k. + (case k of + EName e + \ (case e of + VNam v + \ (init_vals + (table_of (lcls (mbody (mthd (the (methd G C sig)))))) + ((pars (mthd (the (methd G C sig))))[\]pvs)) v + | Res \ Some (default_val (resTy (mthd (the (methd G C sig)))))) + | 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: +"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\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\v fits T) + ArrStore x + ,upd_gobj oref n v s)) + in ((the (cs n),f) + ,abupd (raise_if (\i in_bounds k) IndOutBound \ np a') s)" + +lemma fvar_def2: "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: "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'\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 \ np a') + s)" +apply (unfold avar_def) +apply (simp (no_asm) add: Let_def split_beta) +done + + +section "evaluation judgments" + +consts + eval :: "prog \ (state \ term \ vals \ state) set" + halloc:: "prog \ (state \ obj_tag \ loc \ state) set" + sxalloc:: "prog \ (state \ state) set" + + +syntax +eval ::"[prog,state,term,vals*state]=>bool"("_|-_ -_>-> _" [61,61,80, 61]60) +exec ::"[prog,state,stmt ,state]=>bool"("_|-_ -_-> _" [61,61,65, 61]60) +evar ::"[prog,state,var ,vvar,state]=>bool"("_|-_ -_=>_-> _"[61,61,90,61,61]60) +eval_::"[prog,state,expr ,val, state]=>bool"("_|-_ -_->_-> _"[61,61,80,61,61]60) +evals::"[prog,state,expr list , + val list ,state]=>bool"("_|-_ -_#>_-> _"[61,61,61,61,61]60) +hallo::"[prog,state,obj_tag, + loc,state]=>bool"("_|-_ -halloc _>_-> _"[61,61,61,61,61]60) +sallo::"[prog,state ,state]=>bool"("_|-_ -sxalloc-> _"[61,61, 61]60) + +syntax (xsymbols) +eval ::"[prog,state,term,vals\state]\bool" ("_\_ \_\\ _" [61,61,80, 61]60) +exec ::"[prog,state,stmt ,state]\bool"("_\_ \_\ _" [61,61,65, 61]60) +evar ::"[prog,state,var ,vvar,state]\bool"("_\_ \_=\_\ _"[61,61,90,61,61]60) +eval_::"[prog,state,expr ,val ,state]\bool"("_\_ \_-\_\ _"[61,61,80,61,61]60) +evals::"[prog,state,expr list , + val list ,state]\bool"("_\_ \_\\_\ _"[61,61,61,61,61]60) +hallo::"[prog,state,obj_tag, + loc,state]\bool"("_\_ \halloc _\_\ _"[61,61,61,61,61]60) +sallo::"[prog,state, state]\bool"("_\_ \sxalloc\ _"[61,61, 61]60) + +translations + "G\s \t \\ w___s' " == "(s,t,w___s') \ eval G" + "G\s \t \\ (w, s')" <= "(s,t,w, s') \ eval G" + "G\s \t \\ (w,x,s')" <= "(s,t,w,x,s') \ eval G" + "G\s \c \ (x,s')" <= "G\s \In1r c\\ (\,x,s')" + "G\s \c \ s' " == "G\s \In1r c\\ (\ , s')" + "G\s \e-\v \ (x,s')" <= "G\s \In1l e\\ (In1 v ,x,s')" + "G\s \e-\v \ s' " == "G\s \In1l e\\ (In1 v , s')" + "G\s \e=\vf\ (x,s')" <= "G\s \In2 e\\ (In2 vf,x,s')" + "G\s \e=\vf\ s' " == "G\s \In2 e\\ (In2 vf, s')" + "G\s \e\\v \ (x,s')" <= "G\s \In3 e\\ (In3 v ,x,s')" + "G\s \e\\v \ s' " == "G\s \In3 e\\ (In3 v , s')" + "G\s \halloc oi\a\ (x,s')" <= "(s,oi,a,x,s') \ halloc G" + "G\s \halloc oi\a\ s' " == "(s,oi,a, s') \ halloc G" + "G\s \sxalloc\ (x,s')" <= "(s ,x,s') \ sxalloc G" + "G\s \sxalloc\ s' " == "(s , s') \ sxalloc G" + +inductive "halloc G" intros (* allocating objects on the heap, cf. 12.5 *) + + Abrupt: + "G\(Some x,s) \halloc oi\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\Norm s \halloc oi\a\ (x,init_obj G oi' (Heap a) s)" + +inductive "sxalloc G" intros (* allocating exception objects for + standard exceptions (other than OutOfMemory) *) + + Norm: "G\ Norm s \sxalloc\ Norm s" + + XcptL: "G\(Some (Xcpt (Loc a) ),s) \sxalloc\ (Some (Xcpt (Loc a)),s)" + + SXcpt: "\G\Norm s0 \halloc (CInst (SXcpt xn))\a\ (x,s1)\ \ + G\(Some (Xcpt (Std xn)),s0) \sxalloc\ (Some (Xcpt (Loc a)),s1)" + + +inductive "eval G" intros + +(* propagation of abrupt completion *) + + (* cf. 14.1, 15.5 *) + Abrupt: + "G\(Some xc,s) \t\\ (arbitrary3 t,(Some xc,s))" + + +(* execution of statements *) + + (* cf. 14.5 *) + Skip: "G\Norm s \Skip\ Norm s" + + (* cf. 14.7 *) + Expr: "\G\Norm s0 \e-\v\ s1\ \ + G\Norm s0 \Expr e\ s1" + + Lab: "\G\Norm s0 \c \ s1\ \ + G\Norm s0 \l\ c\ abupd (absorb (Break l)) s1" + (* cf. 14.2 *) + Comp: "\G\Norm s0 \c1 \ s1; + G\ s1 \c2 \ s2\ \ + G\Norm s0 \c1;; c2\ s2" + + + (* cf. 14.8.2 *) + If: "\G\Norm s0 \e-\b\ s1; + G\ s1\(if the_Bool b then c1 else c2)\ s2\ \ + G\Norm s0 \If(e) c1 Else c2 \ s2" + + (* cf. 14.10, 14.10.1 *) + (* G\Norm s0 \If(e) (c;; While(e) c) Else Skip\ s3 *) + (* A "continue jump" from the while body c is handled by + this rule. If a continue jump with the proper label was invoked inside 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 (Lab l (while\). + *) + Loop: "\G\Norm s0 \e-\b\ s1; + if normal s1 \ the_Bool b + then (G\s1 \c\ s2 \ + G\(abupd (absorb (Cont l)) s2) \l\ While(e) c\ s3) + else s3 = s1\ \ + G\Norm s0 \l\ While(e) c\ s3" + + Do: "G\Norm s \Do j\ (Some (Jump j), s)" + + (* cf. 14.16 *) + Throw: "\G\Norm s0 \e-\a'\ s1\ \ + G\Norm s0 \Throw e\ abupd (throw a') s1" + + (* cf. 14.18.1 *) + Try: "\G\Norm s0 \c1\ s1; G\s1 \sxalloc\ s2; + if G,s2\catch C then G\new_xcpt_var vn s2 \c2\ s3 else s3 = s2\ \ + G\Norm s0 \Try c1 Catch(C vn) c2\ s3" + + (* cf. 14.18.2 *) + Fin: "\G\Norm s0 \c1\ (x1,s1); + G\Norm s1 \c2\ s2\ \ + G\Norm s0 \c1 Finally c2\ abupd (abrupt_if (x1\None) x1) s2" + + (* cf. 12.4.2, 8.5 *) + 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))\ s1 \ + G\set_lvars empty s1 \init c\ s2 \ s3 = restore_lvars s1 s2)\ + \ + G\Norm s0 \Init C\ s3" + +(* evaluation of expressions *) + + (* cf. 15.8.1, 12.4.1 *) + NewC: "\G\Norm s0 \Init C\ s1; + G\ s1 \halloc (CInst C)\a\ s2\ \ + G\Norm s0 \NewC C-\Addr a\ s2" + + (* cf. 15.9.1, 12.4.1 *) + NewA: "\G\Norm s0 \init_comp_ty T\ s1; G\s1 \e-\i'\ s2; + G\abupd (check_neg i') s2 \halloc (Arr T (the_Intg i'))\a\ s3\ \ + G\Norm s0 \New T[e]-\Addr a\ s3" + + (* cf. 15.15 *) + Cast: "\G\Norm s0 \e-\v\ s1; + s2 = abupd (raise_if (\G,store s1\v fits T) ClassCast) s1\ \ + G\Norm s0 \Cast T e-\v\ s2" + + (* cf. 15.19.2 *) + Inst: "\G\Norm s0 \e-\v\ s1; + b = (v\Null \ G,store s1\v fits RefT T)\ \ + G\Norm s0 \e InstOf T-\Bool b\ s1" + + (* cf. 15.7.1 *) + Lit: "G\Norm s \Lit v-\v\ Norm s" + + + + (* cf. 15.10.2 *) + Super: "G\Norm s \Super-\val_this s\ Norm s" + + (* cf. 15.2 *) + Acc: "\G\Norm s0 \va=\(v,f)\ s1\ \ + G\Norm s0 \Acc va-\v\ s1" + + (* cf. 15.25.1 *) + Ass: "\G\Norm s0 \va=\(w,f)\ s1; + G\ s1 \e-\v \ s2\ \ + G\Norm s0 \va:=e-\v\ assign f v s2" + + (* cf. 15.24 *) + Cond: "\G\Norm s0 \e0-\b\ s1; + G\ s1 \(if the_Bool b then e1 else e2)-\v\ s2\ \ + G\Norm s0 \e0 ? e1 : e2-\v\ s2" + + + (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5 *) + Call: + "\G\Norm s0 \e-\a'\ s1; G\s1 \args\\vs\ 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\ s3\ + \ + G\Norm s0 \{statT,mode}e\mn({pTs}args)-\v\ (restore_lvars s2 s3)" + + Methd: "\G\Norm s0 \body G D sig-\v\ s1\ \ + G\Norm s0 \Methd D sig-\v\ s1" + + (* cf. 14.15, 12.4.1 *) + Body: "\G\Norm s0 \Init D\ s1; G\s1 \c\ s2\ \ + G\Norm s0 \Body D c -\the (locals (store s2) Result)\abupd (absorb Ret) s2" + +(* evaluation of variables *) + + (* cf. 15.13.1, 15.7.2 *) + LVar: "G\Norm s \LVar vn=\lvar vn s\ Norm s" + + (* cf. 15.10.1, 12.4.1 *) + FVar: "\G\Norm s0 \Init C\ s1; G\s1 \e-\a\ s2; + (v,s2') = fvar C stat fn a s2\ \ + G\Norm s0 \{C,stat}e..fn=\v\ s2'" + + (* cf. 15.12.1, 15.25.1 *) + AVar: "\G\ Norm s0 \e1-\a\ s1; G\s1 \e2-\i\ s2; + (v,s2') = avar G i a s2\ \ + G\Norm s0 \e1.[e2]=\v\ s2'" + + +(* evaluation of expression lists *) + + (* cf. 15.11.4.2 *) + Nil: + "G\Norm s0 \[]\\[]\ Norm s0" + + (* cf. 15.6.4 *) + Cons: "\G\Norm s0 \e -\ v \ s1; + G\ s1 \es\\vs\ s2\ \ + G\Norm s0 \e#es\\v#vs\ s2" + +(* Rearrangement of premisses: +[0,1(Abrupt),2(Skip),8(Do),4(Lab),28(Nil),29(Cons),25(LVar),15(Cast),16(Inst), + 17(Lit),18(Super),19(Acc),3(Expr),5(Comp),23(Methd),24(Body),21(Cond),6(If), + 7(Loop),11(Fin),9(Throw),13(NewC),14(NewA),12(Init),20(Ass),10(Try),26(FVar), + 27(AVar),22(Call)] +*) +ML {* +bind_thm ("eval_induct_", rearrange_prems +[0,1,2,8,4,28,29,25,15,16, + 17,18,19,3,5,23,24,21,6, + 7,11,9,13,14,12,20,10,26, + 27,22] (thm "eval.induct")) +*} + +lemmas eval_induct = eval_induct_ [split_format 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\(Some xc,s) \halloc oi\a\ s'" + "G\(Norm s) \halloc oi\a\ s'" + +inductive_cases sxalloc_elim_cases: + "G\ Norm s \sxalloc\ s'" + "G\(Some (Xcpt (Loc a )),s) \sxalloc\ s'" + "G\(Some (Xcpt (Std xn)),s) \sxalloc\ s'" +inductive_cases sxalloc_cases: "G\s \sxalloc\ s'" + +lemma sxalloc_elim_cases2: "\G\s \sxalloc\ s'; + \s. \s' = Norm 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] +ML_setup {* +simpset_ref() := simpset() delloop "split_all_tac" +*} +inductive_cases eval_cases: "G\s \t\\ vs'" + +inductive_cases eval_elim_cases: + "G\(Some xc,s) \t \\ vs'" + "G\Norm s \In1r Skip \\ xs'" + "G\Norm s \In1r (Do j) \\ xs'" + "G\Norm s \In1r (l\ c) \\ xs'" + "G\Norm s \In3 ([]) \\ vs'" + "G\Norm s \In3 (e#es) \\ vs'" + "G\Norm s \In1l (Lit w) \\ vs'" + "G\Norm s \In2 (LVar vn) \\ vs'" + "G\Norm s \In1l (Cast T e) \\ vs'" + "G\Norm s \In1l (e InstOf T) \\ vs'" + "G\Norm s \In1l (Super) \\ vs'" + "G\Norm s \In1l (Acc va) \\ vs'" + "G\Norm s \In1r (Expr e) \\ xs'" + "G\Norm s \In1r (c1;; c2) \\ xs'" + "G\Norm s \In1l (Methd C sig) \\ xs'" + "G\Norm s \In1l (Body D c) \\ xs'" + "G\Norm s \In1l (e0 ? e1 : e2) \\ vs'" + "G\Norm s \In1r (If(e) c1 Else c2) \\ xs'" + "G\Norm s \In1r (l\ While(e) c) \\ xs'" + "G\Norm s \In1r (c1 Finally c2) \\ xs'" + "G\Norm s \In1r (Throw e) \\ xs'" + "G\Norm s \In1l (NewC C) \\ vs'" + "G\Norm s \In1l (New T[e]) \\ vs'" + "G\Norm s \In1l (Ass va e) \\ vs'" + "G\Norm s \In1r (Try c1 Catch(tn vn) c2) \\ xs'" + "G\Norm s \In2 ({C,stat}e..fn) \\ vs'" + "G\Norm s \In2 (e1.[e2]) \\ vs'" + "G\Norm s \In1l ({statT,mode}e\mn({pT}p)) \\ vs'" + "G\Norm s \In1r (Init C) \\ xs'" +declare not_None_eq [simp] (* IntDef.Zero_def [simp] *) +declare split_paired_All [simp] split_paired_Ex [simp] +ML_setup {* +simpset_ref() := simpset() 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\s \t\\ (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 + +ML_setup {* +fun eval_fun nam inj rhs = +let + val name = "eval_" ^ nam ^ "_eq" + val lhs = "G\s \" ^ inj ^ " t\\ (w, s')" + val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") + (K [Auto_tac, ALLGOALS (ftac (thm "eval_Inj_elim")) THEN Auto_tac]) + fun is_Inj (Const (inj,_) $ _) = true + | is_Inj _ = false + fun pred (_ $ (Const ("Pair",_) $ _ $ + (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ x $ _ ))) $ _ ) = is_Inj x +in + make_simproc name lhs pred (thm name) +end + +val eval_expr_proc =eval_fun "expr" "In1l" "\v. w=In1 v \ G\s \t-\v \ s'" +val eval_var_proc =eval_fun "var" "In2" "\vf. w=In2 vf \ G\s \t=\vf\ s'" +val eval_exprs_proc=eval_fun "exprs""In3" "\vs. w=In3 vs \ G\s \t\\vs\ s'" +val eval_stmt_proc =eval_fun "stmt" "In1r" " w=\ \ G\s \t \ s'"; +Addsimprocs [eval_expr_proc,eval_var_proc,eval_exprs_proc,eval_stmt_proc]; +bind_thms ("AbruptIs", sum3_instantiate (thm "eval.Abrupt")) +*} + +declare halloc.Abrupt [intro!] eval.Abrupt [intro!] AbruptIs [intro!] + + +lemma eval_no_abrupt_lemma: + "\s s'. G\s \t\\ (w,s') \ normal s' \ normal s" +by (erule eval_cases, auto) + +lemma eval_no_abrupt: + "G\(x,s) \t\\ (w,Norm s') = + (x = None \ G\Norm s \t\\ (w,Norm s'))" +apply auto +apply (frule eval_no_abrupt_lemma, auto)+ +done + +ML {* +local + fun is_None (Const ("Option.option.None",_)) = true + | is_None _ = false + fun pred (t as (_ $ (Const ("Pair",_) $ + (Const ("Pair", _) $ x $ _) $ _ ) $ _)) = is_None x +in + val eval_no_abrupt_proc = + make_simproc "eval_no_abrupt" "G\(x,s) \e\\ (w,Norm s')" pred + (thm "eval_no_abrupt") +end; +Addsimprocs [eval_no_abrupt_proc] +*} + + +lemma eval_abrupt_lemma: + "G\s \t\\ (v,s') \ abrupt s=Some xc \ s'= s \ v = arbitrary3 t" +by (erule eval_cases, auto) + +lemma eval_abrupt: + " G\(Some xc,s) \t\\ (w,s') = + (s'=(Some xc,s) \ w=arbitrary3 t \ + G\(Some xc,s) \t\\ (arbitrary3 t,(Some xc,s)))" +apply auto +apply (frule eval_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", _) $ _ $ + x))) $ _ ) = is_Some x +in + val eval_abrupt_proc = + make_simproc "eval_abrupt" + "G\(Some xc,s) \e\\ (w,s')" pred (thm "eval_abrupt") +end; +Addsimprocs [eval_abrupt_proc] +*} + + +lemma LitI: "G\s \Lit v-\(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\s \Skip\ s" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: eval.Skip) + +lemma ExprI: "G\s \e-\v\ s' \ G\s \Expr e\ s'" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: eval.Expr) + +lemma CompI: "\G\s \c1\ s1; G\s1 \c2\ s2\ \ G\s \c1;; c2\ s2" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: eval.Comp) + +lemma CondI: + "\s1. \G\s \e-\b\ s1; G\s1 \(if the_Bool b then e1 else e2)-\v\ s2\ \ + G\s \e ? e1 : e2-\(if normal s1 then v else arbitrary)\ s2" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: eval.Cond) + +lemma IfI: "\G\s \e-\v\ s1; G\s1 \(if the_Bool v then c1 else c2)\ s2\ + \ G\s \If(e) c1 Else c2\ s2" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: eval.If) + +lemma MethdI: "G\s \body G C sig-\v\ s' \ G\s \Methd C sig-\v\ s'" +apply (case_tac "s", case_tac "a = None") +by (auto intro!: eval.Methd) + +lemma eval_Call: "\G\Norm s0 \e-\a'\ s1; G\s1 \ps\\pvs\ 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' pvs s2 + \Methd D \name=mn,parTs=pTs\-\ v\ s3; + s3' = restore_lvars s2 s3\ \ + G\Norm s0 \{statT,mode}e\mn({pTs}ps)-\v\ s3'" +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\Norm (init_class_obj G C s0) + \(if C = Object then Skip else Init (super (the (class G C))))\ s1 \ + G\set_lvars empty s1 \(init (the (class G C)))\ s2 \ + s3 = restore_lvars s1 s2\ \ + G\Norm s0 \Init C\ s3" +apply (rule eval.Init) +apply auto +done + +lemma init_done: "initd C s \ G\s \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\s \StatRef rt-\(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\s \Skip\ s' \ s' = s" +apply (erule eval_cases) +by auto + +lemma Skip_eq [simp]: "G\s \Skip\ s' = (s = s')" +by auto + +(*unused*) +lemma init_retains_locals [rule_format (no_asm)]: "G\s \t\\ (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\(Some xc,s) \halloc oi\a\ s' \ s'=(Some xc,s)" +apply (erule_tac halloc_elim_cases) +by auto + +(* +G\(x,(h,l)) \e\v\ (x',(h',l'))) \ l This = l' This" +G\(x,(h,l)) \s \ (x',(h',l'))) \ l This = l' This" +*) + +lemma eval_Methd: + "G\s \In1l(body G C sig)\\ (w,s') \ G\s \In1l(Methd C sig)\\ (w,s')" +apply (case_tac "s") +apply (case_tac "a") +apply clarsimp+ +apply (erule eval.Methd) +apply (drule eval_abrupt_lemma) +apply force +done + + +section "single valued" + +lemma unique_halloc [rule_format (no_asm)]: + "\s as as'. (s,oi,as)\halloc G \ (s,oi,as')\halloc G \ as'=as" +apply (simp (no_asm_simp) only: split_tupled_all) +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\s \halloc oi\a \ s'}" +apply (unfold single_valued_def) +by (clarsimp, drule (1) unique_halloc, auto) + + +lemma unique_sxalloc [rule_format (no_asm)]: + "\s s'. G\s \sxalloc\ s' \ G\s \sxalloc\ s'' \ s'' = s'" +apply (simp (no_asm_simp) only: split_tupled_all) +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\s \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\s \t\\ ws \ (\ws'. G\s \t\\ ws' \ ws' = ws)" +apply (case_tac "ws") +apply (simp only:) +apply (erule thin_rl) +apply (erule eval_induct) +apply (tactic {* ALLGOALS (EVERY' + [strip_tac, rotate_tac ~1, eresolve_tac (thms "eval_elim_cases")]) *}) +(* 29 subgoals *) +prefer 26 (* Try *) +apply (simp (no_asm_use) only: split add: split_if_asm) +(* 32 subgoals *) +prefer 28 (* Init *) +apply (case_tac "inited C (globs s0)", (simp only: if_True if_False)+) +prefer 24 (* While *) +apply (simp (no_asm_use) only: split add: split_if_asm, blast) +apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp) +apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp) +apply blast +(* 31 subgoals *) +apply (blast dest: unique_sxalloc unique_halloc split_pairD)+ +done + +(* unused *) +lemma single_valued_eval: + "single_valued {((s,t),vs'). G\s \t\\ vs'}" +apply (unfold single_valued_def) +by (clarify, drule (1) unique_eval, auto) + +end 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 diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Example.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Example.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,1248 @@ +(* Title: isabelle/Bali/Example.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* Example Bali program *} + +theory Example = Eval + WellForm: + + +text {* +The following example Bali program includes: +\begin{itemize} +\item class and interface declarations with inheritance, hiding of fields, + overriding of methods (with refined result type), array type, +\item method call (with dynamic binding), parameter access, return expressions, +\item expression statements, sequential composition, literal values, + local assignment, local access, field assignment, type cast, +\item exception generation and propagation, try and catch statement, throw + statement +\item instance creation and (default) static initialization +\end{itemize} +\begin{verbatim} +package java_lang + +public interface HasFoo { + public Base foo(Base z); +} + +public class Base implements HasFoo { + static boolean arr[] = new boolean[2]; + public HasFoo vee; + public Base foo(Base z) { + return z; + } +} + +public class Ext extends Base { + public int vee; + public Ext foo(Base z) { + ((Ext)z).vee = 1; + return null; + } +} + +public class Example { + public static void main(String args[]) throws Throwable { + Base e = new Ext(); + try {e.foo(null); } + catch(NullPointerException z) { + while(Ext.arr[2]) ; + } + } +} +\end{verbatim} +*} + +declare widen.null [intro] + +lemma wf_fdecl_def2: "\fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))" +apply (unfold wf_fdecl_def) +apply (simp (no_asm)) +done + +declare wf_fdecl_def2 [iff] + + +section "type and expression names" + +(** unfortunately cannot simply instantiate tnam **) +datatype tnam_ = HasFoo_ | Base_ | Ext_ +datatype vnam_ = arr_ | vee_ | z_ | e_ +datatype label_ = lab1_ + +consts + + tnam_ :: "tnam_ \ tnam" + vnam_ :: "vnam_ \ vname" + label_:: "label_ \ label" +axioms (** tnam_, vnam_ and label are intended to be isomorphic + to tnam, vname and label **) + + inj_tnam_ [simp]: "(tnam_ x = tnam_ y) = (x = y)" + inj_vnam_ [simp]: "(vnam_ x = vnam_ y) = (x = y)" + inj_label_ [simp]: "(label_ x = label_ y) = (x = y)" + + + surj_tnam_: "\m. n = tnam_ m" + surj_vnam_: "\m. n = vnam_ m" + surj_label_:" \m. n = label_ m" + +syntax + + HasFoo :: qtname + Base :: qtname + Ext :: qtname + arr :: ename + vee :: ename + z :: ename + e :: ename + lab1:: label +translations + + "HasFoo" == "\pid=java_lang,tid=TName (tnam_ HasFoo_)\" + "Base" == "\pid=java_lang,tid=TName (tnam_ Base_)\" + "Ext" == "\pid=java_lang,tid=TName (tnam_ Ext_)\" + "arr" == "(vnam_ arr_)" + "vee" == "(vnam_ vee_)" + "z" == "(vnam_ z_)" + "e" == "(vnam_ e_)" + "lab1" == "label_ lab1_" + + +lemma neq_Base_Object [simp]: "Base\Object" +by (simp add: Object_def) + +lemma neq_Ext_Object [simp]: "Ext\Object" +by (simp add: Object_def) + +lemma neq_Base_SXcpt [simp]: "Base\SXcpt xn" +by (simp add: SXcpt_def) + +lemma neq_Ext_SXcpt [simp]: "Ext\SXcpt xn" +by (simp add: SXcpt_def) + +section "classes and interfaces" + +defs + + Object_mdecls_def: "Object_mdecls \ []" + SXcpt_mdecls_def: "SXcpt_mdecls \ []" + +consts + + foo :: mname + +constdefs + + foo_sig :: sig + "foo_sig \ \name=foo,parTs=[Class Base]\" + + foo_mhead :: mhead + "foo_mhead \ \access=Public,static=False,pars=[z],resT=Class Base\" + +constdefs + + Base_foo :: mdecl + "Base_foo \ (foo_sig, \access=Public,static=False,pars=[z],resT=Class Base, + mbody=\lcls=[],stmt=Return (!!z)\\)" + + Ext_foo :: mdecl + "Ext_foo \ (foo_sig, + \access=Public,static=False,pars=[z],resT=Class Ext, + mbody=\lcls=[] + ,stmt=Expr({Ext,False}Cast (Class Ext) (!!z)..vee := + Lit (Intg 1))\ + \)" + +constdefs + +arr_viewed_from :: "qtname \ var" +"arr_viewed_from C \ {Base,True}StatRef (ClassT C)..arr" + +BaseCl :: class +"BaseCl \ \access=Public, + cfields=[(arr, \access=Public,static=True ,type=PrimT Boolean.[]\), + (vee, \access=Public,static=False,type=Iface HasFoo \)], + methods=[Base_foo], + init=Expr(arr_viewed_from Base :=New (PrimT Boolean)[Lit (Intg 2)]), + super=Object, + superIfs=[HasFoo]\" + +ExtCl :: class +"ExtCl \ \access=Public, + cfields=[(vee, \access=Public,static=False,type= PrimT Integer\)], + methods=[Ext_foo], + init=Skip, + super=Base, + superIfs=[]\" + +constdefs + + HasFooInt :: iface + "HasFooInt \ \access=Public,imethods=[(foo_sig, foo_mhead)],isuperIfs=[]\" + + Ifaces ::"idecl list" + "Ifaces \ [(HasFoo,HasFooInt)]" + + "Classes" ::"cdecl list" + "Classes \ [(Base,BaseCl),(Ext,ExtCl)]@standard_classes" + +lemmas table_classes_defs = + Classes_def standard_classes_def ObjectC_def SXcptC_def + +lemma table_ifaces [simp]: "table_of Ifaces = empty(HasFoo\HasFooInt)" +apply (unfold Ifaces_def) +apply (simp (no_asm)) +done + +lemma table_classes_Object [simp]: + "table_of Classes Object = Some \access=Public,cfields=[] + ,methods=Object_mdecls + ,init=Skip,super=arbitrary,superIfs=[]\" +apply (unfold table_classes_defs) +apply (simp (no_asm) add:Object_def) +done + +lemma table_classes_SXcpt [simp]: + "table_of Classes (SXcpt xn) + = Some \access=Public,cfields=[],methods=SXcpt_mdecls, + init=Skip, + super=if xn = Throwable then Object else SXcpt Throwable, + superIfs=[]\" +apply (unfold table_classes_defs) +apply (induct_tac xn) +apply (simp add: Object_def SXcpt_def)+ +done + +lemma table_classes_HasFoo [simp]: "table_of Classes HasFoo = None" +apply (unfold table_classes_defs) +apply (simp (no_asm) add: Object_def SXcpt_def) +done + +lemma table_classes_Base [simp]: "table_of Classes Base = Some BaseCl" +apply (unfold table_classes_defs ) +apply (simp (no_asm) add: Object_def SXcpt_def) +done + +lemma table_classes_Ext [simp]: "table_of Classes Ext = Some ExtCl" +apply (unfold table_classes_defs ) +apply (simp (no_asm) add: Object_def SXcpt_def) +done + + +section "program" + +syntax + tprg :: prog + +translations + "tprg" == "\ifaces=Ifaces,classes=Classes\" + +constdefs + test :: "(ty)list \ stmt" + "test pTs \ e:==NewC Ext;; + \ Try Expr({ClassT Base,IntVir}!!e\foo({pTs}[Lit Null])) + \ Catch((SXcpt NullPointer) z) + (lab1\ While(Acc (Acc (arr_viewed_from Ext).[Lit (Intg 2)])) Skip)" + + +section "well-structuredness" + +lemma not_Object_subcls_any [elim!]: "(Object, C) \ (subcls1 tprg)^+ \ R" +apply (auto dest!: tranclD subcls1D) +done + +lemma not_Throwable_subcls_SXcpt [elim!]: + "(SXcpt Throwable, SXcpt xn) \ (subcls1 tprg)^+ \ R" +apply (auto dest!: tranclD subcls1D) +apply (simp add: Object_def SXcpt_def) +done + +lemma not_SXcpt_n_subcls_SXcpt_n [elim!]: + "(SXcpt xn, SXcpt xn) \ (subcls1 tprg)^+ \ R" +apply (auto dest!: tranclD subcls1D) +apply (drule rtranclD) +apply auto +done + +lemma not_Base_subcls_Ext [elim!]: "(Base, Ext) \ (subcls1 tprg)^+ \ R" +apply (auto dest!: tranclD subcls1D simp add: BaseCl_def) +done + +lemma not_TName_n_subcls_TName_n [rule_format (no_asm), elim!]: + "(\pid=java_lang,tid=TName tn\, \pid=java_lang,tid=TName tn\) + \ (subcls1 tprg)^+ \ R" +apply (rule_tac n1 = "tn" in surj_tnam_ [THEN exE]) +apply (erule ssubst) +apply (rule tnam_.induct) +apply safe +apply (auto dest!: tranclD subcls1D simp add: BaseCl_def ExtCl_def) +apply (drule rtranclD) +apply auto +done + + +lemma ws_idecl_HasFoo: "ws_idecl tprg HasFoo []" +apply (unfold ws_idecl_def) +apply (simp (no_asm)) +done + +lemma ws_cdecl_Object: "ws_cdecl tprg Object any" +apply (unfold ws_cdecl_def) +apply auto +done + +lemma ws_cdecl_Throwable: "ws_cdecl tprg (SXcpt Throwable) Object" +apply (unfold ws_cdecl_def) +apply auto +done + +lemma ws_cdecl_SXcpt: "ws_cdecl tprg (SXcpt xn) (SXcpt Throwable)" +apply (unfold ws_cdecl_def) +apply auto +done + +lemma ws_cdecl_Base: "ws_cdecl tprg Base Object" +apply (unfold ws_cdecl_def) +apply auto +done + +lemma ws_cdecl_Ext: "ws_cdecl tprg Ext Base" +apply (unfold ws_cdecl_def) +apply auto +done + +lemmas ws_cdecls = ws_cdecl_SXcpt ws_cdecl_Object ws_cdecl_Throwable + ws_cdecl_Base ws_cdecl_Ext + +declare not_Object_subcls_any [rule del] + not_Throwable_subcls_SXcpt [rule del] + not_SXcpt_n_subcls_SXcpt_n [rule del] + not_Base_subcls_Ext [rule del] not_TName_n_subcls_TName_n [rule del] + +lemma ws_idecl_all: + "G=tprg \ (\(I,i)\set Ifaces. ws_idecl G I (isuperIfs i))" +apply (simp (no_asm) add: Ifaces_def HasFooInt_def) +apply (auto intro!: ws_idecl_HasFoo) +done + +lemma ws_cdecl_all: "G=tprg \ (\(C,c)\set Classes. ws_cdecl G C (super c))" +apply (simp (no_asm) add: Classes_def BaseCl_def ExtCl_def) +apply (auto intro!: ws_cdecls simp add: standard_classes_def ObjectC_def + SXcptC_def) +done + +lemma ws_tprg: "ws_prog tprg" +apply (unfold ws_prog_def) +apply (auto intro!: ws_idecl_all ws_cdecl_all) +done + + +section "misc program properties (independent of well-structuredness)" + +lemma single_iface [simp]: "is_iface tprg I = (I = HasFoo)" +apply (unfold Ifaces_def) +apply (simp (no_asm)) +done + +lemma empty_subint1 [simp]: "subint1 tprg = {}" +apply (unfold subint1_def Ifaces_def HasFooInt_def) +apply auto +done + +lemma unique_ifaces: "unique Ifaces" +apply (unfold Ifaces_def) +apply (simp (no_asm)) +done + +lemma unique_classes: "unique Classes" +apply (unfold table_classes_defs ) +apply (simp ) +done + +lemma SXcpt_subcls_Throwable [simp]: "tprg\SXcpt xn\\<^sub>C SXcpt Throwable" +apply (rule SXcpt_subcls_Throwable_lemma) +apply auto +done + +lemma Ext_subclseq_Base [simp]: "tprg\Ext \\<^sub>C Base" +apply (rule subcls_direct1) +apply (simp (no_asm) add: ExtCl_def) +apply (simp add: Object_def) +apply (simp (no_asm)) +done + +lemma Ext_subcls_Base [simp]: "tprg\Ext \\<^sub>C Base" +apply (rule subcls_direct2) +apply (simp (no_asm) add: ExtCl_def) +apply (simp add: Object_def) +apply (simp (no_asm)) +done + + + +section "fields and method lookup" + +lemma fields_tprg_Object [simp]: "DeclConcepts.fields tprg Object = []" +by (rule ws_tprg [THEN fields_emptyI], force+) + +lemma fields_tprg_Throwable [simp]: + "DeclConcepts.fields tprg (SXcpt Throwable) = []" +by (rule ws_tprg [THEN fields_emptyI], force+) + +lemma fields_tprg_SXcpt [simp]: "DeclConcepts.fields tprg (SXcpt xn) = []" +apply (case_tac "xn = Throwable") +apply (simp (no_asm_simp)) +by (rule ws_tprg [THEN fields_emptyI], force+) + +lemmas fields_rec_ = fields_rec [OF _ ws_tprg] + +lemma fields_Base [simp]: +"DeclConcepts.fields tprg Base + = [((arr,Base), \access=Public,static=True ,type=PrimT Boolean.[]\), + ((vee,Base), \access=Public,static=False,type=Iface HasFoo \)]" +apply (subst fields_rec_) +apply (auto simp add: BaseCl_def) +done + +lemma fields_Ext [simp]: + "DeclConcepts.fields tprg Ext + = [((vee,Ext), \access=Public,static=False,type= PrimT Integer\)] + @ DeclConcepts.fields tprg Base" +apply (rule trans) +apply (rule fields_rec_) +apply (auto simp add: ExtCl_def Object_def) +done + +lemmas imethds_rec_ = imethds_rec [OF _ ws_tprg] +lemmas methd_rec_ = methd_rec [OF _ ws_tprg] + +lemma imethds_HasFoo [simp]: + "imethds tprg HasFoo = o2s \ empty(foo_sig\(HasFoo, foo_mhead))" +apply (rule trans) +apply (rule imethds_rec_) +apply (auto simp add: HasFooInt_def) +done + +lemma methd_tprg_Object [simp]: "methd tprg Object = empty" +apply (subst methd_rec_) +apply (auto simp add: Object_mdecls_def) +done + +lemma methd_Base [simp]: + "methd tprg Base = table_of [(\(s,m). (s, Base, m)) Base_foo]" +apply (rule trans) +apply (rule methd_rec_) +apply (auto simp add: BaseCl_def) +done + +(* ### To Table *) +lemma filter_tab_all_False: + "\ k y. t k = Some y \ \ p k y \filter_tab p t = empty" +by (auto simp add: filter_tab_def expand_fun_eq) + + +lemma memberid_Base_foo_simp [simp]: + "memberid (mdecl Base_foo) = mid foo_sig" +by (simp add: Base_foo_def) + +lemma memberid_Ext_foo_simp [simp]: + "memberid (mdecl Ext_foo) = mid foo_sig" +by (simp add: Ext_foo_def) + +lemma Base_declares_foo: + "tprg\mdecl Base_foo declared_in Base" +by (auto simp add: declared_in_def cdeclaredmethd_def BaseCl_def Base_foo_def) + +lemma foo_sig_not_undeclared_in_Base: + "\ tprg\mid foo_sig undeclared_in Base" +proof - + from Base_declares_foo + show ?thesis + by (auto dest!: declared_not_undeclared ) +qed + +lemma Ext_declares_foo: + "tprg\mdecl Ext_foo declared_in Ext" +by (auto simp add: declared_in_def cdeclaredmethd_def ExtCl_def Ext_foo_def) + +lemma foo_sig_not_undeclared_in_Ext: + "\ tprg\mid foo_sig undeclared_in Ext" +proof - + from Ext_declares_foo + show ?thesis + by (auto dest!: declared_not_undeclared ) +qed + +lemma Base_foo_not_inherited_in_Ext: + "\ tprg \ Ext inherits (Base,mdecl Base_foo)" +by (auto simp add: inherits_def foo_sig_not_undeclared_in_Ext) + +lemma Ext_method_inheritance: + "filter_tab (\sig m. tprg \ Ext inherits method sig m) + (empty(fst ((\(s, m). (s, Base, m)) Base_foo)\ + snd ((\(s, m). (s, Base, m)) Base_foo))) + = empty" +proof - + from Base_foo_not_inherited_in_Ext + show ?thesis + by (auto intro: filter_tab_all_False simp add: Base_foo_def) +qed + + +lemma methd_Ext [simp]: "methd tprg Ext = + table_of [(\(s,m). (s, Ext, m)) Ext_foo]" +apply (rule trans) +apply (rule methd_rec_) +apply (auto simp add: ExtCl_def Object_def Ext_method_inheritance) +done + +section "accessibility" + +lemma classesDefined: + "\class tprg C = Some c; C\Object\ \ \ sc. class tprg (super c) = Some sc" +apply (auto simp add: Classes_def standard_classes_def + BaseCl_def ExtCl_def + SXcptC_def ObjectC_def) +done + +lemma superclassesBase [simp]: "superclasses tprg Base={Object}" +proof - + have ws: "ws_prog tprg" by (rule ws_tprg) + then show ?thesis + by (auto simp add: superclasses_rec BaseCl_def) +qed + +lemma superclassesExt [simp]: "superclasses tprg Ext={Base,Object}" +proof - + have ws: "ws_prog tprg" by (rule ws_tprg) + then show ?thesis + by (auto simp add: superclasses_rec ExtCl_def BaseCl_def) +qed + +lemma HasFoo_accessible[simp]:"tprg\(Iface HasFoo) accessible_in P" +by (simp add: accessible_in_RefT_simp is_public_def HasFooInt_def) + +lemma HasFoo_is_acc_iface[simp]: "is_acc_iface tprg P HasFoo" +by (simp add: is_acc_iface_def) + +lemma HasFoo_is_acc_type[simp]: "is_acc_type tprg P (Iface HasFoo)" +by (simp add: is_acc_type_def) + +lemma Base_accessible[simp]:"tprg\(Class Base) accessible_in P" +by (simp add: accessible_in_RefT_simp is_public_def BaseCl_def) + +lemma Base_is_acc_class[simp]: "is_acc_class tprg P Base" +by (simp add: is_acc_class_def) + +lemma Base_is_acc_type[simp]: "is_acc_type tprg P (Class Base)" +by (simp add: is_acc_type_def) + +lemma Ext_accessible[simp]:"tprg\(Class Ext) accessible_in P" +by (simp add: accessible_in_RefT_simp is_public_def ExtCl_def) + +lemma Ext_is_acc_class[simp]: "is_acc_class tprg P Ext" +by (simp add: is_acc_class_def) + +lemma Ext_is_acc_type[simp]: "is_acc_type tprg P (Class Ext)" +by (simp add: is_acc_type_def) + +lemma accmethd_tprg_Object [simp]: "accmethd tprg S Object = empty" +apply (unfold accmethd_def) +apply (simp) +done + +lemma snd_special_simp: "snd ((\(s, m). (s, a, m)) x) = (a,snd x)" +by (cases x) (auto) + +lemma fst_special_simp: "fst ((\(s, m). (s, a, m)) x) = fst x" +by (cases x) (auto) + +lemma foo_sig_undeclared_in_Object: + "tprg\mid foo_sig undeclared_in Object" +by (auto simp add: undeclared_in_def cdeclaredmethd_def Object_mdecls_def) + +(* ### To DeclConcepts *) +lemma undeclared_not_declared: + "G\ memberid m undeclared_in C \ \ G\ m declared_in C" +by (cases m) (auto simp add: declared_in_def undeclared_in_def) + + +lemma unique_sig_Base_foo: + "tprg\ mdecl (sig, snd Base_foo) declared_in Base \ sig=foo_sig" +by (auto simp add: declared_in_def cdeclaredmethd_def + Base_foo_def BaseCl_def) + +lemma Base_foo_no_override: + "tprg,sig\(Base,(snd Base_foo)) overrides old \ P" +apply (drule overrides_commonD) +apply (clarsimp) +apply (frule subclsEval) +apply (rule ws_tprg) +apply (simp) +apply (rule classesDefined) +apply assumption+ +apply (frule unique_sig_Base_foo) +apply (auto dest!: declared_not_undeclared intro: foo_sig_undeclared_in_Object + dest: unique_sig_Base_foo) +done + +lemma Base_foo_no_stat_override: + "tprg,sig\(Base,(snd Base_foo)) overrides\<^sub>S old \ P" +apply (drule stat_overrides_commonD) +apply (clarsimp) +apply (frule subclsEval) +apply (rule ws_tprg) +apply (simp) +apply (rule classesDefined) +apply assumption+ +apply (frule unique_sig_Base_foo) +apply (auto dest!: declared_not_undeclared intro: foo_sig_undeclared_in_Object + dest: unique_sig_Base_foo) +done + + +lemma Base_foo_no_hide: + "tprg,sig\(Base,(snd Base_foo)) hides old \ P" +by (auto dest: hidesD simp add: Base_foo_def member_is_static_simp) + +lemma Ext_foo_no_hide: + "tprg,sig\(Ext,(snd Ext_foo)) hides old \ P" +by (auto dest: hidesD simp add: Ext_foo_def member_is_static_simp) + +lemma unique_sig_Ext_foo: + "tprg\ mdecl (sig, snd Ext_foo) declared_in Ext \ sig=foo_sig" +by (auto simp add: declared_in_def cdeclaredmethd_def + Ext_foo_def ExtCl_def) + +(* ### To DeclConcepts *) +lemma unique_declaration: + "\G\m declared_in C; G\n declared_in C; memberid m = memberid n \ + \ m = n" +apply (cases m) +apply (cases n, + auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)+ +done + + +lemma Ext_foo_override: + "tprg,sig\(Ext,(snd Ext_foo)) overrides old + \ old = (Base,(snd Base_foo))" +apply (drule overrides_commonD) +apply (clarsimp) +apply (frule subclsEval) +apply (rule ws_tprg) +apply (simp) +apply (rule classesDefined) +apply assumption+ +apply (frule unique_sig_Ext_foo) +apply (case_tac "old") +apply (insert Base_declares_foo foo_sig_undeclared_in_Object) +apply (auto simp add: ExtCl_def Ext_foo_def + BaseCl_def Base_foo_def Object_mdecls_def + split_paired_all + member_is_static_simp + dest: declared_not_undeclared unique_declaration) +done + +lemma Ext_foo_stat_override: + "tprg,sig\(Ext,(snd Ext_foo)) overrides\<^sub>S old + \ old = (Base,(snd Base_foo))" +apply (drule stat_overrides_commonD) +apply (clarsimp) +apply (frule subclsEval) +apply (rule ws_tprg) +apply (simp) +apply (rule classesDefined) +apply assumption+ +apply (frule unique_sig_Ext_foo) +apply (case_tac "old") +apply (insert Base_declares_foo foo_sig_undeclared_in_Object) +apply (auto simp add: ExtCl_def Ext_foo_def + BaseCl_def Base_foo_def Object_mdecls_def + split_paired_all + member_is_static_simp + dest: declared_not_undeclared unique_declaration) +done + +(*### weiter hoch *) +lemma Base_foo_member_of_Base: + "tprg\(Base,mdecl Base_foo) member_of Base" +by (auto intro!: members.Immediate Base_declares_foo) + +(*### weiter hoch *) +lemma Ext_foo_member_of_Ext: + "tprg\(Ext,mdecl Ext_foo) member_of Ext" +by (auto intro!: members.Immediate Ext_declares_foo) + +lemma Base_foo_permits_acc: + "tprg \ (Base, mdecl Base_foo) in Base permits_acc_to S" +by ( simp add: permits_acc_def Base_foo_def) + +lemma Base_foo_accessible [simp]: + "tprg\(Base,mdecl Base_foo) of Base accessible_from S" +by (auto intro: accessible_fromR.immediate + Base_foo_member_of_Base Base_foo_permits_acc) + +lemma accmethd_Base [simp]: + "accmethd tprg S Base = methd tprg Base" +apply (simp add: accmethd_def) +apply (rule filter_tab_all_True) +apply (simp add: snd_special_simp fst_special_simp) +done + +lemma Ext_foo_permits_acc: + "tprg \ (Ext, mdecl Ext_foo) in Ext permits_acc_to S" +by ( simp add: permits_acc_def Ext_foo_def) + +lemma Ext_foo_accessible [simp]: + "tprg\(Ext,mdecl Ext_foo) of Ext accessible_from S" +by (auto intro: accessible_fromR.immediate + Ext_foo_member_of_Ext Ext_foo_permits_acc) + +(* +lemma Base_foo_accessible_through_inheritance_in_Ext [simp]: + "tprg\(Base,snd Base_foo) accessible_through_inheritance_in Ext" +apply (rule accessible_through_inheritance.Direct) +apply simp +apply (simp add: accessible_for_inheritance_in_def Base_foo_def) +done +*) + +lemma Ext_foo_overrides_Base_foo: + "tprg\(Ext,Ext_foo) overrides (Base,Base_foo)" +proof (rule overridesR.Direct, simp_all) + show "\ is_static Ext_foo" + by (simp add: member_is_static_simp Ext_foo_def) + show "\ is_static Base_foo" + by (simp add: member_is_static_simp Base_foo_def) + show "accmodi Ext_foo \ Private" + by (simp add: Ext_foo_def) + show "msig (Ext, Ext_foo) = msig (Base, Base_foo)" + by (simp add: Ext_foo_def Base_foo_def) + show "tprg\mdecl Ext_foo declared_in Ext" + by (auto intro: Ext_declares_foo) + show "tprg\mdecl Base_foo declared_in Base" + by (auto intro: Base_declares_foo) + show "tprg \(Base, mdecl Base_foo) inheritable_in java_lang" + by (simp add: inheritable_in_def Base_foo_def) + show "tprg\resTy Ext_foo\resTy Base_foo" + by (simp add: Ext_foo_def Base_foo_def mhead_resTy_simp) +qed + +(* +lemma Base_foo_of_Ext_accessible[simp]: + "tprg\(Base, mdecl Base_foo) of Ext accessible_from S" +apply (auto intro: accessible_fromR.immediate + Base_foo_member_of_Base Base_foo_permits_acc) +apply (rule accessible_fromR.immediate) +apply (rule_tac "old"="(Base,Base_foo)" and sup="Base" + in accessible_fromR.overriding) +apply (auto intro!: Ext_foo_overrides_Base_foo) +apply (auto +apply (insert Ext_foo_overrides_Base_foo) +apply (rule accessible_fromR.overriding, simp_all) +apply (auto intro!: Ext_foo_overrides_Base_foo) +apply (auto intro!: accessible_fromR.overriding + intro: Ext_foo_overrides_Base_foo) +by + Ext_foo_member_of_Ext Ext_foo_permits_acc) +apply (auto intro!: accessible +apply (auto simp add: method_accessible_from_def accessible_from_def) +apply (simp add: Base_foo_def) +done +*) + +lemma accmethd_Ext [simp]: + "accmethd tprg S Ext = methd tprg Ext" +apply (simp add: accmethd_def) +apply (rule filter_tab_all_True) +apply (auto simp add: snd_special_simp fst_special_simp) +done + +(* ### Weiter hoch *) +lemma cls_Ext: "class tprg Ext = Some ExtCl" +by simp +lemma dynmethd_Ext_foo: + "dynmethd tprg Base Ext \name = foo, parTs = [Class Base]\ + = Some (Ext,snd Ext_foo)" +proof - + have "methd tprg Base \name = foo, parTs = [Class Base]\ + = Some (Base,snd Base_foo)" and + "methd tprg Ext \name = foo, parTs = [Class Base]\ + = Some (Ext,snd Ext_foo)" + by (auto simp add: Ext_foo_def Base_foo_def foo_sig_def) + with cls_Ext ws_tprg Ext_foo_overrides_Base_foo + show ?thesis + by (auto simp add: dynmethd_rec simp add: Ext_foo_def Base_foo_def) +qed + +lemma Base_fields_accessible[simp]: + "accfield tprg S Base + = table_of((map (\((n,d),f).(n,(d,f)))) (DeclConcepts.fields tprg Base))" +apply (auto simp add: accfield_def expand_fun_eq Let_def + accessible_in_RefT_simp + is_public_def + BaseCl_def + permits_acc_def + declared_in_def + cdeclaredfield_def + intro!: filter_tab_all_True_Some filter_tab_None + accessible_fromR.immediate + intro: members.Immediate) +done + + +lemma arr_member_of_Base: + "tprg\(Base, fdecl (arr, + \access = Public, static = True, type = PrimT Boolean.[]\)) + member_of Base" +by (auto intro: members.Immediate + simp add: declared_in_def cdeclaredfield_def BaseCl_def) + +lemma arr_member_of_Ext: + "tprg\(Base, fdecl (arr, + \access = Public, static = True, type = PrimT Boolean.[]\)) + member_of Ext" +apply (rule members.Inherited) +apply (simp add: inheritable_in_def) +apply (simp add: undeclared_in_def cdeclaredfield_def ExtCl_def) +apply (auto intro: arr_member_of_Base simp add: subcls1_def ExtCl_def) +done + +lemma Ext_fields_accessible[simp]: +"accfield tprg S Ext + = table_of((map (\((n,d),f).(n,(d,f)))) (DeclConcepts.fields tprg Ext))" +apply (auto simp add: accfield_def expand_fun_eq Let_def + accessible_in_RefT_simp + is_public_def + BaseCl_def + ExtCl_def + permits_acc_def + intro!: filter_tab_all_True_Some filter_tab_None + accessible_fromR.immediate) +apply (auto intro: members.Immediate arr_member_of_Ext + simp add: declared_in_def cdeclaredfield_def ExtCl_def) +done + +lemma array_of_PrimT_acc [simp]: + "is_acc_type tprg java_lang (PrimT t.[])" +apply (simp add: is_acc_type_def accessible_in_RefT_simp) +done + +lemma PrimT_acc [simp]: + "is_acc_type tprg java_lang (PrimT t)" +apply (simp add: is_acc_type_def accessible_in_RefT_simp) +done + +lemma Object_acc [simp]: + "is_acc_class tprg java_lang Object" +apply (auto simp add: is_acc_class_def accessible_in_RefT_simp is_public_def) +done + + +section "well-formedness" + + +lemma wf_HasFoo: "wf_idecl tprg (HasFoo, HasFooInt)" +apply (unfold wf_idecl_def HasFooInt_def) +apply (auto intro!: wf_mheadI ws_idecl_HasFoo + simp add: foo_sig_def foo_mhead_def mhead_resTy_simp + member_is_static_simp ) +done + +declare wt.Skip [rule del] wt.Init [rule del] +lemmas Base_foo_defs = Base_foo_def foo_sig_def foo_mhead_def +lemmas Ext_foo_defs = Ext_foo_def foo_sig_def +ML {* bind_thms ("wt_intros",map (rewrite_rule [id_def]) (thms "wt.intros")) *} +lemmas wtIs = wt_Call wt_Super wt_FVar wt_StatRef wt_intros + +lemma wf_Base_foo: "wf_mdecl tprg Base Base_foo" +apply (unfold Base_foo_defs ) +apply (auto intro!: wf_mdeclI wf_mheadI intro!: wtIs + simp add: mhead_resTy_simp) +done + +lemma wf_Ext_foo: "wf_mdecl tprg Ext Ext_foo" +apply (unfold Ext_foo_defs ) +apply (auto intro!: wf_mdeclI wf_mheadI intro!: wtIs + simp add: mhead_resTy_simp ) +apply (rule wt.Cast) +prefer 2 +apply simp +apply (rule_tac [2] narrow.subcls [THEN cast.narrow]) +apply (auto intro!: wtIs) +done + +declare mhead_resTy_simp [simp add] +declare member_is_static_simp [simp add] + +lemma wf_BaseC: "wf_cdecl tprg (Base,BaseCl)" +apply (unfold wf_cdecl_def BaseCl_def arr_viewed_from_def) +apply (auto intro!: wf_Base_foo) +apply (auto intro!: ws_cdecl_Base simp add: Base_foo_def foo_mhead_def) +apply (auto intro!: wtIs) +apply (auto simp add: Base_foo_defs entails_def Let_def) +apply (insert Base_foo_no_stat_override, simp add: Base_foo_def,blast)+ +apply (insert Base_foo_no_hide , simp add: Base_foo_def,blast) +done + + +lemma wf_ExtC: "wf_cdecl tprg (Ext,ExtCl)" +apply (unfold wf_cdecl_def ExtCl_def) +apply (auto intro!: wf_Ext_foo ws_cdecl_Ext) +apply (auto simp add: entails_def snd_special_simp) +apply (insert Ext_foo_stat_override) +apply (force simp add: qmdecl_def Ext_foo_def Base_foo_def) +apply (force simp add: qmdecl_def Ext_foo_def Base_foo_def) +apply (force simp add: qmdecl_def Ext_foo_def Base_foo_def) +apply (insert Ext_foo_no_hide) +apply (simp_all add: qmdecl_def) +apply blast+ +done + +lemma wf_idecl_all: "p=tprg \ Ball (set Ifaces) (wf_idecl p)" +apply (simp (no_asm) add: Ifaces_def) +apply (simp (no_asm_simp)) +apply (rule wf_HasFoo) +done + +lemma wf_cdecl_all_standard_classes: + "Ball (set standard_classes) (wf_cdecl tprg)" +apply (unfold standard_classes_def Let_def + ObjectC_def SXcptC_def Object_mdecls_def SXcpt_mdecls_def) +apply (simp (no_asm) add: wf_cdecl_def ws_cdecls) +apply (auto simp add:is_acc_class_def accessible_in_RefT_simp SXcpt_def) +apply (auto simp add: Object_def Classes_def standard_classes_def + SXcptC_def SXcpt_def) +done + +lemma wf_cdecl_all: "p=tprg \ Ball (set Classes) (wf_cdecl p)" +apply (simp (no_asm) add: Classes_def) +apply (simp (no_asm_simp)) +apply (rule wf_BaseC [THEN conjI]) +apply (rule wf_ExtC [THEN conjI]) +apply (rule wf_cdecl_all_standard_classes) +done + +theorem wf_tprg: "wf_prog tprg" +apply (unfold wf_prog_def Let_def) +apply (simp (no_asm) add: unique_ifaces unique_classes) +apply (rule conjI) +apply ((simp (no_asm) add: Classes_def standard_classes_def)) +apply (rule conjI) +apply (simp add: Object_mdecls_def) +apply safe +apply (cut_tac xn_cases_old) (* FIXME (insert xn_cases) *) +apply (simp (no_asm_simp) add: Classes_def standard_classes_def) +apply (insert wf_idecl_all) +apply (insert wf_cdecl_all) +apply auto +done + + +section "max spec" + +lemma appl_methds_Base_foo: +"appl_methds tprg S (ClassT Base) \name=foo, parTs=[NT]\ = + {((ClassT Base, \access=Public,static=False,pars=[z],resT=Class Base\) + ,[Class Base])}" +apply (unfold appl_methds_def) +apply (simp (no_asm)) +apply (subgoal_tac "tprg\NT\ Class Base") +apply (auto simp add: cmheads_def Base_foo_defs) +done + +lemma max_spec_Base_foo: "max_spec tprg S (ClassT Base) \name=foo,parTs=[NT]\ = + {((ClassT Base, \access=Public,static=False,pars=[z],resT=Class Base\) + , [Class Base])}" +apply (unfold max_spec_def) +apply (simp (no_asm) add: appl_methds_Base_foo) +apply auto +done + + +section "well-typedness" + +lemma wt_test: "\prg=tprg,cls=S,lcl=empty(VName e\Class Base)\\test ?pTs\\" +apply (unfold test_def arr_viewed_from_def) +(* ?pTs = [Class Base] *) +apply (rule wtIs (* ;; *)) +apply (rule wtIs (* Ass *)) +apply (rule wtIs (* NewC *)) +apply (rule wtIs (* LVar *)) +apply (simp) +apply (simp) +apply (simp) +apply (rule wtIs (* NewC *)) +apply (simp) +apply (simp) +apply (rule wtIs (* Try *)) +prefer 4 +apply (simp) +defer +apply (rule wtIs (* Expr *)) +apply (rule wtIs (* Call *)) +apply (rule wtIs (* Acc *)) +apply (rule wtIs (* LVar *)) +apply (simp) +apply (simp) +apply (rule wtIs (* Cons *)) +apply (rule wtIs (* Lit *)) +apply (simp) +apply (rule wtIs (* Nil *)) +apply (simp) +apply (rule max_spec_Base_foo) +apply (simp) +apply (simp) +apply (simp) +apply (simp) +apply (rule wtIs (* While *)) +apply (rule wtIs (* Acc *)) +apply (rule wtIs (* AVar *)) +apply (rule wtIs (* Acc *)) +apply (rule wtIs (* FVar *)) +apply (rule wtIs (* StatRef *)) +apply (simp) +apply (simp) +apply (simp ) +apply (simp) +apply (rule wtIs (* LVar *)) +apply (simp) +apply (rule wtIs (* Skip *)) +done + + +section "execution" + +lemma alloc_one: "\a obj. \the (new_Addr h) = a; atleast_free h (Suc n)\ \ + new_Addr h = Some a \ atleast_free (h(a\obj)) n" +apply (frule atleast_free_SucD) +apply (drule atleast_free_Suc [THEN iffD1]) +apply clarsimp +apply (frule new_Addr_SomeI) +apply force +done + +declare fvar_def2 [simp] avar_def2 [simp] init_lvars_def2 [simp] +declare init_obj_def [simp] var_tys_def [simp] fields_table_def [simp] +declare BaseCl_def [simp] ExtCl_def [simp] Ext_foo_def [simp] + Base_foo_defs [simp] + +ML {* bind_thms ("eval_intros", map + (simplify (simpset() delsimps [thm "Skip_eq"] + addsimps [thm "lvar_def"]) o + rewrite_rule [thm "assign_def",Let_def]) (thms "eval.intros")) *} +lemmas eval_Is = eval_Init eval_StatRef AbruptIs eval_intros + +consts + a :: loc + b :: loc + c :: loc + +syntax + + tprg :: prog + + obj_a :: obj + obj_b :: obj + obj_c :: obj + arr_N :: "(vn, val) table" + arr_a :: "(vn, val) table" + globs1 :: globs + globs2 :: globs + globs3 :: globs + globs8 :: globs + locs3 :: locals + locs4 :: locals + locs8 :: locals + s0 :: state + s0' :: state + s9' :: state + s1 :: state + s1' :: state + s2 :: state + s2' :: state + s3 :: state + s3' :: state + s4 :: state + s4' :: state + s6' :: state + s7' :: state + s8 :: state + s8' :: state + +translations + + "tprg" == "\ifaces=Ifaces,classes=Classes\" + + "obj_a" <= "\tag=Arr (PrimT Boolean) two + ,values=empty(Inr 0\Bool False)(Inr one\Bool False)\" + "obj_b" <= "\tag=CInst Ext + ,values=(empty(Inl (vee, Base)\Null ) + (Inl (vee, Ext )\Intg 0))\" + "obj_c" == "\tag=CInst (SXcpt NullPointer),values=empty\" + "arr_N" == "empty(Inl (arr, Base)\Null)" + "arr_a" == "empty(Inl (arr, Base)\Addr a)" + "globs1" == "empty(Inr Ext \\tag=arbitrary, values=empty\) + (Inr Base \\tag=arbitrary, values=arr_N\) + (Inr Object\\tag=arbitrary, values=empty\)" + "globs2" == "empty(Inr Ext \\tag=arbitrary, values=empty\) + (Inr Object\\tag=arbitrary, values=empty\) + (Inl a\obj_a) + (Inr Base \\tag=arbitrary, values=arr_a\)" + "globs3" == "globs2(Inl b\obj_b)" + "globs8" == "globs3(Inl c\obj_c)" + "locs3" == "empty(VName e\Addr b)" + "locs4" == "empty(VName z\Null)(Inr()\Addr b)" + "locs8" == "locs3(VName z\Addr c)" + "s0" == " st empty empty" + "s0'" == " Norm s0" + "s1" == " st globs1 empty" + "s1'" == " Norm s1" + "s2" == " st globs2 empty" + "s2'" == " Norm s2" + "s3" == " st globs3 locs3 " + "s3'" == " Norm s3" + "s4" == " st globs3 locs4" + "s4'" == " Norm s4" + "s6'" == "(Some (Xcpt (Std NullPointer)), s4)" + "s7'" == "(Some (Xcpt (Std NullPointer)), s3)" + "s8" == " st globs8 locs8" + "s8'" == " Norm s8" + "s9'" == "(Some (Xcpt (Std IndOutBound)), s8)" + + +syntax "four"::nat + "tree"::nat + "two" ::nat + "one" ::nat +translations + "one" == "Suc 0" + "two" == "Suc one" + "tree" == "Suc two" + "four" == "Suc tree" + +declare Pair_eq [simp del] +lemma exec_test: +"\the (new_Addr (heap s1)) = a; + the (new_Addr (heap ?s2)) = b; + the (new_Addr (heap ?s3)) = c\ \ + atleast_free (heap s0) four \ + tprg\s0' \test [Class Base]\ ?s9'" +apply (unfold test_def arr_viewed_from_def) +(* ?s9' = s9' *) +apply (simp (no_asm_use)) +apply (drule (1) alloc_one,clarsimp) +apply (rule eval_Is (* ;; *)) +apply (erule_tac V = "the (new_Addr ?h) = c" in thin_rl) +apply (erule_tac [2] V = "new_Addr ?h = Some a" in thin_rl) +apply (erule_tac [2] V = "atleast_free ?h four" in thin_rl) +apply (rule eval_Is (* Expr *)) +apply (rule eval_Is (* Ass *)) +apply (rule eval_Is (* LVar *)) +apply (rule eval_Is (* NewC *)) + (* begin init Ext *) +apply (erule_tac V = "the (new_Addr ?h) = b" in thin_rl) +apply (erule_tac V = "atleast_free ?h tree" in thin_rl) +apply (erule_tac [2] V = "atleast_free ?h four" in thin_rl) +apply (erule_tac [2] V = "new_Addr ?h = Some a" in thin_rl) +apply (rule eval_Is (* Init Ext *)) +apply (simp) +apply (rule conjI) +prefer 2 apply (rule conjI HOL.refl)+ +apply (rule eval_Is (* Init Base *)) +apply (simp add: arr_viewed_from_def) +apply (rule conjI) +apply (rule eval_Is (* Init Object *)) +apply (simp) +apply (rule conjI, rule HOL.refl)+ +apply (rule HOL.refl) +apply (simp) +apply (rule conjI, rule_tac [2] HOL.refl) +apply (rule eval_Is (* Expr *)) +apply (rule eval_Is (* Ass *)) +apply (rule eval_Is (* FVar *)) +apply (rule init_done, simp) +apply (rule eval_Is (* StatRef *)) +apply (simp) +apply (rule eval_Is (* NewA *)) +apply (simp) +apply (rule eval_Is (* Lit *)) +apply (simp) +apply (rule halloc.New) +apply (simp (no_asm_simp)) +apply (drule atleast_free_weaken,rotate_tac -1,drule atleast_free_weaken) +apply (simp (no_asm_simp)) +apply (simp add: upd_gobj_def) + (* end init Ext *) +apply (rule halloc.New) +apply (drule alloc_one) +prefer 2 apply fast +apply (simp (no_asm_simp)) +apply (drule atleast_free_weaken) +apply force +apply (simp) +apply (drule alloc_one) +apply (simp (no_asm_simp)) +apply clarsimp +apply (erule_tac V = "atleast_free ?h tree" in thin_rl) +apply (drule_tac x = "a" in new_AddrD2 [THEN spec]) +apply (simp (no_asm_use)) +apply (rule eval_Is (* Try *)) +apply (rule eval_Is (* Expr *)) + (* begin method call *) +apply (rule eval_Is (* Call *)) +apply (rule eval_Is (* Acc *)) +apply (rule eval_Is (* LVar *)) +apply (rule eval_Is (* Cons *)) +apply (rule eval_Is (* Lit *)) +apply (rule eval_Is (* Nil *)) +apply (simp) +apply (simp) +apply (rule eval_Is (* Methd *)) +apply (simp add: body_def Let_def) +apply (rule eval_Is (* Body *)) +apply (rule init_done, simp) +apply (simp add: invocation_declclass_def dynlookup_def dynmethd_Ext_foo) +apply (simp add: invocation_declclass_def dynlookup_def dynmethd_Ext_foo) +apply (rule eval_Is (* Expr *)) +apply (rule eval_Is (* Ass *)) +apply (rule eval_Is (* FVar *)) +apply (rule init_done, simp) +apply (rule eval_Is (* Cast *)) +apply (rule eval_Is (* Acc *)) +apply (rule eval_Is (* LVar *)) +apply (simp) +apply (simp split del: split_if) +apply (rule eval_Is (* XcptE *)) +apply (simp) + (* end method call *) +apply (rule sxalloc.intros) +apply (rule halloc.New) +apply (erule alloc_one [THEN conjunct1]) +apply (simp (no_asm_simp)) +apply (simp (no_asm_simp)) +apply (simp add: gupd_def lupd_def obj_ty_def split del: split_if) +apply (drule alloc_one [THEN conjunct1]) +apply (simp (no_asm_simp)) +apply (erule_tac V = "atleast_free ?h two" in thin_rl) +apply (drule_tac x = "a" in new_AddrD2 [THEN spec]) +apply simp +apply (rule eval_Is (* While *)) +apply (rule eval_Is (* Acc *)) +apply (rule eval_Is (* AVar *)) +apply (rule eval_Is (* Acc *)) +apply (rule eval_Is (* FVar *)) +apply (rule init_done, simp) +apply (rule eval_Is (* StatRef *)) +apply (simp) +apply (rule eval_Is (* Lit *)) +apply (simp (no_asm_simp)) +apply (auto simp add: in_bounds_def) +done +declare Pair_eq [simp] + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Name.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Name.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,123 @@ +(* Title: isabelle/Bali/Name.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* Java names *} + +theory Name = Basis: + +(* cf. 6.5 *) +typedecl tnam (* ordinary type name, i.e. class or interface name *) +typedecl pname (* package name *) +typedecl mname (* method name *) +typedecl vname (* variable or field name *) +typedecl label (* label as destination of break or continue *) + +arities + tnam :: "type" + pname :: "type" + vname :: "type" + mname :: "type" + label :: "type" + + +datatype ename (* expression name *) + = VNam vname + | Res (* special name to model the return value of methods *) + +datatype lname (* names for local variables and the This pointer *) + = EName ename + | This +syntax + VName :: "vname \ lname" + Result :: lname + +translations + "VName n" == "EName (VNam n)" + "Result" == "EName Res" + +datatype xname (* names of standard exceptions *) + = Throwable + | NullPointer | OutOfMemory | ClassCast + | NegArrSize | IndOutBound | ArrStore + +lemma xn_cases: + "xn = Throwable \ xn = NullPointer \ + xn = OutOfMemory \ xn = ClassCast \ + xn = NegArrSize \ xn = IndOutBound \ xn = ArrStore" +apply (induct xn) +apply auto +done + +lemma xn_cases_old: (* FIXME cf. Example.thy *) + "ALL xn. xn = Throwable \ xn = NullPointer \ + xn = OutOfMemory \ xn = ClassCast \ + xn = NegArrSize \ xn = IndOutBound \ xn = ArrStore" +apply (rule allI) +apply (induct_tac xn) +apply auto +done + +datatype tname (* type names for standard classes and other type names *) + = Object_ + | SXcpt_ xname + | TName tnam + +record qtname = (* qualified tname cf. 6.5.3, 6.5.4*) + pid :: pname + tid :: tname + +axclass has_pname < "type" +consts pname::"'a::has_pname \ pname" + +instance pname::has_pname +by (intro_classes) + +defs (overloaded) +pname_pname_def: "pname (p::pname) \ p" + +axclass has_tname < "type" +consts tname::"'a::has_tname \ tname" + +instance tname::has_tname +by (intro_classes) + +defs (overloaded) +tname_tname_def: "tname (t::tname) \ t" + +axclass has_qtname < "type" +consts qtname:: "'a::has_qtname \ qtname" + +(* Declare qtname as instance of has_qtname *) +instance pid_field_type::(has_pname,"type") has_qtname +by intro_classes + +defs (overloaded) +qtname_qtname_def: "qtname (q::qtname) \ q" + +translations + "mname" <= "Name.mname" + "xname" <= "Name.xname" + "tname" <= "Name.tname" + "ename" <= "Name.ename" + "qtname" <= (type) "\pid::pname,tid::tname\" + (type) "'a qtname_scheme" <= (type) "\pid::pname,tid::tname,\::'a\" + + +consts java_lang::pname (* package java.lang *) + +consts + Object :: qtname + SXcpt :: "xname \ qtname" +defs + Object_def: "Object \ \pid = java_lang, tid = Object_\" + SXcpt_def: "SXcpt \ \x. \pid = java_lang, tid = SXcpt_ x\" + +lemma Object_neq_SXcpt [simp]: "Object \ SXcpt xn" +by (simp add: Object_def SXcpt_def) + +lemma SXcpt_inject [simp]: "(SXcpt xn = SXcpt xm) = (xn = xm)" +by (simp add: SXcpt_def) +end + diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/ROOT.ML Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,13 @@ + +use_thy "WellForm"; + +(*The dynamic part of Bali, including type-safety*) +use_thy "Evaln"; +use_thy "Example"; +use_thy "TypeSafe"; +(*###use_thy "Trans";*) + +(*The Hoare logic for Bali*) +use_thy "AxExample"; +use_thy "AxSound"; +use_thy "AxCompl"; diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/State.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/State.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,739 @@ +(* Title: isabelle/Bali/State.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* State for evaluation of Java expressions and statements *} + +theory State = DeclConcepts: + +text {* +design issues: +\begin{itemize} +\item all kinds of objects (class instances, arrays, and class objects) + are handeled via a general object abstraction +\item the heap and the map for class objects are combined into a single table + @{text "(recall (loc, obj) table \ (qtname, obj) table ~= (loc + qtname, obj) table)"} +\end{itemize} +*} + +section "objects" + +datatype obj_tag = (* tag for generic object *) + CInst qtname (* class instance *) + | Arr ty int (* array with component type and length *) + (* | CStat the tag is irrelevant for a class object, + i.e. the static fields of a class *) + +types vn = "fspec + int" (* variable name *) +record obj = + tag :: "obj_tag" (* generalized object *) + values :: "(vn, val) table" + +translations + "fspec" <= (type) "vname \ qtname" + "vn" <= (type) "fspec + int" + "obj" <= (type) "\tag::obj_tag, values::vn \ val option\" + "obj" <= (type) "\tag::obj_tag, values::vn \ val option,\::'a\" + +constdefs + + the_Arr :: "obj option \ ty \ int \ (vn, val) table" + "the_Arr obj \ \(T,k,t). obj = Some \tag=Arr T k,values=t\" + +lemma the_Arr_Arr [simp]: "the_Arr (Some \tag=Arr T k,values=cs\) = (T,k,cs)" +apply (auto simp: the_Arr_def) +done + +lemma the_Arr_Arr1 [simp,intro,dest]: + "\tag obj = Arr T k\ \ the_Arr (Some obj) = (T,k,values obj)" +apply (auto simp add: the_Arr_def) +done + +constdefs + + upd_obj :: "vn \ val \ obj \ obj" + "upd_obj n v \ \ obj . obj \values:=(values obj)(n\v)\" + +lemma upd_obj_def2 [simp]: + "upd_obj n v obj = obj \values:=(values obj)(n\v)\" +apply (auto simp: upd_obj_def) +done + +constdefs + obj_ty :: "obj \ ty" + "obj_ty obj \ case tag obj of + CInst C \ Class C + | Arr T k \ T.[]" + +lemma obj_ty_eq [intro!]: "obj_ty \tag=oi,values=x\ = obj_ty \tag=oi,values=y\" +by (simp add: obj_ty_def) + + +lemma obj_ty_eq1 [intro!,dest]: + "tag obj = tag obj' \ obj_ty obj = obj_ty obj'" +by (simp add: obj_ty_def) + +lemma obj_ty_cong [simp]: + "obj_ty (obj \values:=vs\) = obj_ty obj" +by auto +(* +lemma obj_ty_cong [simp]: + "obj_ty (obj \values:=vs(n\v)\) = obj_ty (obj \values:=vs\)" +by auto +*) +lemma obj_ty_CInst [simp]: + "obj_ty \tag=CInst C,values=vs\ = Class C" +by (simp add: obj_ty_def) + +lemma obj_ty_CInst1 [simp,intro!,dest]: + "\tag obj = CInst C\ \ obj_ty obj = Class C" +by (simp add: obj_ty_def) + +lemma obj_ty_Arr [simp]: + "obj_ty \tag=Arr T i,values=vs\ = T.[]" +by (simp add: obj_ty_def) + +lemma obj_ty_Arr1 [simp,intro!,dest]: + "\tag obj = Arr T i\ \ obj_ty obj = T.[]" +by (simp add: obj_ty_def) + +lemma obj_ty_widenD: + "G\obj_ty obj\RefT t \ (\C. tag obj = CInst C) \ (\T k. tag obj = Arr T k)" +apply (unfold obj_ty_def) +apply (auto split add: obj_tag.split_asm) +done + +constdefs + + obj_class :: "obj \ qtname" + "obj_class obj \ case tag obj of + CInst C \ C + | Arr T k \ Object" + + +lemma obj_class_CInst [simp]: "obj_class \tag=CInst C,values=vs\ = C" +by (auto simp: obj_class_def) + +lemma obj_class_CInst1 [simp,intro!,dest]: + "tag obj = CInst C \ obj_class obj = C" +by (auto simp: obj_class_def) + +lemma obj_class_Arr [simp]: "obj_class \tag=Arr T k,values=vs\ = Object" +by (auto simp: obj_class_def) + +lemma obj_class_Arr1 [simp,intro!,dest]: + "tag obj = Arr T k \ obj_class obj = Object" +by (auto simp: obj_class_def) + +lemma obj_ty_obj_class: "G\obj_ty obj\ Class statC = G\obj_class obj \\<^sub>C statC" +apply (case_tac "tag obj") +apply (auto simp add: obj_ty_def obj_class_def) +apply (case_tac "statC = Object") +apply (auto dest: widen_Array_Class) +done + +section "object references" + +types oref = "loc + qtname" (* generalized object reference *) +syntax + Heap :: "loc \ oref" + Stat :: "qtname \ oref" + +translations + "Heap" => "Inl" + "Stat" => "Inr" + "oref" <= (type) "loc + qtname" + +constdefs + fields_table:: + "prog \ qtname \ (fspec \ field \ bool) \ (fspec, ty) table" + "fields_table G C P + \ option_map type \ table_of (filter (split P) (DeclConcepts.fields G C))" + +lemma fields_table_SomeI: +"\table_of (DeclConcepts.fields G C) n = Some f; P n f\ + \ fields_table G C P n = Some (type f)" +apply (unfold fields_table_def) +apply clarsimp +apply (rule exI) +apply (rule conjI) +apply (erule map_of_filter_in) +apply assumption +apply simp +done + +(* unused *) +lemma fields_table_SomeD': "fields_table G C P fn = Some T \ + \f. (fn,f)\set(DeclConcepts.fields G C) \ type f = T" +apply (unfold fields_table_def) +apply clarsimp +apply (drule map_of_SomeD) +apply auto +done + +lemma fields_table_SomeD: +"\fields_table G C P fn = Some T; unique (DeclConcepts.fields G C)\ \ + \f. table_of (DeclConcepts.fields G C) fn = Some f \ type f = T" +apply (unfold fields_table_def) +apply clarsimp +apply (rule exI) +apply (rule conjI) +apply (erule table_of_filter_unique_SomeD) +apply assumption +apply simp +done + +constdefs + in_bounds :: "int \ int \ bool" ("(_/ in'_bounds _)" [50, 51] 50) + "i in_bounds k \ 0 \ i \ i < k" + + arr_comps :: "'a \ int \ int \ 'a option" + "arr_comps T k \ \i. if i in_bounds k then Some T else None" + + var_tys :: "prog \ obj_tag \ oref \ (vn, ty) table" +"var_tys G oi r + \ case r of + Heap a \ (case oi of + CInst C \ fields_table G C (\n f. \static f) (+) empty + | Arr T k \ empty (+) arr_comps T k) + | Stat C \ fields_table G C (\fn f. declclassf fn = C \ static f) + (+) empty" + +lemma var_tys_Some_eq: + "var_tys G oi r n = Some T + = (case r of + Inl a \ (case oi of + CInst C \ (\nt. n = Inl nt \ fields_table G C (\n f. + \static f) nt = Some T) + | Arr t k \ (\ i. n = Inr i \ i in_bounds k \ t = T)) + | Inr C \ (\nt. n = Inl nt \ + fields_table G C (\fn f. declclassf fn = C \ static f) nt + = Some T))" +apply (unfold var_tys_def arr_comps_def) +apply (force split add: sum.split_asm sum.split obj_tag.split) +done + + +section "stores" + +types globs (* global variables: heap and static variables *) + = "(oref , obj) table" + heap + = "(loc , obj) table" + locals + = "(lname, val) table" (* local variables *) + +translations + "globs" <= (type) "(oref , obj) table" + "heap" <= (type) "(loc , obj) table" + "locals" <= (type) "(lname, val) table" + +datatype st = (* pure state, i.e. contents of all variables *) + st globs locals + +subsection "access" + +constdefs + + globs :: "st \ globs" + "globs \ st_case (\g l. g)" + + locals :: "st \ locals" + "locals \ st_case (\g l. l)" + + heap :: "st \ heap" + "heap s \ globs s \ Heap" + + +lemma globs_def2 [simp]: " globs (st g l) = g" +by (simp add: globs_def) + +lemma locals_def2 [simp]: "locals (st g l) = l" +by (simp add: locals_def) + +lemma heap_def2 [simp]: "heap s a=globs s (Heap a)" +by (simp add: heap_def) + + +syntax + val_this :: "st \ val" + lookup_obj :: "st \ val \ obj" + +translations + "val_this s" == "the (locals s This)" + "lookup_obj s a'" == "the (heap s (the_Addr a'))" + +subsection "memory allocation" + +constdefs + new_Addr :: "heap \ loc option" + "new_Addr h \ if (\a. h a \ None) then None else Some (\a. h a = None)" + +lemma new_AddrD: "new_Addr h = Some a \ h a = None" +apply (unfold new_Addr_def) +apply auto +apply (case_tac "h (SOME a\loc. h a = None)") +apply simp +apply (fast intro: someI2) +done + +lemma new_AddrD2: "new_Addr h = Some a \ \b. h b \ None \ b \ a" +apply (drule new_AddrD) +apply auto +done + +lemma new_Addr_SomeI: "h a = None \ \b. new_Addr h = Some b \ h b = None" +apply (unfold new_Addr_def) +apply (frule not_Some_eq [THEN iffD2]) +apply auto +apply (drule not_Some_eq [THEN iffD2]) +apply auto +apply (fast intro!: someI2) +done + + +subsection "initialization" + +syntax + + init_vals :: "('a, ty) table \ ('a, val) table" + +translations + "init_vals vs" == "option_map default_val \ vs" + +lemma init_arr_comps_base [simp]: "init_vals (arr_comps T 0) = empty" +apply (unfold arr_comps_def in_bounds_def) +apply (rule ext) +apply auto +done + +lemma init_arr_comps_step [simp]: +"0 < j \ init_vals (arr_comps T j ) = + init_vals (arr_comps T (j - 1))(j - 1\default_val T)" +apply (unfold arr_comps_def in_bounds_def) +apply (rule ext) +apply auto +done + +subsection "update" + +constdefs + gupd :: "oref \ obj \ st \ st" ("gupd'(_\_')"[10,10]1000) + "gupd r obj \ st_case (\g l. st (g(r\obj)) l)" + + lupd :: "lname \ val \ st \ st" ("lupd'(_\_')"[10,10]1000) + "lupd vn v \ st_case (\g l. st g (l(vn\v)))" + + upd_gobj :: "oref \ vn \ val \ st \ st" + "upd_gobj r n v \ st_case (\g l. st (chg_map (upd_obj n v) r g) l)" + + set_locals :: "locals \ st \ st" + "set_locals l \ st_case (\g l'. st g l)" + + init_obj :: "prog \ obj_tag \ oref \ st \ st" + "init_obj G oi r \ gupd(r\\tag=oi, values=init_vals (var_tys G oi r)\)" + +syntax + init_class_obj :: "prog \ qtname \ st \ st" + +translations + "init_class_obj G C" == "init_obj G arbitrary (Inr C)" + +lemma gupd_def2 [simp]: "gupd(r\obj) (st g l) = st (g(r\obj)) l" +apply (unfold gupd_def) +apply (simp (no_asm)) +done + +lemma lupd_def2 [simp]: "lupd(vn\v) (st g l) = st g (l(vn\v))" +apply (unfold lupd_def) +apply (simp (no_asm)) +done + +lemma globs_gupd [simp]: "globs (gupd(r\obj) s) = globs s(r\obj)" +apply (induct "s") +by (simp add: gupd_def) + +lemma globs_lupd [simp]: "globs (lupd(vn\v ) s) = globs s" +apply (induct "s") +by (simp add: lupd_def) + +lemma locals_gupd [simp]: "locals (gupd(r\obj) s) = locals s" +apply (induct "s") +by (simp add: gupd_def) + +lemma locals_lupd [simp]: "locals (lupd(vn\v ) s) = locals s(vn\v )" +apply (induct "s") +by (simp add: lupd_def) + +lemma globs_upd_gobj_new [rule_format (no_asm), simp]: + "globs s r = None \ globs (upd_gobj r n v s) = globs s" +apply (unfold upd_gobj_def) +apply (induct "s") +apply auto +done + +lemma globs_upd_gobj_upd [rule_format (no_asm), simp]: +"globs s r=Some obj\ globs (upd_gobj r n v s) = globs s(r\upd_obj n v obj)" +apply (unfold upd_gobj_def) +apply (induct "s") +apply auto +done + +lemma locals_upd_gobj [simp]: "locals (upd_gobj r n v s) = locals s" +apply (induct "s") +by (simp add: upd_gobj_def) + + +lemma globs_init_obj [simp]: "globs (init_obj G oi r s) t = + (if t=r then Some \tag=oi,values=init_vals (var_tys G oi r)\ else globs s t)" +apply (unfold init_obj_def) +apply (simp (no_asm)) +done + +lemma locals_init_obj [simp]: "locals (init_obj G oi r s) = locals s" +by (simp add: init_obj_def) + +lemma surjective_st [simp]: "st (globs s) (locals s) = s" +apply (induct "s") +by auto + +lemma surjective_st_init_obj: + "st (globs (init_obj G oi r s)) (locals s) = init_obj G oi r s" +apply (subst locals_init_obj [THEN sym]) +apply (rule surjective_st) +done + +lemma heap_heap_upd [simp]: + "heap (st (g(Inl a\obj)) l) = heap (st g l)(a\obj)" +apply (rule ext) +apply (simp (no_asm)) +done +lemma heap_stat_upd [simp]: "heap (st (g(Inr C\obj)) l) = heap (st g l)" +apply (rule ext) +apply (simp (no_asm)) +done +lemma heap_local_upd [simp]: "heap (st g (l(vn\v))) = heap (st g l)" +apply (rule ext) +apply (simp (no_asm)) +done + +lemma heap_gupd_Heap [simp]: "heap (gupd(Heap a\obj) s) = heap s(a\obj)" +apply (rule ext) +apply (simp (no_asm)) +done +lemma heap_gupd_Stat [simp]: "heap (gupd(Stat C\obj) s) = heap s" +apply (rule ext) +apply (simp (no_asm)) +done +lemma heap_lupd [simp]: "heap (lupd(vn\v) s) = heap s" +apply (rule ext) +apply (simp (no_asm)) +done + +(* +lemma heap_upd_gobj_Heap: "!!a. heap (upd_gobj (Heap a) n v s) = ?X" +apply (rule ext) +apply (simp (no_asm)) +apply (case_tac "globs s (Heap a)") +apply auto +*) + +lemma heap_upd_gobj_Stat [simp]: "heap (upd_gobj (Stat C) n v s) = heap s" +apply (rule ext) +apply (simp (no_asm)) +apply (case_tac "globs s (Stat C)") +apply auto +done + +lemma set_locals_def2 [simp]: "set_locals l (st g l') = st g l" +apply (unfold set_locals_def) +apply (simp (no_asm)) +done + +lemma set_locals_id [simp]: "set_locals (locals s) s = s" +apply (unfold set_locals_def) +apply (induct_tac "s") +apply (simp (no_asm)) +done + +lemma set_set_locals [simp]: "set_locals l (set_locals l' s) = set_locals l s" +apply (unfold set_locals_def) +apply (induct_tac "s") +apply (simp (no_asm)) +done + +lemma locals_set_locals [simp]: "locals (set_locals l s) = l" +apply (unfold set_locals_def) +apply (induct_tac "s") +apply (simp (no_asm)) +done + +lemma globs_set_locals [simp]: "globs (set_locals l s) = globs s" +apply (unfold set_locals_def) +apply (induct_tac "s") +apply (simp (no_asm)) +done + +lemma heap_set_locals [simp]: "heap (set_locals l s) = heap s" +apply (unfold heap_def) +apply (induct_tac "s") +apply (simp (no_asm)) +done + + +section "abrupt completion" + + +datatype xcpt (* exception *) + = Loc loc (* location of allocated execption object *) + | Std xname (* intermediate standard exception, see Eval.thy *) + +datatype abrupt (* abrupt completion *) + = Xcpt xcpt (* exception *) + | Jump jump (* break, continue, return *) +consts + + the_Xcpt :: "abrupt \ xcpt" + the_Jump :: "abrupt => jump" + the_Loc :: "xcpt \ loc" + the_Std :: "xcpt \ xname" + +primrec "the_Xcpt (Xcpt x) = x" +primrec "the_Jump (Jump j) = j" +primrec "the_Loc (Loc a) = a" +primrec "the_Std (Std x) = x" + +types + abopt = "abrupt option" + + +constdefs + abrupt_if :: "bool \ abopt \ abopt \ abopt" + "abrupt_if c x' x \ if c \ (x = None) then x' else x" + +lemma abrupt_if_True_None [simp]: "abrupt_if True x None = x" +by (simp add: abrupt_if_def) + +lemma abrupt_if_True_not_None [simp]: "x \ None \ abrupt_if True x y \ None" +by (simp add: abrupt_if_def) + +lemma abrupt_if_False [simp]: "abrupt_if False x y = y" +by (simp add: abrupt_if_def) + +lemma abrupt_if_Some [simp]: "abrupt_if c x (Some y) = Some y" +by (simp add: abrupt_if_def) + +lemma abrupt_if_not_None [simp]: "y \ None \ abrupt_if c x y = y" +apply (simp add: abrupt_if_def) +by auto + + +lemma split_abrupt_if: +"P (abrupt_if c x' x) = + ((c \ x = None \ P x') \ (\ (c \ x = None) \ P x))" +apply (unfold abrupt_if_def) +apply (split split_if) +apply auto +done + +syntax + + raise_if :: "bool \ xname \ abopt \ abopt" + np :: "val \ \ abopt \ abopt" + check_neg:: "val \ \ abopt \ abopt" + +translations + + "raise_if c xn" == "abrupt_if c (Some (Xcpt (Std xn)))" + "np v" == "raise_if (v = Null) NullPointer" + "check_neg i'" == "raise_if (the_Intg i'<0) NegArrSize" + +lemma raise_if_None [simp]: "(raise_if c x y = None) = (\c \ y = None)" +apply (simp add: abrupt_if_def) +by auto +declare raise_if_None [THEN iffD1, dest!] + +lemma if_raise_if_None [simp]: + "((if b then y else raise_if c x y) = None) = ((c \ b) \ y = None)" +apply (simp add: abrupt_if_def) +apply auto +done + +lemma raise_if_SomeD [dest!]: + "raise_if c x y = Some z \ c \ z=(Xcpt (Std x)) \ y=None \ (y=Some z)" +apply (case_tac y) +apply (case_tac c) +apply (simp add: abrupt_if_def) +apply (simp add: abrupt_if_def) +apply auto +done + +constdefs + absorb :: "jump \ abopt \ abopt" + "absorb j a \ if a=Some (Jump j) then None else a" + +lemma absorb_SomeD [dest!]: "absorb j a = Some x \ a = Some x" +by (auto simp add: absorb_def) + +lemma absorb_same [simp]: "absorb j (Some (Jump j)) = None" +by (auto simp add: absorb_def) + +lemma absorb_other [simp]: "a \ Some (Jump j) \ absorb j a = a" +by (auto simp add: absorb_def) + + +section "full program state" + +types + state = "abopt \ st" (* state including exception information *) + +syntax + Norm :: "st \ state" + abrupt :: "state \ abopt" + store :: "state \ st" + +translations + + "Norm s" == "(None,s)" + "abrupt" => "fst" + "store" => "snd" + "abopt" <= (type) "State.abrupt option" + "abopt" <= (type) "abrupt option" + "state" <= (type) "abopt \ State.st" + "state" <= (type) "abopt \ st" + + + +lemma single_stateE: "\Z. Z = (s::state) \ False" +apply (erule_tac x = "(Some k,y)" in all_dupE) +apply (erule_tac x = "(None,y)" in allE) +apply clarify +done + +lemma state_not_single: "All (op = (x::state)) \ R" +apply (drule_tac x = "(if abrupt x = None then Some ?x else None,?y)" in spec) +apply clarsimp +done + +constdefs + + normal :: "state \ bool" + "normal \ \s. abrupt s = None" + +lemma normal_def2 [simp]: "normal s = (abrupt s = None)" +apply (unfold normal_def) +apply (simp (no_asm)) +done + +constdefs + heap_free :: "nat \ state \ bool" + "heap_free n \ \s. atleast_free (heap (store s)) n" + +lemma heap_free_def2 [simp]: "heap_free n s = atleast_free (heap (store s)) n" +apply (unfold heap_free_def) +apply simp +done + +subsection "update" + +constdefs + + abupd :: "(abopt \ abopt) \ state \ state" + "abupd f \ prod_fun f id" + + supd :: "(st \ st) \ state \ state" + "supd \ prod_fun id" + +lemma abupd_def2 [simp]: "abupd f (x,s) = (f x,s)" +by (simp add: abupd_def) + +lemma abupd_abrupt_if_False [simp]: "\ s. abupd (abrupt_if False xo) s = s" +by simp + +lemma supd_def2 [simp]: "supd f (x,s) = (x,f s)" +by (simp add: supd_def) + +lemma supd_lupd [simp]: + "\ s. supd (lupd vn v ) s = (abrupt s,lupd vn v (store s))" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (simp (no_asm)) +done + + +lemma supd_gupd [simp]: + "\ s. supd (gupd r obj) s = (abrupt s,gupd r obj (store s))" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (simp (no_asm)) +done + +lemma supd_init_obj [simp]: + "supd (init_obj G oi r) s = (abrupt s,init_obj G oi r (store s))" +apply (unfold init_obj_def) +apply (simp (no_asm)) +done + +syntax + + set_lvars :: "locals \ state \ state" + restore_lvars :: "state \ state \ state" + +translations + + "set_lvars l" == "supd (set_locals l)" + "restore_lvars s' s" == "set_lvars (locals (store s')) s" + +lemma set_set_lvars [simp]: "\ s. set_lvars l (set_lvars l' s) = set_lvars l s" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (simp (no_asm)) +done + +lemma set_lvars_id [simp]: "\ s. set_lvars (locals (store s)) s = s" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (simp (no_asm)) +done + +section "initialisation test" + +constdefs + + inited :: "qtname \ globs \ bool" + "inited C g \ g (Stat C) \ None" + + initd :: "qtname \ state \ bool" + "initd C \ inited C \ globs \ store" + +lemma not_inited_empty [simp]: "\inited C empty" +apply (unfold inited_def) +apply (simp (no_asm)) +done + +lemma inited_gupdate [simp]: "inited C (g(r\obj)) = (inited C g \ r = Stat C)" +apply (unfold inited_def) +apply (auto split add: st.split) +done + +lemma inited_init_class_obj [intro!]: "inited C (globs (init_class_obj G C s))" +apply (unfold inited_def) +apply (simp (no_asm)) +done + +lemma not_initedD: "\ inited C g \ g (Stat C) = None" +apply (unfold inited_def) +apply (erule notnotD) +done + +lemma initedD: "inited C g \ \ obj. g (Stat C) = Some obj" +apply (unfold inited_def) +apply auto +done + +lemma initd_def2 [simp]: "initd C s = inited C (globs (store s))" +apply (unfold initd_def) +apply (simp (no_asm)) +done + + + +end + diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Table.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Table.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,445 @@ +(* Title: isabelle/Bali/Table.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* Abstract tables and their implementation as lists *} + +theory Table = Basis: + +text {* +design issues: +\begin{itemize} +\item definition of table: infinite map vs. list vs. finite set + list chosen, because: + \begin{itemize} + \item[+] a priori finite + \item[+] lookup is more operational than for finite set + \item[-] not very abstract, but function table converts it to abstract + mapping + \end{itemize} +\item coding of lookup result: Some/None vs. value/arbitrary + Some/None chosen, because: + \begin{itemize} + \item[++] makes definedness check possible (applies also to finite set), + which is important for the type standard, hiding/overriding, etc. + (though it may perhaps be possible at least for the operational semantics + to treat programs as infinite, i.e. where classes, fields, methods etc. + of any name are considered to be defined) + \item[-] sometimes awkward case distinctions, alleviated by operator 'the' + \end{itemize} +\end{itemize} +*} + + +types ('a, 'b) table (* table with key type 'a and contents type 'b *) + = "'a \ 'b" + ('a, 'b) tables (* non-unique table with key 'a and contents 'b *) + = "'a \ 'b set" + + +section "map of / table of" + +syntax + table_of :: "('a \ 'b) list \ ('a, 'b) table" (* concrete table *) + +translations + "table_of" == "map_of" + + (type)"'a \ 'b" <= (type)"'a \ 'b Option.option" + (type)"('a, 'b) table" <= (type)"'a \ 'b" + +(* ### To map *) +lemma override_find_left[simp]: +"n k = None \ (m ++ n) k = m k" +by (simp add: override_def) + +section {* Conditional Override *} +constdefs +cond_override:: + "('b \'b \ bool) \ ('a, 'b)table \ ('a, 'b)table \ ('a, 'b) table" + +(* when merging tables old and new, only override an entry of table old when + the condition cond holds *) +"cond_override cond old new \ + \ k. + (case new k of + None \ old k + | Some new_val \ (case old k of + None \ Some new_val + | Some old_val \ (if cond new_val old_val + then Some new_val + else Some old_val)))" + +lemma cond_override_empty1[simp]: "cond_override c empty t = t" +by (simp add: cond_override_def expand_fun_eq) + +lemma cond_override_empty2[simp]: "cond_override c t empty = t" +by (simp add: cond_override_def expand_fun_eq) + +lemma cond_override_None[simp]: + "old k = None \ (cond_override c old new) k = new k" +by (simp add: cond_override_def) + +lemma cond_override_override: + "\old k = Some ov;new k = Some nv; C nv ov\ + \ (cond_override C old new) k = Some nv" +by (auto simp add: cond_override_def) + +lemma cond_override_noOverride: + "\old k = Some ov;new k = Some nv; \ (C nv ov)\ + \ (cond_override C old new) k = Some ov" +by (auto simp add: cond_override_def) + +lemma dom_cond_override: "dom (cond_override C s t) \ dom s \ dom t" +by (auto simp add: cond_override_def dom_def) + +lemma finite_dom_cond_override: + "\ finite (dom s); finite (dom t) \ \ finite (dom (cond_override C s t))" +apply (rule_tac B="dom s \ dom t" in finite_subset) +apply (rule dom_cond_override) +by (rule finite_UnI) + +section {* Filter on Tables *} + +constdefs +filter_tab:: "('a \ 'b \ bool) \ ('a, 'b) table \ ('a, 'b) table" +"filter_tab c t \ \ k. (case t k of + None \ None + | Some x \ if c k x then Some x else None)" + +lemma filter_tab_empty[simp]: "filter_tab c empty = empty" +by (simp add: filter_tab_def empty_def) + +lemma filter_tab_True[simp]: "filter_tab (\x y. True) t = t" +by (simp add: expand_fun_eq filter_tab_def) + +lemma filter_tab_False[simp]: "filter_tab (\x y. False) t = empty" +by (simp add: expand_fun_eq filter_tab_def empty_def) + +lemma filter_tab_ran_subset: "ran (filter_tab c t) \ ran t" +by (auto simp add: filter_tab_def ran_def) + +lemma filter_tab_range_subset: "range (filter_tab c t) \ range t \ {None}" +apply (auto simp add: filter_tab_def) +apply (drule sym, blast) +done + +lemma finite_range_filter_tab: + "finite (range t) \ finite (range (filter_tab c t))" +apply (rule_tac B="range t \ {None} " in finite_subset) +apply (rule filter_tab_range_subset) +apply (auto intro: finite_UnI) +done + +lemma filter_tab_SomeD[dest!]: +"filter_tab c t k = Some x \ (t k = Some x) \ c k x" +by (auto simp add: filter_tab_def) + +lemma filter_tab_SomeI: "\t k = Some x;C k x\ \filter_tab C t k = Some x" +by (simp add: filter_tab_def) + +lemma filter_tab_all_True: + "\ k y. t k = Some y \ p k y \filter_tab p t = t" +apply (auto simp add: filter_tab_def expand_fun_eq) +done + +lemma filter_tab_all_True_Some: + "\\ k y. t k = Some y \ p k y; t k = Some v\ \ filter_tab p t k = Some v" +by (auto simp add: filter_tab_def expand_fun_eq) + +lemma filter_tab_None: "t k = None \ filter_tab p t k = None" +apply (simp add: filter_tab_def expand_fun_eq) +done + +lemma filter_tab_dom_subset: "dom (filter_tab C t) \ dom t" +by (auto simp add: filter_tab_def dom_def) + +lemma filter_tab_eq: "\a=b\ \ filter_tab C a = filter_tab C b" +by (auto simp add: expand_fun_eq filter_tab_def) + +lemma finite_dom_filter_tab: +"finite (dom t) \ finite (dom (filter_tab C t))" +apply (rule_tac B="dom t" in finite_subset) +by (rule filter_tab_dom_subset) + + +lemma filter_tab_weaken: +"\\ a \ t k: \ b \ s k: P a b; + \ k x y. \t k = Some x;s k = Some y\ \ cond k x \ cond k y + \ \ \ a \ filter_tab cond t k: \ b \ filter_tab cond s k: P a b" +apply (auto simp add: filter_tab_def) +done + +lemma cond_override_filter: + "\\ k old new. \s k = Some new; t k = Some old\ + \ (\ overC new old \ \ filterC k new) \ + (overC new old \ filterC k old \ filterC k new) + \ \ + cond_override overC (filter_tab filterC t) (filter_tab filterC s) + = filter_tab filterC (cond_override overC t s)" +by (auto simp add: expand_fun_eq cond_override_def filter_tab_def ) + +section {* Misc. *} + +lemma Ball_set_table: "(\ (x,y)\ set l. P x y) \ \ x. \ y\ map_of l x: P x y" +apply (erule make_imp) +apply (induct l) +apply simp +apply (simp (no_asm)) +apply auto +done + +lemma Ball_set_tableD: + "\(\ (x,y)\ set l. P x y); x \ o2s (table_of l xa)\ \ P xa x" +apply (frule Ball_set_table) +by auto + +declare map_of_SomeD [elim] + +lemma table_of_Some_in_set: +"table_of l k = Some x \ (k,x) \ set l" +by auto + +lemma set_get_eq: + "unique l \ (k, the (table_of l k)) \ set l = (table_of l k \ None)" +apply safe +apply (fast dest!: weak_map_of_SomeI) +apply auto +done + +lemma inj_Pair_const2: "inj (\k. (k, C))" +apply (rule injI) +apply auto +done + +lemma table_of_mapconst_SomeI: + "\table_of t k = Some y'; snd y=y'; fst y=c\ \ + table_of (map (\(k,x). (k,c,x)) t) k = Some y" +apply (induct t) +apply auto +done + +lemma table_of_mapconst_NoneI: + "\table_of t k = None\ \ + table_of (map (\(k,x). (k,c,x)) t) k = None" +apply (induct t) +apply auto +done + +lemmas table_of_map2_SomeI = inj_Pair_const2 [THEN map_of_mapk_SomeI, standard] + +lemma table_of_map_SomeI [rule_format (no_asm)]: "table_of t k = Some x \ + table_of (map (\(k,x). (k, f x)) t) k = Some (f x)" +apply (induct_tac "t") +apply auto +done + +lemma table_of_remap_SomeD [rule_format (no_asm)]: + "table_of (map (\((k,k'),x). (k,(k',x))) t) k = Some (k',x) \ + table_of t (k, k') = Some x" +apply (induct_tac "t") +apply auto +done + +lemma table_of_mapf_Some [rule_format (no_asm)]: "\x y. f x = f y \ x = y \ + table_of (map (\(k,x). (k,f x)) t) k = Some (f x) \ table_of t k = Some x" +apply (induct_tac "t") +apply auto +done + +lemma table_of_mapf_SomeD [rule_format (no_asm), dest!]: +"table_of (map (\(k,x). (k, f x)) t) k = Some z \ (\y\table_of t k: z=f y)" +apply (induct_tac "t") +apply auto +done + +lemma table_of_mapf_NoneD [rule_format (no_asm), dest!]: +"table_of (map (\(k,x). (k, f x)) t) k = None \ (table_of t k = None)" +apply (induct_tac "t") +apply auto +done + +lemma table_of_mapkey_SomeD [rule_format (no_asm), dest!]: + "table_of (map (\(k,x). ((k,C),x)) t) (k,D) = Some x \ C = D \ table_of t k = Some x" +apply (induct_tac "t") +apply auto +done +lemma table_of_mapkey_SomeD2 [rule_format (no_asm), dest!]: + "table_of (map (\(k,x). ((k,C),x)) t) ek = Some x + \ C = snd ek \ table_of t (fst ek) = Some x" +apply (induct_tac "t") +apply auto +done + +lemma table_append_Some_iff: "table_of (xs@ys) k = Some z = + (table_of xs k = Some z \ (table_of xs k = None \ table_of ys k = Some z))" +apply (simp only: map_of_override [THEN sym]) +apply (rule override_Some_iff) +done + +lemma table_of_filter_unique_SomeD [rule_format (no_asm)]: + "table_of (filter P xs) k = Some z \ unique xs \ table_of xs k = Some z" +apply (induct xs) +apply (auto del: map_of_SomeD intro!: map_of_SomeD) +done + + +consts + Un_tables :: "('a, 'b) tables set \ ('a, 'b) tables" + overrides_t :: "('a, 'b) tables \ ('a, 'b) tables \ + ('a, 'b) tables" (infixl "\\" 100) + hidings_entails:: "('a, 'b) tables \ ('a, 'c) tables \ + ('b \ 'c \ bool) \ bool" ("_ hidings _ entails _" 20) + (* variant for unique table: *) + hiding_entails :: "('a, 'b) table \ ('a, 'c) table \ + ('b \ 'c \ bool) \ bool" ("_ hiding _ entails _" 20) + (* variant for a unique table and conditional overriding: *) + cond_hiding_entails :: "('a, 'b) table \ ('a, 'c) table + \ ('b \ 'c \ bool) \ ('b \ 'c \ bool) \ bool" + ("_ hiding _ under _ entails _" 20) + +defs + Un_tables_def: "Un_tables ts\\\ \k. \t\ts. t k" + overrides_t_def: "s \\ t \ \k. if t k = {} then s k else t k" + hidings_entails_def: "t hidings s entails R \ \k. \x\t k. \y\s k. R x y" + hiding_entails_def: "t hiding s entails R \ \k. \x\t k: \y\s k: R x y" + cond_hiding_entails_def: "t hiding s under C entails R + \ \k. \x\t k: \y\s k: C x y \ R x y" + +section "Untables" + +lemma Un_tablesI [intro]: "\x. \t \ ts; x \ t k\ \ x \ Un_tables ts k" +apply (simp add: Un_tables_def) +apply auto +done + +lemma Un_tablesD [dest!]: "\x. x \ Un_tables ts k \ \t. t \ ts \ x \ t k" +apply (simp add: Un_tables_def) +apply auto +done + +lemma Un_tables_empty [simp]: "Un_tables {} = (\k. {})" +apply (unfold Un_tables_def) +apply (simp (no_asm)) +done + + +section "overrides" + +lemma empty_overrides_t [simp]: "(\k. {}) \\ m = m" +apply (unfold overrides_t_def) +apply (simp (no_asm)) +done +lemma overrides_empty_t [simp]: "m \\ (\k. {}) = m" +apply (unfold overrides_t_def) +apply (simp (no_asm)) +done + +lemma overrides_t_Some_iff: + "(x \ (s \\ t) k) = (x \ t k \ t k = {} \ x \ s k)" +by (simp add: overrides_t_def) + +lemmas overrides_t_SomeD = overrides_t_Some_iff [THEN iffD1, dest!] + +lemma overrides_t_right_empty [simp]: "n k = {} \ (m \\ n) k = m k" +by (simp add: overrides_t_def) + +lemma overrides_t_find_right [simp]: "n k \ {} \ (m \\ n) k = n k" +by (simp add: overrides_t_def) + +section "hiding entails" + +lemma hiding_entailsD: + "\t hiding s entails R; t k = Some x; s k = Some y\ \ R x y" +by (simp add: hiding_entails_def) + +lemma empty_hiding_entails: "empty hiding s entails R" +by (simp add: hiding_entails_def) + +lemma hiding_empty_entails: "t hiding empty entails R" +by (simp add: hiding_entails_def) +declare empty_hiding_entails [simp] hiding_empty_entails [simp] + +section "cond hiding entails" + +lemma cond_hiding_entailsD: +"\t hiding s under C entails R; t k = Some x; s k = Some y; C x y\ \ R x y" +by (simp add: cond_hiding_entails_def) + +lemma empty_cond_hiding_entails[simp]: "empty hiding s under C entails R" +by (simp add: cond_hiding_entails_def) + +lemma cond_hiding_empty_entails[simp]: "t hiding empty under C entails R" +by (simp add: cond_hiding_entails_def) + +lemma hidings_entailsD: "\t hidings s entails R; x \ t k; y \ s k\ \ R x y" +by (simp add: hidings_entails_def) + +lemma hidings_empty_entails: "t hidings (\k. {}) entails R" +apply (unfold hidings_entails_def) +apply (simp (no_asm)) +done + +lemma empty_hidings_entails: + "(\k. {}) hidings s entails R"apply (unfold hidings_entails_def) +by (simp (no_asm)) +declare empty_hidings_entails [intro!] hidings_empty_entails [intro!] + + + +(*###TO Map?*) +consts + atleast_free :: "('a ~=> 'b) => nat => bool" +primrec + "atleast_free m 0 = True" + atleast_free_Suc: + "atleast_free m (Suc n) = (? a. m a = None & (!b. atleast_free (m(a|->b)) n))" + +lemma atleast_free_weaken [rule_format (no_asm)]: + "!m. atleast_free m (Suc n) \ atleast_free m n" +apply (induct_tac "n") +apply (simp (no_asm)) +apply clarify +apply (simp (no_asm)) +apply (drule atleast_free_Suc [THEN iffD1]) +apply fast +done + +lemma atleast_free_SucI: +"[| h a = None; !obj. atleast_free (h(a|->obj)) n |] ==> atleast_free h (Suc n)" +by force + +declare fun_upd_apply [simp del] +lemma atleast_free_SucD_lemma [rule_format (no_asm)]: +" !m a. m a = None --> (!c. atleast_free (m(a|->c)) n) --> + (!b d. a ~= b --> atleast_free (m(b|->d)) n)" +apply (induct_tac "n") +apply auto +apply (rule_tac x = "a" in exI) +apply (rule conjI) +apply (force simp add: fun_upd_apply) +apply (erule_tac V = "m a = None" in thin_rl) +apply clarify +apply (subst fun_upd_twist) +apply (erule not_sym) +apply (rename_tac "ba") +apply (drule_tac x = "ba" in spec) +apply clarify +apply (tactic "smp_tac 2 1") +apply (erule (1) notE impE) +apply (case_tac "aa = b") +apply fast+ +done +declare fun_upd_apply [simp] + +lemma atleast_free_SucD [rule_format (no_asm)]: "atleast_free h (Suc n) ==> atleast_free (h(a|->b)) n" +apply auto +apply (case_tac "aa = a") +apply auto +apply (erule atleast_free_SucD_lemma) +apply auto +done + +declare atleast_free_Suc [simp del] +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Term.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Term.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,231 @@ +(* Title: isabelle/Bali/Term.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) + +header {* Java expressions and statements *} + +theory Term = Value: + +text {* +design issues: +\begin{itemize} +\item invocation frames for local variables could be reduced to special static + objects (one per method). This would reduce redundancy, but yield a rather + non-standard execution model more difficult to understand. +\item method bodies separated from calls to handle assumptions in axiomat. + semantics + NB: Body is intended to be in the environment of the called method. +\item class initialization is regarded as (auxiliary) statement + (required for AxSem) +\item result expression of method return is handled by a special result variable + result variable is treated uniformly with local variables + \begin{itemize} + \item[+] welltypedness and existence of the result/return expression is + ensured without extra efford + \end{itemize} +\end{itemize} + +simplifications: +\begin{itemize} +\item expression statement allowed for any expression +\item no unary, binary, etc, operators +\item This is modeled as a special non-assignable local variable +\item Super is modeled as a general expression with the same value as This +\item access to field x in current class via This.x +\item NewA creates only one-dimensional arrays; + initialization of further subarrays may be simulated with nested NewAs +\item The 'Lit' constructor is allowed to contain a reference value. + But this is assumed to be prohibited in the input language, which is enforced + by the type-checking rules. +\item a call of a static method via a type name may be simulated by a dummy + variable +\item no nested blocks with inner local variables +\item no synchronized statements +\item no secondary forms of if, while (e.g. no for) (may be easily simulated) +\item no switch (may be simulated with if) +\item the @{text try_catch_finally} statement is divided into the + @{text try_catch} statement + and a finally statement, which may be considered as try..finally with + empty catch +\item the @{text try_catch} statement has exactly one catch clause; + multiple ones can be + simulated with instanceof +\item the compiler is supposed to add the annotations {@{text _}} during + type-checking. This + transformation is left out as its result is checked by the type rules anyway +\end{itemize} +*} + +datatype inv_mode (* invocation mode for method calls *) + = Static (* static *) + | SuperM (* super *) + | IntVir (* interface or virtual *) + +record sig = (* signature of a method, cf. 8.4.2 *) + name ::"mname" (* acutally belongs to Decl.thy *) + parTs::"ty list" + +translations + "sig" <= (type) "\name::mname,parTs::ty list\" + "sig" <= (type) "\name::mname,parTs::ty list,\::'a\" + +datatype jump + = Break label (* break *) + | Cont label (* continue *) + | Ret (* return from method *) + +datatype var + = LVar lname(* local variable (incl. parameters) *) + | FVar qtname bool expr vname(*class field*)("{_,_}_.._"[10,10,85,99]90) + | AVar expr expr (* array component *) ("_.[_]"[90,10 ]90) + +and expr + = NewC qtname (* class instance creation *) + | NewA ty expr (* array creation *) ("New _[_]"[99,10 ]85) + | Cast ty expr (* type cast *) + | Inst expr ref_ty (* instanceof *) ("_ InstOf _"[85,99] 85) + | Lit val (* literal value, references not allowed *) + | Super (* special Super keyword *) + | Acc var (* variable access *) + | Ass var expr (* variable assign *) ("_:=_" [90,85 ]85) + | Cond expr expr expr (* conditional *) ("_ ? _ : _" [85,85,80]80) + | Call ref_ty inv_mode expr mname "(ty list)" (* method call *) + "(expr list)" ("{_,_}_\_'( {_}_')"[10,10,85,99,10,10]85) + | Methd qtname sig (* (folded) method (see below) *) + | Body qtname stmt (* (unfolded) method body *) +and stmt + = Skip (* empty statement *) + | Expr expr (* expression statement *) + | Lab label stmt ("_\ _"(* labeled statement*)[ 99,66]66) + (* handles break *) + | Comp stmt stmt ("_;; _" [ 66,65]65) + | If_ expr stmt stmt ("If'(_') _ Else _" [ 80,79,79]70) + | Loop label expr stmt ("_\ While'(_') _" [ 99,80,79]70) + | Do jump (* break, continue, return *) + | Throw expr + | TryC stmt + qtname vname stmt ("Try _ Catch'(_ _') _" [79,99,80,79]70) + | Fin stmt stmt ("_ Finally _" [ 79,79]70) + | Init qtname (* class initialization *) + + +text {* +The expressions Methd and Body are artificial program constructs, in the +sense that they are not used to define a concrete Bali program. In the +evaluation semantic definition they are "generated on the fly" to decompose +the task to define the behaviour of the Call expression. They are crucial +for the axiomatic semantics to give a syntactic hook to insert +some assertions (cf. AxSem.thy, Eval.thy). +Also the Init statement (to initialize a class on its first use) is inserted +in various places by the evaluation semantics. +*} + +types "term" = "(expr+stmt, var, expr list) sum3" +translations + "sig" <= (type) "mname \ ty list" + "var" <= (type) "Term.var" + "expr" <= (type) "Term.expr" + "stmt" <= (type) "Term.stmt" + "term" <= (type) "(expr+stmt, var, expr list) sum3" + +syntax + + this :: expr + LAcc :: "vname \ expr" ("!!") + LAss :: "vname \ expr \ stmt" ("_:==_" [90,85] 85) + Return :: "expr \ stmt" + StatRef :: "ref_ty \ expr" + +translations + + "this" == "Acc (LVar This)" + "!!v" == "Acc (LVar (EName (VNam v)))" + "v:==e" == "Expr (Ass (LVar (EName (VNam v))) e)" + "Return e" == "Expr (Ass (LVar (EName Res)) e);; Do Ret" + (* Res := e;; Do Ret *) + "StatRef rt" == "Cast (RefT rt) (Lit Null)" + +constdefs + + is_stmt :: "term \ bool" + "is_stmt t \ \c. t=In1r c" + +ML {* +bind_thms ("is_stmt_rews", sum3_instantiate (thm "is_stmt_def")); +*} + +declare is_stmt_rews [simp] + + +(* ############# Just testing syntax *) +(** unfortunately cannot simply instantiate tnam **) +(* +datatype tnam_ = HasFoo_ | Base_ | Ext_ +datatype vnam_ = arr_ | vee_ | z_ | e_ +datatype label_ = lab1_ + +consts + + tnam_ :: "tnam_ \ tnam" + vnam_ :: "vnam_ \ vname" + label_:: "label_ \ label" +axioms + + inj_tnam_ [simp]: "(tnam_ x = tnam_ y) = (x = y)" + inj_vnam_ [simp]: "(vnam_ x = vnam_ y) = (x = y)" + inj_label_ [simp]: "(label_ x = label_ y) = (x = y)" + + + surj_tnam_: "\m. n = tnam_ m" + surj_vnam_: "\m. n = vnam_ m" + surj_label_:" \m. n = label_ m" + +syntax + + HasFoo :: qtname + Base :: qtname + Ext :: qtname + arr :: ename + vee :: ename + z :: ename + e :: ename + lab1:: label + +consts + + foo :: mname +translations + + "HasFoo" == "\pid=java_lang,tid=TName (tnam_ HasFoo_)\" + "Base" == "\pid=java_lang,tid=TName (tnam_ Base_)\" + "Ext" == "\pid=java_lang,tid=TName (tnam_ Ext_)\" + "arr" == "(vnam_ arr_)" + "vee" == "(vnam_ vee_)" + "z" == "(vnam_ z_)" + "e" == "(vnam_ e_)" + "lab1" == "label_ lab1_" + +constdefs test::stmt +"test \ +(lab1@ While(Acc + (Acc ({Base,True}StatRef (ClassT Object).arr).[Lit (Intg #2)])) Skip)" + +consts + pTs :: "ty list" + +constdefs + +test1::stmt +"test1 \ + Expr({ClassT Base,IntVir}!!e\foo({pTs}[Lit Null]))" + + + +constdefs test::stmt +"test \ +(lab1\ While(Acc + (Acc ({Base,True}StatRef (ClassT Object)..arr).[Lit (Intg #2)])) Skip)" +*) +end \ No newline at end of file diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Trans.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Trans.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,144 @@ +(* Title: isabelle/Bali/Trans.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen + +Operational transition (small-step) semantics of the +execution of Java expressions and statements + +PRELIMINARY!!!!!!!! + +*) + +Trans = Eval + + +consts + texpr_tstmt :: "prog \ (((expr \ state) \ (expr \ state)) + + ((stmt \ state) \ (stmt \ state))) set" + +syntax (symbols) + texpr :: "[prog, expr \ state, expr \ state] \ bool "("_\_ \1 _"[61,82,82] 81) + tstmt :: "[prog, stmt \ state, stmt \ state] \ bool "("_\_ \1 _"[61,82,82] 81) + Ref :: "loc \ expr" + +translations + + "G\e_s \1 ex_s'" == "Inl (e_s, ex_s') \ texpr_tstmt G" + "G\s_s \1 s'_s'" == "Inr (s_s, s'_s') \ texpr_tstmt G" + "Ref a" == "Lit (Addr a)" + +inductive "texpr_tstmt G" intrs + +(* evaluation of expression *) + (* cf. 15.5 *) + XcptE "\\v. e \ Lit v\ \\ + G\(e,Some xc,s) \1 (Lit arbitrary,Some xc,s)" + + (* cf. 15.8.1 *) + NewC "\new_Addr (heap s) = Some (a,x); + s' = assign (hupd[a\init_Obj G C]s) (x,s)\ \\ + G\(NewC C,None,s) \1 (Ref a,s')" + + (* cf. 15.9.1 *) + NewA1 "\G\(e,None,s) \1 (e',s')\ \\ + G\(New T[e],None,s) \1 (New T[e'],s')" + NewA "\i = the_Intg i'; new_Addr (heap s) = Some (a, x); + s' = assign (hupd[a\init_Arr T i]s)(raise_if (i<#0) NegArrSize x,s)\ \\ + G\(New T[Lit i'],None,s) \1 (Ref a,s')" + (* cf. 15.15 *) + Cast1 "\G\(e,None,s) \1 (e',s')\ \\ + G\(Cast T e,None,s) \1 (Cast T e',s')" + Cast "\x'= raise_if (\G,s\v fits T) ClassCast None\ \\ + G\(Cast T (Lit v),None,s) \1 (Lit v,x',s)" + + (* cf. 15.7.1 *) +(*Lit "G\(Lit v,None,s) \1 (Lit v,None,s)"*) + + (* cf. 15.13.1, 15.2 *) + LAcc "\v = the (locals s vn)\ \\ + G\(LAcc vn,None,s) \1 (Lit v,None,s)" + + (* cf. 15.25.1 *) + LAss1 "\G\(e,None,s) \1 (e',s')\ \\ + G\(f vn:=e,None,s) \1 (vn:=e',s')" + LAss "G\(f vn:=Lit v,None,s) \1 (Lit v,None,lupd[vn\v]s)" + + (* cf. 15.10.1, 15.2 *) + FAcc1 "\G\(e,None,s) \1 (e',s')\ \\ + G\({T}e..fn,None,s) \1 ({T}e'..fn,s')" + FAcc "\v = the (snd (the_Obj (heap s (the_Addr a'))) (fn,T))\ \\ + G\({T}Lit a'..fn,None,s) \1 (Lit v,np a' None,s)" + + (* cf. 15.25.1 *) + FAss1 "\G\(e1,None,s) \1 (e1',s')\ \\ + G\(f ({T}e1..fn):=e2,None,s) \1 (f({T}e1'..fn):=e2,s')" + FAss2 "\G\(e2,np a' None,s) \1 (e2',s')\ \\ + G\(f({T}Lit a'..fn):=e2,None,s) \1 (f({T}Lit a'..fn):=e2',s')" + FAss "\a = the_Addr a'; (c,fs) = the_Obj (heap s a); + s'= assign (hupd[a\Obj c (fs[(fn,T)\v])]s) (None,s)\ \\ + G\(f({T}Lit a'..fn):=Lit v,None,s) \1 (Lit v,s')" + + + + + + (* cf. 15.12.1 *) + AAcc1 "\G\(e1,None,s) \1 (e1',s')\ \\ + G\(e1[e2],None,s) \1 (e1'[e2],s')" + AAcc2 "\G\(e2,None,s) \1 (e2',s')\ \\ + G\(Lit a'[e2],None,s) \1 (Lit a'[e2'],s')" + AAcc "\vo = snd (the_Arr (heap s (the_Addr a'))) (the_Intg i'); + x' = raise_if (vo = None) IndOutBound (np a' None)\ \\ + G\(Lit a'[Lit i'],None,s) \1 (Lit (the vo),x',s)" + + + (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *) + Call1 "\G\(e,None,s) \1 (e',s')\ \\ + G\(e..mn({pT}p),None,s) \1 (e'..mn({pT}p),s')" + Call2 "\G\(p,None,s) \1 (p',s')\ \\ + G\(Lit a'..mn({pT}p),None,s) \1 (Lit a'..mn({pT}p'),s')" + Call "\a = the_Addr a'; (md,(pn,rT),lvars,blk,res) = + the (cmethd G (fst (the_Obj (h a))) (mn,pT))\ \\ + G\(Lit a'..mn({pT}Lit pv),None,(h,l)) \1 + (Body blk res l,np a' x,(h,init_vals lvars[This\a'][Super\a'][pn\pv]))" + Body1 "\G\(s0,None,s) \1 (s0',s')\ \\ + G\(Body s0 e l,None,s) \1 (Body s0' e l,s')" + Body2 "\G\(e ,None,s) \1 (e',s')\ \\ + G\(Body Skip e l,None,s) \1 (Body Skip e' l,s')" + Body "G\(Body Skip (Lit v) l,None,s) \1 (Lit v,None,(heap s,l))" + +(* execution of statements *) + + (* cf. 14.1 *) + XcptS "\s0 \ Skip\ \\ + G\(s0,Some xc,s) \1 (Skip,Some xc,s)" + + (* cf. 14.5 *) +(*Skip "G\(Skip,None,s) \1 (Skip,None,s)"*) + + (* cf. 14.2 *) + Comp1 "\G\(s1,None,s) \1 (s1',s')\ \\ + G\(s1;; s2,None,s) \1 (s1';; s2,s')" + Comp "G\(Skip;; s2,None,s) \1 (s2,None,s)" + + + + + + + (* cf. 14.7 *) + Expr1 "\G\(e ,None,s) \1 (e',s')\ \\ + G\(Expr e,None,s) \1 (Expr e',s')" + Expr "G\(Expr (Lit v),None,s) \1 (Skip,None,s)" + + (* cf. 14.8.2 *) + If1 "\G\(e ,None,s) \1 (e',s')\ \\ + G\(If(e) s1 Else s2,None,s) \1 (If(e') s1 Else s2,s')" + If "G\(If(Lit v) s1 Else s2,None,s) \1 + (if the_Bool v then s1 else s2,None,s)" + + (* cf. 14.10, 14.10.1 *) + Loop "G\(While(e) s0,None,s) \1 + (If(e) (s0;; While(e) s0) Else Skip,None,s)" + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Type.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Type.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,59 @@ +(* Title: isabelle/Bali/Type.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) + +header {* Java types *} + +theory Type = Name: + +text {* +simplifications: +\begin{itemize} +\item only the most important primitive types +\item the null type is regarded as reference type +\end{itemize} +*} + +datatype prim_ty (* primitive type, cf. 4.2 *) + = Void (* 'result type' of void methods *) + | Boolean + | Integer + + +datatype ref_ty (* reference type, cf. 4.3 *) + = NullT (* null type, cf. 4.1 *) + | IfaceT qtname (* interface type *) + | ClassT qtname (* class type *) + | ArrayT ty (* array type *) + +and ty (* any type, cf. 4.1 *) + = PrimT prim_ty (* primitive type *) + | RefT ref_ty (* reference type *) + +translations + "prim_ty" <= (type) "Type.prim_ty" + "ref_ty" <= (type) "Type.ref_ty" + "ty" <= (type) "Type.ty" + +syntax + NT :: " \ ty" + Iface :: "qtname \ ty" + Class :: "qtname \ ty" + Array :: "ty \ ty" ("_.[]" [90] 90) + +translations + "NT" == "RefT NullT" + "Iface I" == "RefT (IfaceT I)" + "Class C" == "RefT (ClassT C)" + "T.[]" == "RefT (ArrayT T)" + +constdefs + the_Class :: "ty \ qtname" + "the_Class T \ \C. T = Class C" (** primrec not possible here **) + +lemma the_Class_eq [simp]: "the_Class (Class C)= C" +by (auto simp add: the_Class_def) + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/TypeRel.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/TypeRel.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,667 @@ +(* Title: isabelle/Bali/TypeRel.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* The relations between Java types *} + +theory TypeRel = Decl: + +text {* +simplifications: +\begin{itemize} +\item subinterface, subclass and widening relation includes identity +\end{itemize} +improvements over Java Specification 1.0: +\begin{itemize} +\item narrowing reference conversion also in cases where the return types of a + pair of methods common to both types are in widening (rather identity) + relation +\item one could add similar constraints also for other cases +\end{itemize} +design issues: +\begin{itemize} +\item the type relations do not require @{text is_type} for their arguments +\item the subint1 and subcls1 relations imply @{text is_iface}/@{text is_class} + for their first arguments, which is required for their finiteness +\end{itemize} +*} + +consts + +(*subint1, in Decl.thy*)  (* direct subinterface *) +(*subint , by translation*)   (* subinterface (+ identity) *) +(*subcls1, in Decl.thy*)  (* direct subclass *) +(*subcls , by translation*)   (* subclass *) +(*subclseq, by translation*)  (* subclass + identity *) + implmt1 :: "prog \ (qtname \ qtname) set" (* direct implementation *) + implmt :: "prog \ (qtname \ qtname) set" (* implementation *) + widen :: "prog \ (ty \ ty ) set" (* widening *) + narrow :: "prog \ (ty \ ty ) set" (* narrowing *) + cast :: "prog \ (ty \ ty ) set"  (* casting *) + +syntax + + "@subint1" :: "prog => [qtname, qtname] => bool" ("_|-_<:I1_" [71,71,71] 70) + "@subint" :: "prog => [qtname, qtname] => bool" ("_|-_<=:I _"[71,71,71] 70) + (* Defined in Decl.thy: + "@subcls1" :: "prog => [qtname, qtname] => bool" ("_|-_<:C1_" [71,71,71] 70) + "@subclseq":: "prog => [qtname, qtname] => bool" ("_|-_<=:C _"[71,71,71] 70) + "@subcls" :: "prog => [qtname, qtname] => bool" ("_|-_<:C _"[71,71,71] 70) + *) + "@implmt1" :: "prog => [qtname, qtname] => bool" ("_|-_~>1_" [71,71,71] 70) + "@implmt" :: "prog => [qtname, qtname] => bool" ("_|-_~>_" [71,71,71] 70) + "@widen" :: "prog => [ty , ty ] => bool" ("_|-_<=:_" [71,71,71] 70) + "@narrow" :: "prog => [ty , ty ] => bool" ("_|-_:>_" [71,71,71] 70) + "@cast" :: "prog => [ty , ty ] => bool" ("_|-_<=:? _"[71,71,71] 70) + +syntax (symbols) + + "@subint1" :: "prog \ [qtname, qtname] \ bool" ("_\_\I1_" [71,71,71] 70) + "@subint" :: "prog \ [qtname, qtname] \ bool" ("_\_\I _" [71,71,71] 70) + (* Defined in Decl.thy: +\ "@subcls1" :: "prog \ [qtname, qtname] \ bool" ("_\_\\<^sub>C\<^sub>1_" [71,71,71] 70) + "@subclseq":: "prog \ [qtname, qtname] \ bool" ("_\_\\<^sub>C _" [71,71,71] 70) + "@subcls" :: "prog \ [qtname, qtname] \ bool" ("_\_\\<^sub>C _" [71,71,71] 70) + *) + "@implmt1" :: "prog \ [qtname, qtname] \ bool" ("_\_\1_" [71,71,71] 70) + "@implmt" :: "prog \ [qtname, qtname] \ bool" ("_\_\_" [71,71,71] 70) + "@widen" :: "prog \ [ty , ty ] \ bool" ("_\_\_" [71,71,71] 70) + "@narrow" :: "prog \ [ty , ty ] \ bool" ("_\_\_" [71,71,71] 70) + "@cast" :: "prog \ [ty , ty ] \ bool" ("_\_\? _" [71,71,71] 70) + +translations + + "G\I \I1 J" == "(I,J) \ subint1 G" + "G\I \I J" == "(I,J) \(subint1 G)^*" (* cf. 9.1.3 *) + (* Defined in Decl.thy: + "G\C \\<^sub>C\<^sub>1 D" == "(C,D) \ subcls1 G" + "G\C \\<^sub>C D" == "(C,D) \(subcls1 G)^*" + *) + "G\C \1 I" == "(C,I) \ implmt1 G" + "G\C \ I" == "(C,I) \ implmt G" + "G\S \ T" == "(S,T) \ widen G" + "G\S \ T" == "(S,T) \ narrow G" + "G\S \? T" == "(S,T) \ cast G" + + +section "subclass and subinterface relations" + + (* direct subinterface in Decl.thy, cf. 9.1.3 *) + (* direct subclass in Decl.thy, cf. 8.1.3 *) + +lemmas subcls_direct = subcls1I [THEN r_into_rtrancl, standard] + +lemma subcls_direct1: + "\class G C = Some c; C \ Object;D=super c\ \ G\C\\<^sub>C D" +apply (auto dest: subcls_direct) +done + +lemma subcls1I1: + "\class G C = Some c; C \ Object;D=super c\ \ G\C\\<^sub>C\<^sub>1 D" +apply (auto dest: subcls1I) +done + +lemma subcls_direct2: + "\class G C = Some c; C \ Object;D=super c\ \ G\C\\<^sub>C D" +apply (auto dest: subcls1I1) +done + +lemma subclseq_trans: "\G\A \\<^sub>C B; G\B \\<^sub>C C\ \ G\A \\<^sub>C C" +by (blast intro: rtrancl_trans) + +lemma subcls_trans: "\G\A \\<^sub>C B; G\B \\<^sub>C C\ \ G\A \\<^sub>C C" +by (blast intro: trancl_trans) + +lemma SXcpt_subcls_Throwable_lemma: +"\class G (SXcpt xn) = Some xc; + super xc = (if xn = Throwable then Object else SXcpt Throwable)\ +\ G\SXcpt xn\\<^sub>C SXcpt Throwable" +apply (case_tac "xn = Throwable") +apply simp_all +apply (drule subcls_direct) +apply (auto dest: sym) +done + +lemma subcls_ObjectI: "\is_class G C; ws_prog G\ \ G\C\\<^sub>C Object" +apply (erule ws_subcls1_induct) +apply clarsimp +apply (case_tac "C = Object") +apply (fast intro: r_into_rtrancl [THEN rtrancl_trans])+ +done + +lemma subclseq_ObjectD [dest!]: "G\Object\\<^sub>C C \ C = Object" +apply (erule rtrancl_induct) +apply (auto dest: subcls1D) +done + +lemma subcls_ObjectD [dest!]: "G\Object\\<^sub>C C \ False" +apply (erule trancl_induct) +apply (auto dest: subcls1D) +done + +lemma subcls_ObjectI1 [intro!]: + "\C \ Object;is_class G C;ws_prog G\ \ G\C \\<^sub>C Object" +apply (drule (1) subcls_ObjectI) +apply (auto intro: rtrancl_into_trancl3) +done + +lemma subcls_is_class: "(C,D) \ (subcls1 G)^+ \ is_class G C" +apply (erule trancl_trans_induct) +apply (auto dest!: subcls1D) +done + +lemma subcls_is_class2 [rule_format (no_asm)]: + "G\C\\<^sub>C D \ is_class G D \ is_class G C" +apply (erule rtrancl_induct) +apply (drule_tac [2] subcls1D) +apply auto +done + +lemma single_inheritance: +"\G\A \\<^sub>C\<^sub>1 B; G\A \\<^sub>C\<^sub>1 C\ \ B = C" +by (auto simp add: subcls1_def) + +lemma subcls_compareable: +"\G\A \\<^sub>C X; G\A \\<^sub>C Y + \ \ G\X \\<^sub>C Y \ G\Y \\<^sub>C X" +by (rule triangle_lemma) (auto intro: single_inheritance) + +lemma subcls1_irrefl: "\G\C \\<^sub>C\<^sub>1 D; ws_prog G \ + \ C \ D" +proof + assume ws: "ws_prog G" and + subcls1: "G\C \\<^sub>C\<^sub>1 D" and + eq_C_D: "C=D" + from subcls1 obtain c + where + neq_C_Object: "C\Object" and + clsC: "class G C = Some c" and + super_c: "super c = D" + by (auto simp add: subcls1_def) + with super_c subcls1 eq_C_D + have subcls_super_c_C: "G\super c \\<^sub>C C" + by auto + from ws clsC neq_C_Object + have "\ G\super c \\<^sub>C C" + by (auto dest: ws_prog_cdeclD) + from this subcls_super_c_C + show "False" + by (rule notE) +qed + +lemma no_subcls_Object: "G\C \\<^sub>C D \ C \ Object" +by (erule converse_trancl_induct) (auto dest: subcls1D) + +lemma subcls_acyclic: "\G\C \\<^sub>C D; ws_prog G\ \ \ G\D \\<^sub>C C" +proof - + assume ws: "ws_prog G" + assume subcls_C_D: "G\C \\<^sub>C D" + then show ?thesis + proof (induct rule: converse_trancl_induct) + fix C + assume subcls1_C_D: "G\C \\<^sub>C\<^sub>1 D" + then obtain c where + "C\Object" and + "class G C = Some c" and + "super c = D" + by (auto simp add: subcls1_def) + with ws + show "\ G\D \\<^sub>C C" + by (auto dest: ws_prog_cdeclD) + next + fix C Z + assume subcls1_C_Z: "G\C \\<^sub>C\<^sub>1 Z" and + subcls_Z_D: "G\Z \\<^sub>C D" and + nsubcls_D_Z: "\ G\D \\<^sub>C Z" + show "\ G\D \\<^sub>C C" + proof + assume subcls_D_C: "G\D \\<^sub>C C" + show "False" + proof - + from subcls_D_C subcls1_C_Z + have "G\D \\<^sub>C Z" + by (auto dest: r_into_trancl trancl_trans) + with nsubcls_D_Z + show ?thesis + by (rule notE) + qed + qed + qed +qed + +lemma subclseq_cases [consumes 1, case_names Eq Subcls]: + "\G\C \\<^sub>C D; C = D \ P; G\C \\<^sub>C D \ P\ \ P" +by (blast intro: rtrancl_cases) + +lemma subclseq_acyclic: + "\G\C \\<^sub>C D; G\D \\<^sub>C C; ws_prog G\ \ C=D" +by (auto elim: subclseq_cases dest: subcls_acyclic) + +lemma subcls_irrefl: "\G\C \\<^sub>C D; ws_prog G\ + \ C \ D" +proof - + assume ws: "ws_prog G" + assume subcls: "G\C \\<^sub>C D" + then show ?thesis + proof (induct rule: converse_trancl_induct) + fix C + assume "G\C \\<^sub>C\<^sub>1 D" + with ws + show "C\D" + by (blast dest: subcls1_irrefl) + next + fix C Z + assume subcls1_C_Z: "G\C \\<^sub>C\<^sub>1 Z" and + subcls_Z_D: "G\Z \\<^sub>C D" and + neq_Z_D: "Z \ D" + show "C\D" + proof + assume eq_C_D: "C=D" + show "False" + proof - + from subcls1_C_Z eq_C_D + have "G\D \\<^sub>C Z" + by (auto) + also + from subcls_Z_D ws + have "\ G\D \\<^sub>C Z" + by (rule subcls_acyclic) + ultimately + show ?thesis + by - (rule notE) + qed + qed + qed +qed + +lemma invert_subclseq: +"\G\C \\<^sub>C D; ws_prog G\ + \ \ G\D \\<^sub>C C" +proof - + assume ws: "ws_prog G" and + subclseq_C_D: "G\C \\<^sub>C D" + show ?thesis + proof (cases "D=C") + case True + with ws + show ?thesis + by (auto dest: subcls_irrefl) + next + case False + with subclseq_C_D + have "G\C \\<^sub>C D" + by (blast intro: rtrancl_into_trancl3) + with ws + show ?thesis + by (blast dest: subcls_acyclic) + qed +qed + +lemma invert_subcls: +"\G\C \\<^sub>C D; ws_prog G\ + \ \ G\D \\<^sub>C C" +proof - + assume ws: "ws_prog G" and + subcls_C_D: "G\C \\<^sub>C D" + then + have nsubcls_D_C: "\ G\D \\<^sub>C C" + by (blast dest: subcls_acyclic) + show ?thesis + proof + assume "G\D \\<^sub>C C" + then show "False" + proof (cases rule: subclseq_cases) + case Eq + with ws subcls_C_D + show ?thesis + by (auto dest: subcls_irrefl) + next + case Subcls + with nsubcls_D_C + show ?thesis + by blast + qed + qed +qed + +lemma subcls_superD: + "\G\C \\<^sub>C D; class G C = Some c\ \ G\(super c) \\<^sub>C D" +proof - + assume clsC: "class G C = Some c" + assume subcls_C_C: "G\C \\<^sub>C D" + then obtain S where + "G\C \\<^sub>C\<^sub>1 S" and + subclseq_S_D: "G\S \\<^sub>C D" + by (blast dest: tranclD) + with clsC + have "S=super c" + by (auto dest: subcls1D) + with subclseq_S_D show ?thesis by simp +qed + + +lemma subclseq_superD: + "\G\C \\<^sub>C D; C\D;class G C = Some c\ \ G\(super c) \\<^sub>C D" +proof - + assume neq_C_D: "C\D" + assume clsC: "class G C = Some c" + assume subclseq_C_D: "G\C \\<^sub>C D" + then show ?thesis + proof (cases rule: subclseq_cases) + case Eq with neq_C_D show ?thesis by contradiction + next + case Subcls + with clsC show ?thesis by (blast dest: subcls_superD) + qed +qed + +section "implementation relation" + +defs + (* direct implementation, cf. 8.1.3 *) +implmt1_def:"implmt1 G\{(C,I). C\Object \ (\c\class G C: I\set (superIfs c))}" + +lemma implmt1D: "G\C\1I \ C\Object \ (\c\class G C: I\set (superIfs c))" +apply (unfold implmt1_def) +apply auto +done + + +inductive "implmt G" intros (* cf. 8.1.4 *) + + direct: "G\C\1J \\ \ G\C\J" + subint: "\G\C\1I; G\I\I J\ \ G\C\J" + subcls1: "\G\C\\<^sub>C\<^sub>1D; G\D\J \ \ G\C\J" + +lemma implmtD: "G\C\J \ (\I. G\C\1I \ G\I\I J) \ (\D. G\C\\<^sub>C\<^sub>1D \ G\D\J)" +apply (erule implmt.induct) +apply fast+ +done + +lemma implmt_ObjectE [elim!]: "G\Object\I \ R" +by (auto dest!: implmtD implmt1D subcls1D) + +lemma subcls_implmt [rule_format (no_asm)]: "G\A\\<^sub>C B \ G\B\K \ G\A\K" +apply (erule rtrancl_induct) +apply (auto intro: implmt.subcls1) +done + +lemma implmt_subint2: "\ G\A\J; G\J\I K\ \ G\A\K" +apply (erule make_imp, erule implmt.induct) +apply (auto dest: implmt.subint rtrancl_trans implmt.subcls1) +done + +lemma implmt_is_class: "G\C\I \ is_class G C" +apply (erule implmt.induct) +apply (blast dest: implmt1D subcls1D)+ +done + + +section "widening relation" + +inductive "widen G" intros(*widening, viz. method invocation conversion, cf. 5.3 + i.e. kind of syntactic subtyping *) + refl: "G\T\T"(*identity conversion, cf. 5.1.1 *) + subint: "G\I\I J \ G\Iface I\ Iface J"(*wid.ref.conv.,cf. 5.1.4 *) + int_obj: "G\Iface I\ Class Object" + subcls: "G\C\\<^sub>C D \ G\Class C\ Class D" + implmt: "G\C\I \ G\Class C\ Iface I" + null: "G\NT\ RefT R" + arr_obj: "G\T.[]\ Class Object" + array: "G\RefT S\RefT T \ G\RefT S.[]\ RefT T.[]" + +declare widen.refl [intro!] +declare widen.intros [simp] + +(* too strong in general: +lemma widen_PrimT: "G\PrimT x\T \ T = PrimT x" +*) +lemma widen_PrimT: "G\PrimT x\T \ (\y. T = PrimT y)" +apply (ind_cases "G\S\T") +by auto + +(* too strong in general: +lemma widen_PrimT2: "G\S\PrimT x \ S = PrimT x" +*) +lemma widen_PrimT2: "G\S\PrimT x \ \y. S = PrimT y" +apply (ind_cases "G\S\T") +by auto + +lemma widen_RefT: "G\RefT R\T \ \t. T=RefT t" +apply (ind_cases "G\S\T") +by auto + +lemma widen_RefT2: "G\S\RefT R \ \t. S=RefT t" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Iface: "G\Iface I\T \ T=Class Object \ (\J. T=Iface J)" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Iface2: "G\S\ Iface J \ S = NT \ (\I. S = Iface I) \ (\D. S = Class D)" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Iface_Iface: "G\Iface I\ Iface J \ G\I\I J" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Iface_Iface_eq [simp]: "G\Iface I\ Iface J = G\I\I J" +apply (rule iffI) +apply (erule widen_Iface_Iface) +apply (erule widen.subint) +done + +lemma widen_Class: "G\Class C\T \ (\D. T=Class D) \ (\I. T=Iface I)" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Class2: "G\S\ Class C \ C = Object \ S = NT \ (\D. S = Class D)" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Class_Class: "G\Class C\ Class cm \ G\C\\<^sub>C cm" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Class_Class_eq [simp]: "G\Class C\ Class cm = G\C\\<^sub>C cm" +apply (rule iffI) +apply (erule widen_Class_Class) +apply (erule widen.subcls) +done + +lemma widen_Class_Iface: "G\Class C\ Iface I \ G\C\I" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Class_Iface_eq [simp]: "G\Class C\ Iface I = G\C\I" +apply (rule iffI) +apply (erule widen_Class_Iface) +apply (erule widen.implmt) +done + +lemma widen_Array: "G\S.[]\T \ T=Class Object \ (\T'. T=T'.[] \ G\S\T')" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Array2: "G\S\T.[] \ S = NT \ (\S'. S=S'.[] \ G\S'\T)" +apply (ind_cases "G\S\T") +by auto + + +lemma widen_ArrayPrimT: "G\PrimT t.[]\T \ T=Class Object \ T=PrimT t.[]" +apply (ind_cases "G\S\T") +by auto + +lemma widen_ArrayRefT: + "G\RefT t.[]\T \ T=Class Object \ (\s. T=RefT s.[] \ G\RefT t\RefT s)" +apply (ind_cases "G\S\T") +by auto + +lemma widen_ArrayRefT_ArrayRefT_eq [simp]: + "G\RefT T.[]\RefT T'.[] = G\RefT T\RefT T'" +apply (rule iffI) +apply (drule widen_ArrayRefT) +apply simp +apply (erule widen.array) +done + +lemma widen_Array_Array: "G\T.[]\T'.[] \ G\T\T'" +apply (drule widen_Array) +apply auto +done + +lemma widen_Array_Class: "G\S.[] \ Class C \ C=Object" +by (auto dest: widen_Array) + +(* +qed_typerel "widen_NT2" "G\S\NT \ S = NT" + [prove_widen_lemma "G\S\T \ T = NT \ S = NT"] +*) +lemma widen_NT2: "G\S\NT \ S = NT" +apply (ind_cases "G\S\T") +by auto + +lemma widen_Object:"\isrtype G T;ws_prog G\ \ G\RefT T \ Class Object" +apply (case_tac T) +apply (auto) +apply (subgoal_tac "G\pid_field_type\\<^sub>C Object") +apply (auto intro: subcls_ObjectI) +done + +lemma widen_trans_lemma [rule_format (no_asm)]: + "\G\S\U; \C. is_class G C \ G\C\\<^sub>C Object\ \ \T. G\U\T \ G\S\T" +apply (erule widen.induct) +apply safe +prefer 5 apply (drule widen_RefT) apply clarsimp +apply (frule_tac [1] widen_Iface) +apply (frule_tac [2] widen_Class) +apply (frule_tac [3] widen_Class) +apply (frule_tac [4] widen_Iface) +apply (frule_tac [5] widen_Class) +apply (frule_tac [6] widen_Array) +apply safe +apply (rule widen.int_obj) +prefer 6 apply (drule implmt_is_class) apply simp +apply (tactic "ALLGOALS (etac thin_rl)") +prefer 6 apply simp +apply (rule_tac [9] widen.arr_obj) +apply (rotate_tac [9] -1) +apply (frule_tac [9] widen_RefT) +apply (auto elim!: rtrancl_trans subcls_implmt implmt_subint2) +done + +lemma ws_widen_trans: "\G\S\U; G\U\T; ws_prog G\ \ G\S\T" +by (auto intro: widen_trans_lemma subcls_ObjectI) + +lemma widen_antisym_lemma [rule_format (no_asm)]: "\G\S\T; + \I J. G\I\I J \ G\J\I I \ I = J; + \C D. G\C\\<^sub>C D \ G\D\\<^sub>C C \ C = D; + \I . G\Object\I \ False\ \ G\T\S \ S = T" +apply (erule widen.induct) +apply (auto dest: widen_Iface widen_NT2 widen_Class) +done + +lemmas subint_antisym = + subint1_acyclic [THEN acyclic_impl_antisym_rtrancl, standard] +lemmas subcls_antisym = + subcls1_acyclic [THEN acyclic_impl_antisym_rtrancl, standard] + +lemma widen_antisym: "\G\S\T; G\T\S; ws_prog G\ \ S=T" +by (fast elim: widen_antisym_lemma subint_antisym [THEN antisymD] + subcls_antisym [THEN antisymD]) + +lemma widen_ObjectD [dest!]: "G\Class Object\T \ T=Class Object" +apply (frule widen_Class) +apply (fast dest: widen_Class_Class widen_Class_Iface) +done + +constdefs + widens :: "prog \ [ty list, ty list] \ bool" ("_\_[\]_" [71,71,71] 70) + "G\Ts[\]Ts' \ list_all2 (\T T'. G\T\T') Ts Ts'" + +lemma widens_Nil [simp]: "G\[][\][]" +apply (unfold widens_def) +apply auto +done + +lemma widens_Cons [simp]: "G\(S#Ss)[\](T#Ts) = (G\S\T \ G\Ss[\]Ts)" +apply (unfold widens_def) +apply auto +done + + +section "narrowing relation" + +(* all properties of narrowing and casting conversions we actually need *) +(* these can easily be proven from the definitions below *) +(* +rules + cast_RefT2 "G\S\? RefT R \ \t. S=RefT t" + cast_PrimT2 "G\S\? PrimT pt \ \t. S=PrimT t \ G\PrimT t\PrimT pt" +*) + +(* more detailed than necessary for type-safety, see above rules. *) +inductive "narrow G" intros (* narrowing reference conversion, cf. 5.1.5 *) + + subcls: "G\C\\<^sub>C D \ G\ Class D\Class C" + implmt: "\G\C\I \ G\ Class C\Iface I" + obj_arr: "G\Class Object\T.[]" + int_cls: "G\ Iface I\Class C" + subint: "imethds G I hidings imethds G J entails + (\(md, mh ) (md',mh').\G\mrt mh\mrt mh') \ + \G\I\I J \\\\ G\ Iface I\Iface J" + array: "G\RefT S\RefT T \\ G\ RefT S.[]\RefT T.[]" + +(*unused*) +lemma narrow_RefT: "G\RefT R\T \ \t. T=RefT t" +apply (ind_cases "G\S\T") +by auto + +lemma narrow_RefT2: "G\S\RefT R \ \t. S=RefT t" +apply (ind_cases "G\S\T") +by auto + +(*unused*) +lemma narrow_PrimT: "G\PrimT pt\T \ \t. T=PrimT t" +apply (ind_cases "G\S\T") +by auto + +lemma narrow_PrimT2: "G\S\PrimT pt \ + \t. S=PrimT t \ G\PrimT t\PrimT pt" +apply (ind_cases "G\S\T") +by auto + + +section "casting relation" + +inductive "cast G" intros (* casting conversion, cf. 5.5 *) + + widen: "G\S\T \ G\S\? T" + narrow: "G\S\T \ G\S\? T" + +(* +lemma ??unknown??: "\G\S\T; G\S\T\ \ R" + deferred *) + +(*unused*) +lemma cast_RefT: "G\RefT R\? T \ \t. T=RefT t" +apply (ind_cases "G\S\? T") +by (auto dest: widen_RefT narrow_RefT) + +lemma cast_RefT2: "G\S\? RefT R \ \t. S=RefT t" +apply (ind_cases "G\S\? T") +by (auto dest: widen_RefT2 narrow_RefT2) + +(*unused*) +lemma cast_PrimT: "G\PrimT pt\? T \ \t. T=PrimT t" +apply (ind_cases "G\S\? T") +by (auto dest: widen_PrimT narrow_PrimT) + +lemma cast_PrimT2: "G\S\? PrimT pt \ \t. S=PrimT t \ G\PrimT t\PrimT pt" +apply (ind_cases "G\S\? T") +by (auto dest: widen_PrimT2 narrow_PrimT2) + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/TypeSafe.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/TypeSafe.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,1211 @@ +(* Title: isabelle/Bali/TypeSafe.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* The type soundness proof for Java *} + + +theory TypeSafe = Eval + WellForm + Conform: + +section "result conformance" + +constdefs + assign_conforms :: "st \ (val \ state \ state) \ ty \ env_ \ bool" + ("_\|_\_\\_" [71,71,71,71] 70) + "s\|f\T\\E \ + \s' w. Norm s'\\E \ fst E,s'\w\\T \ s\|s' \ assign f w (Norm s')\\E" + + rconf :: "prog \ lenv \ st \ term \ vals \ tys \ bool" + ("_,_,_\_\_\\_" [71,71,71,71,71,71] 70) + "G,L,s\t\v\\T + \ case T of + Inl T \ if (\vf. t=In2 vf) + then G,s\fst (the_In2 v)\\T \ s\|snd (the_In2 v)\T\\(G,L) + else G,s\the_In1 v\\T + | Inr Ts \ list_all2 (conf G s) (the_In3 v) Ts" + +lemma rconf_In1 [simp]: + "G,L,s\In1 ec\In1 v \\Inl T = G,s\v\\T" +apply (unfold rconf_def) +apply (simp (no_asm)) +done + +lemma rconf_In2 [simp]: + "G,L,s\In2 va\In2 vf\\Inl T = (G,s\fst vf\\T \ s\|snd vf\T\\(G,L))" +apply (unfold rconf_def) +apply (simp (no_asm)) +done + +lemma rconf_In3 [simp]: + "G,L,s\In3 es\In3 vs\\Inr Ts = list_all2 (\v T. G,s\v\\T) vs Ts" +apply (unfold rconf_def) +apply (simp (no_asm)) +done + +section "fits and conf" + +(* unused *) +lemma conf_fits: "G,s\v\\T \ G,s\v fits T" +apply (unfold fits_def) +apply clarify +apply (erule swap, simp (no_asm_use)) +apply (drule conf_RefTD) +apply auto +done + +lemma fits_conf: + "\G,s\v\\T; G\T\? T'; G,s\v fits T'; ws_prog G\ \ G,s\v\\T'" +apply (auto dest!: fitsD cast_PrimT2 cast_RefT2) +apply (force dest: conf_RefTD intro: conf_AddrI) +done + +lemma fits_Array: + "\G,s\v\\T; G\T'.[]\T.[]; G,s\v fits T'; ws_prog G\ \ G,s\v\\T'" +apply (auto dest!: fitsD widen_ArrayPrimT widen_ArrayRefT) +apply (force dest: conf_RefTD intro: conf_AddrI) +done + + + +section "gext" + +lemma halloc_gext: "\s1 s2. G\s1 \halloc oi\a\ s2 \ snd s1\|snd s2" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (erule halloc.induct) +apply (auto dest!: new_AddrD) +done + +lemma sxalloc_gext: "\s1 s2. G\s1 \sxalloc\ s2 \ snd s1\|snd s2" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (erule sxalloc.induct) +apply (auto dest!: halloc_gext) +done + +lemma eval_gext_lemma [rule_format (no_asm)]: + "G\s \t\\ (w,s') \ snd s\|snd s' \ (case w of + In1 v \ True + | In2 vf \ normal s \ (\v x s. s\|snd (assign (snd vf) v (x,s))) + | In3 vs \ True)" +apply (erule eval_induct) +prefer 24 + apply (case_tac "inited C (globs s0)", clarsimp, erule thin_rl) (* Init *) +apply (auto del: conjI dest!: not_initedD gext_new sxalloc_gext halloc_gext + simp add: lvar_def fvar_def2 avar_def2 init_lvars_def2 + split del: split_if_asm split add: sum3.split) +(* 6 subgoals *) +apply force+ +done + +lemma evar_gext_f: + "G\Norm s1 \e=\vf \ s2 \ s\|snd (assign (snd vf) v (x,s))" +apply (drule eval_gext_lemma [THEN conjunct2]) +apply auto +done + +lemmas eval_gext = eval_gext_lemma [THEN conjunct1] + +lemma eval_gext': "G\(x1,s1) \t\\ (w,x2,s2) \ s1\|s2" +apply (drule eval_gext) +apply auto +done + +lemma init_yields_initd: "G\Norm s1 \Init C\ s2 \ initd C s2" +apply (erule eval_cases , auto split del: split_if_asm) +apply (case_tac "inited C (globs s1)") +apply (clarsimp split del: split_if_asm)+ +apply (drule eval_gext')+ +apply (drule init_class_obj_inited) +apply (erule inited_gext) +apply (simp (no_asm_use)) +done + + +section "Lemmas" + +lemma obj_ty_obj_class1: + "\wf_prog G; is_type G (obj_ty obj)\ \ is_class G (obj_class obj)" +apply (case_tac "tag obj") +apply (auto simp add: obj_ty_def obj_class_def) +done + +lemma oconf_init_obj: + "\wf_prog G; + (case r of Heap a \ is_type G (obj_ty obj) | Stat C \ is_class G C) +\ \ G,s\obj \values:=init_vals (var_tys G (tag obj) r)\\\\r" +apply (auto intro!: oconf_init_obj_lemma unique_fields) +done + +(* +lemma obj_split: "P obj = (\ oi vs. obj = \tag=oi,values=vs\ \ ?P \tag=oi,values=vs\)" +apply auto +apply (case_tac "obj") +apply auto +*) + +lemma conforms_newG: "\globs s oref = None; (x, s)\\(G,L); + wf_prog G; case oref of Heap a \ is_type G (obj_ty \tag=oi,values=vs\) + | Stat C \ is_class G C\ \ + (x, init_obj G oi oref s)\\(G, L)" +apply (unfold init_obj_def) +apply (auto elim!: conforms_gupd dest!: oconf_init_obj + simp add: obj.update_defs) +done + +lemma conforms_init_class_obj: + "\(x,s)\\(G, L); wf_prog G; class G C=Some y; \ inited C (globs s)\ \ + (x,init_class_obj G C s)\\(G, L)" +apply (rule not_initedD [THEN conforms_newG]) +apply (auto) +done + + +lemma fst_init_lvars[simp]: + "fst (init_lvars G C sig (invmode m e) a' pvs (x,s)) = + (if static m then x else (np a') x)" +apply (simp (no_asm) add: init_lvars_def2) +done + + +lemma halloc_conforms: "\s1. \G\s1 \halloc oi\a\ s2; wf_prog G; s1\\(G, L); + is_type G (obj_ty \tag=oi,values=fs\)\ \ s2\\(G, L)" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (case_tac "aa") +apply (auto elim!: halloc_elim_cases dest!: new_AddrD + intro!: conforms_newG [THEN conforms_xconf] conf_AddrI) +done + +lemma halloc_type_sound: "\s1. \G\s1 \halloc oi\a\ (x,s); wf_prog G; s1\\(G, L); + T = obj_ty \tag=oi,values=fs\; is_type G T\ \ + (x,s)\\(G, L) \ (x = None \ G,s\Addr a\\T)" +apply (auto elim!: halloc_conforms) +apply (case_tac "aa") +apply (subst obj_ty_eq) +apply (auto elim!: halloc_elim_cases dest!: new_AddrD intro!: conf_AddrI) +done + +lemma sxalloc_type_sound: + "\s1 s2. \G\s1 \sxalloc\ s2; wf_prog G\ \ case fst s1 of + None \ s2 = s1 | Some x \ + (\a. fst s2 = Some(Xcpt (Loc a)) \ (\L. s1\\(G,L) \ s2\\(G,L)))" +apply (simp (no_asm_simp) only: split_tupled_all) +apply (erule sxalloc.induct) +apply auto +apply (rule halloc_conforms [THEN conforms_xconf]) +apply (auto elim!: halloc_elim_cases dest!: new_AddrD intro!: conf_AddrI) +done + +lemma wt_init_comp_ty: +"is_acc_type G (pid C) T \ \prg=G,cls=C,lcl=L\\init_comp_ty T\\" +apply (unfold init_comp_ty_def) +apply (clarsimp simp add: accessible_in_RefT_simp + is_acc_type_def is_acc_class_def) +done + + +declare fun_upd_same [simp] + +declare fun_upd_apply [simp del] + + +constdefs + DynT_prop::"[prog,inv_mode,qtname,ref_ty] \ bool" + ("_\_\_\_"[71,71,71,71]70) + "G\mode\D\t \ mode = IntVir \ is_class G D \ + (if (\T. t=ArrayT T) then D=Object else G\Class D\RefT t)" + +lemma DynT_propI: + "\(x,s)\\(G, L); G,s\a'\\RefT statT; wf_prog G; mode = IntVir \ a' \ Null\ + \ G\mode\invocation_class mode s a' statT\statT" +proof (unfold DynT_prop_def) + assume state_conform: "(x,s)\\(G, L)" + and statT_a': "G,s\a'\\RefT statT" + and wf: "wf_prog G" + and mode: "mode = IntVir \ a' \ Null" + let ?invCls = "(invocation_class mode s a' statT)" + let ?IntVir = "mode = IntVir" + let ?Concl = "\invCls. is_class G invCls \ + (if \T. statT = ArrayT T + then invCls = Object + else G\Class invCls\RefT statT)" + show "?IntVir \ ?Concl ?invCls" + proof + assume modeIntVir: ?IntVir + with mode have not_Null: "a' \ Null" .. + from statT_a' not_Null state_conform + obtain a obj + where obj_props: "a' = Addr a" "globs s (Inl a) = Some obj" + "G\obj_ty obj\RefT statT" "is_type G (obj_ty obj)" + by (blast dest: conforms_RefTD) + show "?Concl ?invCls" + proof (cases "tag obj") + case CInst + with modeIntVir obj_props + show ?thesis + by (auto dest!: widen_Array2 split add: split_if) + next + case Arr + from Arr obtain T where "obj_ty obj = T.[]" by (blast dest: obj_ty_Arr1) + moreover from Arr have "obj_class obj = Object" + by (blast dest: obj_class_Arr1) + moreover note modeIntVir obj_props wf + ultimately show ?thesis by (auto dest!: widen_Array ) + qed + qed +qed + +lemma invocation_methd: +"\wf_prog G; statT \ NullT; + (\ statC. statT = ClassT statC \ is_class G statC); + (\ I. statT = IfaceT I \ is_iface G I \ mode \ SuperM); + (\ T. statT = ArrayT T \ mode \ SuperM); + G\mode\invocation_class mode s a' statT\statT; + dynlookup G statT (invocation_class mode s a' statT) sig = Some m \ +\ methd G (invocation_declclass G mode s a' statT sig) sig = Some m" +proof - + assume wf: "wf_prog G" + and not_NullT: "statT \ NullT" + and statC_prop: "(\ statC. statT = ClassT statC \ is_class G statC)" + and statI_prop: "(\ I. statT = IfaceT I \ is_iface G I \ mode \ SuperM)" + and statA_prop: "(\ T. statT = ArrayT T \ mode \ SuperM)" + and invC_prop: "G\mode\invocation_class mode s a' statT\statT" + and dynlookup: "dynlookup G statT (invocation_class mode s a' statT) sig + = Some m" + show ?thesis + proof (cases statT) + case NullT + with not_NullT show ?thesis by simp + next + case IfaceT + with statI_prop obtain I + where statI: "statT = IfaceT I" and + is_iface: "is_iface G I" and + not_SuperM: "mode \ SuperM" by blast + + show ?thesis + proof (cases mode) + case Static + with wf dynlookup statI is_iface + show ?thesis + by (auto simp add: invocation_declclass_def dynlookup_def + dynimethd_def dynmethd_C_C + intro: dynmethd_declclass + dest!: wf_imethdsD + dest: table_of_map_SomeI + split: split_if_asm) + next + case SuperM + with not_SuperM show ?thesis .. + next + case IntVir + with wf dynlookup IfaceT invC_prop show ?thesis + by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def + DynT_prop_def + intro: methd_declclass dynmethd_declclass + split: split_if_asm) + qed + next + case ClassT + show ?thesis + proof (cases mode) + case Static + with wf ClassT dynlookup statC_prop + show ?thesis by (auto simp add: invocation_declclass_def dynlookup_def + intro: dynmethd_declclass) + next + case SuperM + with wf ClassT dynlookup statC_prop + show ?thesis by (auto simp add: invocation_declclass_def dynlookup_def + intro: dynmethd_declclass) + next + case IntVir + with wf ClassT dynlookup statC_prop invC_prop + show ?thesis + by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def + DynT_prop_def + intro: dynmethd_declclass) + qed + next + case ArrayT + show ?thesis + proof (cases mode) + case Static + with wf ArrayT dynlookup show ?thesis + by (auto simp add: invocation_declclass_def dynlookup_def + dynimethd_def dynmethd_C_C + intro: dynmethd_declclass + dest: table_of_map_SomeI) + next + case SuperM + with ArrayT statA_prop show ?thesis by blast + next + case IntVir + with wf ArrayT dynlookup invC_prop show ?thesis + by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def + DynT_prop_def dynmethd_C_C + intro: dynmethd_declclass + dest: table_of_map_SomeI) + qed + qed +qed + +declare split_paired_All [simp del] split_paired_Ex [simp del] +ML_setup {* +simpset_ref() := simpset() delloop "split_all_tac"; +claset_ref () := claset () delSWrapper "split_all_tac" +*} +lemma DynT_mheadsD: +"\G\invmode (mhd sm) e\invC\statT; + wf_prog G; \prg=G,cls=C,lcl=L\\e\-RefT statT; + sm \ mheads G C statT sig; + invC = invocation_class (invmode (mhd sm) e) s a' statT; + declC =invocation_declclass G (invmode (mhd sm) e) s a' statT sig + \ \ + \ dm. + methd G declC sig = Some dm \ G\resTy (mthd dm)\resTy (mhd sm) \ + wf_mdecl G declC (sig, mthd dm) \ + declC = declclass dm \ + is_static dm = is_static sm \ + is_class G invC \ is_class G declC \ G\invC\\<^sub>C declC \ + (if invmode (mhd sm) e = IntVir + then (\ statC. statT=ClassT statC \ G\invC \\<^sub>C statC) + else ( (\ statC. statT=ClassT statC \ G\statC\\<^sub>C declC) + \ (\ statC. statT\ClassT statC \ declC=Object)) \ + (declrefT sm) = ClassT (declclass dm))" +proof - + assume invC_prop: "G\invmode (mhd sm) e\invC\statT" + and wf: "wf_prog G" + and wt_e: "\prg=G,cls=C,lcl=L\\e\-RefT statT" + and sm: "sm \ mheads G C statT sig" + and invC: "invC = invocation_class (invmode (mhd sm) e) s a' statT" + and declC: "declC = + invocation_declclass G (invmode (mhd sm) e) s a' statT sig" + from wt_e wf have type_statT: "is_type G (RefT statT)" + by (auto dest: ty_expr_is_type) + from sm have not_Null: "statT \ NullT" by auto + from type_statT + have wf_C: "(\ statC. statT = ClassT statC \ is_class G statC)" + by (auto) + from type_statT wt_e + have wf_I: "(\I. statT = IfaceT I \ is_iface G I \ + invmode (mhd sm) e \ SuperM)" + by (auto dest: invocationTypeExpr_noClassD) + from wt_e + have wf_A: "(\ T. statT = ArrayT T \ invmode (mhd sm) e \ SuperM)" + by (auto dest: invocationTypeExpr_noClassD) + show ?thesis + proof (cases "invmode (mhd sm) e = IntVir") + case True + with invC_prop not_Null + have invC_prop': " is_class G invC \ + (if (\T. statT=ArrayT T) then invC=Object + else G\Class invC\RefT statT)" + by (auto simp add: DynT_prop_def) + from True + have "\ is_static sm" + by (simp add: invmode_IntVir_eq) + with invC_prop' not_Null + have "G,statT \ invC valid_lookup_cls_for (is_static sm)" + by (cases statT) auto + with sm wf type_statT obtain dm where + dm: "dynlookup G statT invC sig = Some dm" and + resT_dm: "G\resTy (mthd dm)\resTy (mhd sm)" and + static: "is_static dm = is_static sm" + by - (drule dynamic_mheadsD,auto) + with declC invC not_Null + have declC': "declC = (declclass dm)" + by (auto simp add: invocation_declclass_def) + with wf invC declC not_Null wf_C wf_I wf_A invC_prop dm + have dm': "methd G declC sig = Some dm" + by - (drule invocation_methd,auto) + from wf dm invC_prop' declC' type_statT + have declC_prop: "G\invC \\<^sub>C declC \ is_class G declC" + by (auto dest: dynlookup_declC) + from wf dm' declC_prop declC' + have wf_dm: "wf_mdecl G declC (sig,(mthd dm))" + by (auto dest: methd_wf_mdecl) + from invC_prop' + have statC_prop: "(\ statC. statT=ClassT statC \ G\invC \\<^sub>C statC)" + by auto + from True dm' resT_dm wf_dm invC_prop' declC_prop statC_prop declC' static + show ?thesis by auto + next + case False + with type_statT wf invC not_Null wf_I wf_A + have invC_prop': "is_class G invC \ + ((\ statC. statT=ClassT statC \ invC=statC) \ + (\ statC. statT\ClassT statC \ invC=Object)) " + by (case_tac "statT") (auto simp add: invocation_class_def + split: inv_mode.splits) + with not_Null wf + have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig" + by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_C_C + dynimethd_def) + from sm wf wt_e not_Null False invC_prop' obtain "dm" where + dm: "methd G invC sig = Some dm" and + eq_declC_sm_dm:"declrefT sm = ClassT (declclass dm)" and + eq_mheads:"mhd sm=mhead (mthd dm) " + by - (drule static_mheadsD, auto dest: accmethd_SomeD) + then have static: "is_static dm = is_static sm" by - (case_tac "sm",auto) + with declC invC dynlookup_static dm + have declC': "declC = (declclass dm)" + by (auto simp add: invocation_declclass_def) + from invC_prop' wf declC' dm + have dm': "methd G declC sig = Some dm" + by (auto intro: methd_declclass) + from wf dm invC_prop' declC' type_statT + have declC_prop: "G\invC \\<^sub>C declC \ is_class G declC" + by (auto dest: methd_declC ) + then have declC_prop1: "invC=Object \ declC=Object" by auto + from wf dm' declC_prop declC' + have wf_dm: "wf_mdecl G declC (sig,(mthd dm))" + by (auto dest: methd_wf_mdecl) + from invC_prop' declC_prop declC_prop1 + have statC_prop: "( (\ statC. statT=ClassT statC \ G\statC\\<^sub>C declC) + \ (\ statC. statT\ClassT statC \ declC=Object))" + by auto + from False dm' wf_dm invC_prop' declC_prop statC_prop declC' + eq_declC_sm_dm eq_mheads static + show ?thesis by auto + qed +qed + +(* +lemma DynT_mheadsD: +"\G\invmode (mhd sm) e\invC\statT; + wf_prog G; \prg=G,cls=C,lcl=L\\e\-RefT statT; + sm \ mheads G C statT sig; + invC = invocation_class (invmode (mhd sm) e) s a' statT; + declC =invocation_declclass G (invmode (mhd sm) e) s a' statT sig + \ \ + \ dm. + methd G declC sig = Some dm \ G\resTy (mthd dm)\resTy (mhd sm) \ + wf_mdecl G declC (sig, mthd dm) \ + is_class G invC \ is_class G declC \ G\invC\\<^sub>C declC \ + (if invmode (mhd sm) e = IntVir + then (\ statC. statT=ClassT statC \ G\invC \\<^sub>C statC) + else (\ statC. statT=ClassT statC \ G\statC \\<^sub>C declC) \ + (declrefT sm) = ClassT (declclass dm))" +proof - + assume invC_prop: "G\invmode (mhd sm) e\invC\statT" + and wf: "wf_prog G" + and wt_e: "\prg=G,cls=C,lcl=L\\e\-RefT statT" + and sm: "sm \ mheads G C statT sig" + and invC: "invC = invocation_class (invmode (mhd sm) e) s a' statT" + and declC: "declC = + invocation_declclass G (invmode (mhd sm) e) s a' statT sig" + from wt_e wf have type_statT: "is_type G (RefT statT)" + by (auto dest: ty_expr_is_type) + from sm have not_Null: "statT \ NullT" by auto + from type_statT + have wf_C: "(\ statC. statT = ClassT statC \ is_class G statC)" + by (auto) + from type_statT wt_e + have wf_I: "(\I. statT = IfaceT I \ is_iface G I \ + invmode (mhd sm) e \ SuperM)" + by (auto dest: invocationTypeExpr_noClassD) + from wt_e + have wf_A: "(\ T. statT = ArrayT T \ invmode (mhd sm) e \ SuperM)" + by (auto dest: invocationTypeExpr_noClassD) + show ?thesis + proof (cases "invmode (mhd sm) e = IntVir") + case True + with invC_prop not_Null + have invC_prop': "is_class G invC \ + (if (\T. statT=ArrayT T) then invC=Object + else G\Class invC\RefT statT)" + by (auto simp add: DynT_prop_def) + with sm wf type_statT not_Null obtain dm where + dm: "dynlookup G statT invC sig = Some dm" and + resT_dm: "G\resTy (mthd dm)\resTy (mhd sm)" + by - (clarify,drule dynamic_mheadsD,auto) + with declC invC not_Null + have declC': "declC = (declclass dm)" + by (auto simp add: invocation_declclass_def) + with wf invC declC not_Null wf_C wf_I wf_A invC_prop dm + have dm': "methd G declC sig = Some dm" + by - (drule invocation_methd,auto) + from wf dm invC_prop' declC' type_statT + have declC_prop: "G\invC \\<^sub>C declC \ is_class G declC" + by (auto dest: dynlookup_declC) + from wf dm' declC_prop declC' + have wf_dm: "wf_mdecl G declC (sig,(mthd dm))" + by (auto dest: methd_wf_mdecl) + from invC_prop' + have statC_prop: "(\ statC. statT=ClassT statC \ G\invC \\<^sub>C statC)" + by auto + from True dm' resT_dm wf_dm invC_prop' declC_prop statC_prop + show ?thesis by auto + next + case False + + with type_statT wf invC not_Null wf_I wf_A + have invC_prop': "is_class G invC \ + ((\ statC. statT=ClassT statC \ invC=statC) \ + (\ statC. statT\ClassT statC \ invC=Object)) " + + by (case_tac "statT") (auto simp add: invocation_class_def + split: inv_mode.splits) + with not_Null + have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig" + by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_def + dynimethd_def) + from sm wf wt_e not_Null False invC_prop' obtain "dm" where + dm: "methd G invC sig = Some dm" and + eq_declC_sm_dm:"declrefT sm = ClassT (declclass dm)" and + eq_mheads:"mhd sm=mhead (mthd dm) " + by - (drule static_mheadsD, auto dest: accmethd_SomeD) + with declC invC dynlookup_static dm + have declC': "declC = (declclass dm)" + by (auto simp add: invocation_declclass_def) + from invC_prop' wf declC' dm + have dm': "methd G declC sig = Some dm" + by (auto intro: methd_declclass) + from wf dm invC_prop' declC' type_statT + have declC_prop: "G\invC \\<^sub>C declC \ is_class G declC" + by (auto dest: methd_declC ) + from wf dm' declC_prop declC' + have wf_dm: "wf_mdecl G declC (sig,(mthd dm))" + by (auto dest: methd_wf_mdecl) + from invC_prop' declC_prop + have statC_prop: "(\ statC. statT=ClassT statC \ G\statC \\<^sub>C declC)" + by auto + from False dm' wf_dm invC_prop' declC_prop statC_prop + eq_declC_sm_dm eq_mheads + show ?thesis by auto + qed +qed +*) + +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] +ML_setup {* +simpset_ref() := simpset() delloop "split_all_tac"; +claset_ref () := claset () delSWrapper "split_all_tac" +*} + +lemma DynT_conf: "\G\invocation_class mode s a' statT \\<^sub>C declC; wf_prog G; + isrtype G (statT); + G,s\a'\\RefT statT; mode = IntVir \ a' \ Null; + mode \ IntVir \ (\ statC. statT=ClassT statC \ G\statC\\<^sub>C declC) + \ (\ statC. statT\ClassT statC \ declC=Object)\ + \G,s\a'\\ Class declC" +apply (case_tac "mode = IntVir") +apply (drule conf_RefTD) +apply (force intro!: conf_AddrI + simp add: obj_class_def split add: obj_tag.split_asm) +apply clarsimp +apply safe +apply (erule (1) widen.subcls [THEN conf_widen]) +apply (erule wf_ws_prog) + +apply (frule widen_Object) apply (erule wf_ws_prog) +apply (erule (1) conf_widen) apply (erule wf_ws_prog) +done + + +lemma Ass_lemma: + "\G\Norm s0 \va=\(w, f)\ Norm s1; G\Norm s1 \e-\v\ Norm s2; G,s2\v\\T'; + s1\|s2 \ assign f v (Norm s2)\\(G, L) + \ \ assign f v (Norm s2)\\(G, L) \ + (\(x',s'). x' = None \ G,s'\v\\T') (assign f v (Norm s2))" +apply (drule_tac x = "None" and s = "s2" and v = "v" in evar_gext_f) +apply (drule eval_gext', clarsimp) +apply (erule conf_gext) +apply simp +done + +lemma Throw_lemma: "\G\tn\\<^sub>C SXcpt Throwable; wf_prog G; (x1,s1)\\(G, L); + x1 = None \ G,s1\a'\\ Class tn\ \ (throw a' x1, s1)\\(G, L)" +apply (auto split add: split_abrupt_if simp add: throw_def2) +apply (erule conforms_xconf) +apply (frule conf_RefTD) +apply (auto elim: widen.subcls [THEN conf_widen]) +done + +lemma Try_lemma: "\G\obj_ty (the (globs s1' (Heap a)))\ Class tn; + (Some (Xcpt (Loc a)), s1')\\(G, L); wf_prog G\ + \ Norm (lupd(vn\Addr a) s1')\\(G, L(vn\Class tn))" +apply (rule conforms_allocL) +apply (erule conforms_NormI) +apply (drule conforms_XcptLocD [THEN conf_RefTD],rule HOL.refl) +apply (auto intro!: conf_AddrI) +done + +lemma Fin_lemma: +"\G\Norm s1 \c2\ (x2,s2); wf_prog G; (Some a, s1)\\(G, L); (x2,s2)\\(G, L)\ +\ (abrupt_if True (Some a) x2, s2)\\(G, L)" +apply (force elim: eval_gext' conforms_xgext split add: split_abrupt_if) +done + +lemma FVar_lemma1: "\table_of (DeclConcepts.fields G Ca) (fn, C) = Some f ; + x2 = None \ G,s2\a\\ Class Ca; wf_prog G; G\Ca\\<^sub>C C; C \ Object; + class G C = Some y; (x2,s2)\\(G, L); s1\|s2; inited C (globs s1); + (if static f then id else np a) x2 = None\ + \ + \obj. globs s2 (if static f then Inr C else Inl (the_Addr a)) = Some obj \ + var_tys G (tag obj) (if static f then Inr C else Inl(the_Addr a)) + (Inl(fn,C)) = Some (type f)" +apply (drule initedD) +apply (frule subcls_is_class2, simp (no_asm_simp)) +apply (case_tac "static f") +apply clarsimp +apply (drule (1) rev_gext_objD, clarsimp) +apply (frule fields_declC, erule wf_ws_prog, simp (no_asm_simp)) +apply (rule var_tys_Some_eq [THEN iffD2]) +apply clarsimp +apply (erule fields_table_SomeI, simp (no_asm)) +apply clarsimp +apply (drule conf_RefTD, clarsimp, rule var_tys_Some_eq [THEN iffD2]) +apply (auto dest!: widen_Array split add: obj_tag.split) +apply (rotate_tac -1) (* to front: tag obja = CInst pid_field_type to enable + conditional rewrite *) +apply (rule fields_table_SomeI) +apply (auto elim!: fields_mono subcls_is_class2) +done + +lemma FVar_lemma: +"\((v, f), Norm s2') = fvar C (static field) fn a (x2, s2); G\Ca\\<^sub>C C; + table_of (DeclConcepts.fields G Ca) (fn, C) = Some field; wf_prog G; + x2 = None \ G,s2\a\\Class Ca; C \ Object; class G C = Some y; + (x2, s2)\\(G, L); s1\|s2; inited C (globs s1)\ \ + G,s2'\v\\type field \ s2'\|f\type field\\(G, L)" +apply (unfold assign_conforms_def) +apply (drule sym) +apply (clarsimp simp add: fvar_def2) +apply (drule (9) FVar_lemma1) +apply (clarsimp) +apply (drule (2) conforms_globsD [THEN oconf_lconf, THEN lconfD]) +apply clarsimp +apply (drule (1) rev_gext_objD) +apply (auto elim!: conforms_upd_gobj) +done + + +lemma AVar_lemma1: "\globs s (Inl a) = Some obj;tag obj=Arr ty i; + the_Intg i' in_bounds i; wf_prog G; G\ty.[]\Tb.[]; Norm s\\(G, L) +\ \ G,s\the ((values obj) (Inr (the_Intg i')))\\Tb" +apply (erule widen_Array_Array [THEN conf_widen]) +apply (erule_tac [2] wf_ws_prog) +apply (drule (1) conforms_globsD [THEN oconf_lconf, THEN lconfD]) +defer apply assumption +apply (force intro: var_tys_Some_eq [THEN iffD2]) +done + +lemma obj_split: "\ obj. \ t vs. obj = \tag=t,values=vs\" +proof record_split + fix tag values more + show "\t vs. \tag = tag, values = values, \ = more\ = \tag = t, values = vs\" + by auto +qed + +lemma AVar_lemma: "\wf_prog G; G\(x1, s1) \e2-\i\ (x2, s2); + ((v,f), Norm s2') = avar G i a (x2, s2); x1 = None \ G,s1\a\\Ta.[]; + (x2, s2)\\(G, L); s1\|s2\ \ G,s2'\v\\Ta \ s2'\|f\Ta\\(G, L)" +apply (unfold assign_conforms_def) +apply (drule sym) +apply (clarsimp simp add: avar_def2) +apply (drule (1) conf_gext) +apply (drule conf_RefTD, clarsimp) +apply (subgoal_tac "\ t vs. obj = \tag=t,values=vs\") +defer +apply (rule obj_split) +apply clarify +apply (frule obj_ty_widenD) +apply (auto dest!: widen_Class) +apply (force dest: AVar_lemma1) +apply (auto split add: split_if) +apply (force elim!: fits_Array dest: gext_objD + intro: var_tys_Some_eq [THEN iffD2] conforms_upd_gobj) +done + + +section "Call" +lemma conforms_init_lvars_lemma: "\wf_prog G; + wf_mhead G P sig mh; + Ball (set lvars) (split (\vn. is_type G)); + list_all2 (conf G s) pvs pTsa; G\pTsa[\](parTs sig)\ \ + G,s\init_vals (table_of lvars)(pars mh[\]pvs) + [\\]table_of lvars(pars mh[\]parTs sig)" +apply (unfold wf_mhead_def) +apply clarify +apply (erule (2) Ball_set_table [THEN lconf_init_vals, THEN lconf_ext_list]) +apply (drule wf_ws_prog) +apply (erule (2) conf_list_widen) +done + + +lemma lconf_map_lname [simp]: + "G,s\(lname_case l1 l2)[\\](lname_case L1 L2) + = + (G,s\l1[\\]L1 \ G,s\(\x::unit . l2)[\\](\x::unit. L2))" +apply (unfold lconf_def) +apply safe +apply (case_tac "n") +apply (force split add: lname.split)+ +done + +lemma lconf_map_ename [simp]: + "G,s\(ename_case l1 l2)[\\](ename_case L1 L2) + = + (G,s\l1[\\]L1 \ G,s\(\x::unit. l2)[\\](\x::unit. L2))" +apply (unfold lconf_def) +apply safe +apply (force split add: ename.split)+ +done + + +lemma defval_conf1 [rule_format (no_asm), elim]: + "is_type G T \ (\v\Some (default_val T): G,s\v\\T)" +apply (unfold conf_def) +apply (induct "T") +apply (auto intro: prim_ty.induct) +done + + +lemma conforms_init_lvars: +"\wf_mhead G (pid declC) sig (mhead (mthd dm)); wf_prog G; + list_all2 (conf G s) pvs pTsa; G\pTsa[\](parTs sig); + (x, s)\\(G, L); + methd G declC sig = Some dm; + isrtype G statT; + G\invC\\<^sub>C declC; + G,s\a'\\RefT statT; + invmode (mhd sm) e = IntVir \ a' \ Null; + invmode (mhd sm) e \ IntVir \ + (\ statC. statT=ClassT statC \ G\statC\\<^sub>C declC) + \ (\ statC. statT\ClassT statC \ declC=Object); + invC = invocation_class (invmode (mhd sm) e) s a' statT; + declC = invocation_declclass G (invmode (mhd sm) e) s a' statT sig; + Ball (set (lcls (mbody (mthd dm)))) + (split (\vn. is_type G)) + \ \ + init_lvars G declC sig (invmode (mhd sm) e) a' + pvs (x,s)\\(G,\ k. + (case k of + EName e \ (case e of + VNam v + \ (table_of (lcls (mbody (mthd dm))) + (pars (mthd dm)[\]parTs sig)) v + | Res \ Some (resTy (mthd dm))) + | This \ if (static (mthd sm)) + then None else Some (Class declC)))" +apply (simp add: init_lvars_def2) +apply (rule conforms_set_locals) +apply (simp (no_asm_simp) split add: split_if) +apply (drule (4) DynT_conf) +apply clarsimp +(* apply intro *) +apply (drule (4) conforms_init_lvars_lemma) +apply (case_tac "dm",simp) +apply (rule conjI) +apply (unfold lconf_def, clarify) +apply (rule defval_conf1) +apply (clarsimp simp add: wf_mhead_def is_acc_type_def) +apply (case_tac "is_static sm") +apply simp_all +done + + +lemma Call_type_sound: "\wf_prog G; G\(x1, s1) \ps\\pvs\ (x2, s2); + declC + = invocation_declclass G (invmode (mhd esm) e) s2 a' statT \name=mn,parTs=pTs\; +s2'=init_lvars G declC \name=mn,parTs=pTs\ (invmode (mhd esm) e) a' pvs (x2,s2); + G\s2' \Methd declC \name=mn,parTs=pTs\-\v\ (x3, s3); + \L. s2'\\(G, L) + \ (\T. \prg=G,cls=declC,lcl=L\\ Methd declC \name=mn,parTs=pTs\\-T + \ (x3, s3)\\(G, L) \ (x3 = None \ G,s3\v\\T)); + Norm s0\\(G, L); + \prg=G,cls=C,lcl=L\\e\-RefT statT; \prg=G,cls=C,lcl=L\\ps\\pTsa; + max_spec G C statT \name=mn,parTs=pTsa\ = {(esm, pTs)}; + (x1, s1)\\(G, L); + x1 = None \ G,s1\a'\\RefT statT; (x2, s2)\\(G, L); + x2 = None \ list_all2 (conf G s2) pvs pTsa; + sm=(mhd esm)\ \ + (x3, set_locals (locals s2) s3)\\(G, L) \ + (x3 = None \ G,s3\v\\resTy sm)" +apply clarify +apply (case_tac "x2") +defer +apply (clarsimp split add: split_if_asm simp add: init_lvars_def2) +apply (case_tac "a' = Null \ \ (static (mhd esm)) \ e \ Super") +apply (clarsimp simp add: init_lvars_def2) +apply clarsimp +apply (drule eval_gext') +apply (frule (1) conf_gext) +apply (drule max_spec2mheads, clarsimp) +apply (subgoal_tac "invmode (mhd esm) e = IntVir \ a' \ Null") +defer +apply (clarsimp simp add: invmode_IntVir_eq) +apply (frule (6) DynT_mheadsD [OF DynT_propI,of _ "s2"],(rule HOL.refl)+) +apply clarify +apply (drule wf_mdeclD1, clarsimp) +apply (frule ty_expr_is_type) apply simp +apply (frule (2) conforms_init_lvars) +apply simp +apply assumption+ +apply simp +apply assumption+ +apply clarsimp +apply (rule HOL.refl) +apply simp +apply (rule Ball_weaken) +apply assumption +apply (force simp add: is_acc_type_def) +apply (tactic "smp_tac 1 1") +apply (frule (2) wt_MethdI, clarsimp) +apply (subgoal_tac "is_static dm = (static (mthd esm))") +apply (simp only:) +apply (tactic "smp_tac 1 1") +apply (rule conjI) +apply (erule conforms_return) +apply blast + +apply (force dest!: eval_gext del: impCE simp add: init_lvars_def2) +apply clarsimp +apply (drule (2) widen_trans, erule (1) conf_widen) +apply (erule wf_ws_prog) + +apply auto +done + + +subsection "accessibility" + +theorem dynamic_field_access_ok: + (assumes wf: "wf_prog G" and + eval_e: "G\s1 \e-\a\ s2" and + not_Null: "a\Null" and + conform_a: "G,(store s2)\a\\ Class statC" and + conform_s2: "s2\\(G, L)" and + normal_s2: "normal s2" and + wt_e: "\prg=G,cls=accC,lcl=L\,dt\e\-Class statC" and + f: "accfield G accC statC fn = Some f" and + dynC: "if stat then dynC=statC + else dynC=obj_class (lookup_obj (store s2) a)" + ) "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f) \ + G\Field fn f in dynC dyn_accessible_from accC" +proof (cases "stat") + case True + with dynC + have dynC': "dynC=statC" by simp + with f + have "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)" + by (auto simp add: accfield_def Let_def intro!: table_of_remap_SomeD) + with dynC' f + show ?thesis + by (auto intro!: static_to_dynamic_accessible_from + dest: accfield_accessibleD accessible_from_commonD) +next + case False + with wf conform_a not_Null conform_s2 dynC + obtain subclseq: "G\dynC \\<^sub>C statC" and + "is_class G dynC" + by (auto dest!: conforms_RefTD [of _ _ _ _ "(fst s2)" L] + dest: obj_ty_obj_class1 + simp add: obj_ty_obj_class ) + with wf f + have "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)" + by (auto simp add: accfield_def Let_def + dest: fields_mono + dest!: table_of_remap_SomeD) + moreover + from f subclseq + have "G\Field fn f in dynC dyn_accessible_from accC" + by (auto intro!: static_to_dynamic_accessible_from + dest: accfield_accessibleD) + ultimately show ?thesis + by blast +qed + +lemma call_access_ok: +(assumes invC_prop: "G\invmode (mhd statM) e\invC\statT" + and wf: "wf_prog G" + and wt_e: "\prg=G,cls=C,lcl=L\\e\-RefT statT" + and statM: "statM \ mheads G accC statT sig" + and invC: "invC = invocation_class (invmode (mhd statM) e) s a statT" +)"\ dynM. dynlookup G statT invC sig = Some dynM \ + G\Methd sig dynM in invC dyn_accessible_from accC" +proof - + from wt_e wf have type_statT: "is_type G (RefT statT)" + by (auto dest: ty_expr_is_type) + from statM have not_Null: "statT \ NullT" by auto + from type_statT wt_e + have wf_I: "(\I. statT = IfaceT I \ is_iface G I \ + invmode (mhd statM) e \ SuperM)" + by (auto dest: invocationTypeExpr_noClassD) + from wt_e + have wf_A: "(\ T. statT = ArrayT T \ invmode (mhd statM) e \ SuperM)" + by (auto dest: invocationTypeExpr_noClassD) + show ?thesis + proof (cases "invmode (mhd statM) e = IntVir") + case True + with invC_prop not_Null + have invC_prop': "is_class G invC \ + (if (\T. statT=ArrayT T) then invC=Object + else G\Class invC\RefT statT)" + by (auto simp add: DynT_prop_def) + with True not_Null + have "G,statT \ invC valid_lookup_cls_for is_static statM" + by (cases statT) (auto simp add: invmode_def + split: split_if split_if_asm) (* was deleted above *) + with statM type_statT wf + show ?thesis + by - (rule dynlookup_access,auto) + next + case False + with type_statT wf invC not_Null wf_I wf_A + have invC_prop': " is_class G invC \ + ((\ statC. statT=ClassT statC \ invC=statC) \ + (\ statC. statT\ClassT statC \ invC=Object)) " + by (case_tac "statT") (auto simp add: invocation_class_def + split: inv_mode.splits) + with not_Null wf + have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig" + by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_C_C + dynimethd_def) + from statM wf wt_e not_Null False invC_prop' obtain dynM where + "accmethd G accC invC sig = Some dynM" + by (auto dest!: static_mheadsD) + from invC_prop' False not_Null wf_I + have "G,statT \ invC valid_lookup_cls_for is_static statM" + by (cases statT) (auto simp add: invmode_def + split: split_if split_if_asm) (* was deleted above *) + with statM type_statT wf + show ?thesis + by - (rule dynlookup_access,auto) + qed +qed + +section "main proof of type safety" + +ML {* +val forward_hyp_tac = EVERY' [smp_tac 1, + FIRST'[mp_tac,etac exI,smp_tac 2,smp_tac 1,EVERY'[etac impE,etac exI]], + REPEAT o (etac conjE)]; +val typD_tac = eresolve_tac (thms "wt_elim_cases") THEN_ALL_NEW + EVERY' [full_simp_tac (simpset() setloop (K no_tac)), + clarify_tac(claset() addSEs[])] +*} + +lemma conforms_locals [rule_format]: + "(a,b)\\(G, L) \ L x = Some T \ G,b\the (locals b x)\\T" +apply (force simp: conforms_def Let_def lconf_def) +done + +lemma eval_type_sound [rule_format (no_asm)]: + "wf_prog G \ G\s0 \t\\ (v,s1) \ (\L. s0\\(G,L) \ + (\C T. \prg=G,cls=C,lcl=L\\t\T \ s1\\(G,L) \ + (let (x,s) = s1 in x = None \ G,L,s\t\v\\T)))" +apply (erule eval_induct) + +(* 29 subgoals *) +(* Xcpt, Inst, Methd, Nil, Skip, Expr, Comp *) +apply (simp_all (no_asm_use) add: Let_def body_def) +apply (tactic "ALLGOALS (EVERY'[Clarify_tac, TRY o typD_tac, + TRY o forward_hyp_tac])") +apply (tactic"ALLGOALS(EVERY'[asm_simp_tac(simpset()),TRY o Clarify_tac])") + +(* 20 subgoals *) + +(* Break *) +apply (erule conforms_absorb) + +(* Cons *) +apply (erule_tac V = "G\Norm s0 \?ea\\ ?vs1" in thin_rl) +apply (frule eval_gext') +apply force + +(* LVar *) +apply (force elim: conforms_localD [THEN lconfD] conforms_lupd + simp add: assign_conforms_def lvar_def) + +(* Cast *) +apply (force dest: fits_conf) + +(* Lit *) +apply (rule conf_litval) +apply (simp add: empty_dt_def) + +(* Super *) +apply (rule conf_widen) +apply (erule (1) subcls_direct [THEN widen.subcls]) +apply (erule (1) conforms_localD [THEN lconfD]) +apply (erule wf_ws_prog) + +(* Acc *) +apply fast + +(* Body *) +apply (rule conjI) +apply (rule conforms_absorb) +apply (fast) +apply (fast intro: conforms_locals) + +(* Cond *) +apply (simp split: split_if_asm) +apply (tactic "forward_hyp_tac 1", force) +apply (tactic "forward_hyp_tac 1", force) + +(* If *) +apply (force split add: split_if_asm) + +(* Loop *) +apply (drule (1) wt.Loop) +apply (clarsimp split: split_if_asm) +apply (fast intro: conforms_absorb) + +(* Fin *) +apply (case_tac "x1", force) +apply (drule spec, erule impE, erule conforms_NormI) +apply (erule impE) +apply blast +apply (clarsimp) +apply (erule (3) Fin_lemma) + +(* Throw *) +apply (erule (3) Throw_lemma) + +(* NewC *) +apply (clarsimp simp add: is_acc_class_def) +apply (drule (1) halloc_type_sound,blast, rule HOL.refl, simp, simp) + +(* NewA *) +apply (tactic "smp_tac 1 1",frule wt_init_comp_ty,erule impE,blast) +apply (tactic "forward_hyp_tac 1") +apply (case_tac "check_neg i' ab") +apply (clarsimp simp add: is_acc_type_def) +apply (drule (2) halloc_type_sound, rule HOL.refl, simp, simp) +apply force + +(* Level 34, 6 subgoals *) + +(* Init *) +apply (case_tac "inited C (globs s0)") +apply (clarsimp) +apply (clarsimp) +apply (frule (1) wf_prog_cdecl) +apply (drule spec, erule impE, erule (3) conforms_init_class_obj) +apply (drule_tac "psi" = "class G C = ?x" in asm_rl,erule impE, + force dest!: wf_cdecl_supD split add: split_if simp add: is_acc_class_def) +apply (drule spec, erule impE, erule conforms_set_locals, rule lconf_empty) +apply (erule impE) apply (rule exI) apply (erule wf_cdecl_wt_init) +apply (drule (1) conforms_return, force dest: eval_gext', assumption) + + +(* Ass *) +apply (tactic "forward_hyp_tac 1") +apply (rename_tac x1 s1 x2 s2 v va w L C Ta T', case_tac x1) +prefer 2 apply force +apply (case_tac x2) +prefer 2 apply force +apply (simp, drule conjunct2) +apply (drule (1) conf_widen) +apply (erule wf_ws_prog) +apply (erule (2) Ass_lemma) +apply (clarsimp simp add: assign_conforms_def) + +(* Try *) +apply (drule (1) sxalloc_type_sound, simp (no_asm_use)) +apply (case_tac a) +apply clarsimp +apply clarsimp +apply (tactic "smp_tac 1 1") +apply (simp split add: split_if_asm) +apply (fast dest: conforms_deallocL Try_lemma) + +(* FVar *) + +apply (frule accfield_fields) +apply (frule ty_expr_is_type [THEN type_is_class],simp) +apply simp +apply (frule wf_ws_prog) +apply (frule (1) fields_declC,simp) +apply clarsimp +(*b y EVERY'[datac cfield_defpl_is_class 2, Clarsimp_tac] 1; not useful here*) +apply (tactic "smp_tac 1 1") +apply (tactic "forward_hyp_tac 1") +apply (rule conjI, force split add: split_if simp add: fvar_def2) +apply (drule init_yields_initd, frule eval_gext') +apply clarsimp +apply (case_tac "C=Object") +apply clarsimp +apply (erule (9) FVar_lemma) + +(* AVar *) +apply (tactic "forward_hyp_tac 1") +apply (erule_tac V = "G\Norm s0 \?e1-\?a'\ (?x1 1, ?s1)" in thin_rl, + frule eval_gext') +apply (rule conjI) +apply (clarsimp simp add: avar_def2) +apply clarsimp +apply (erule (5) AVar_lemma) + +(* Call *) +apply (tactic "forward_hyp_tac 1") +apply (rule Call_type_sound) +apply auto +done + + +declare fun_upd_apply [simp] +declare split_paired_All [simp] split_paired_Ex [simp] +declare split_if [split] split_if_asm [split] + option.split [split] option.split_asm [split] +ML_setup {* +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac); +claset_ref() := claset () addSbefore ("split_all_tac", split_all_tac) +*} + +theorem eval_ts: + "\G\s \e-\v \ (x',s'); wf_prog G; s\\(G,L); \prg=G,cls=C,lcl=L\\e\-T\ +\ (x',s')\\(G,L) \ (x'=None \ G,s'\v\\T)" +apply (drule (3) eval_type_sound) +apply (unfold Let_def) +apply clarsimp +done + +theorem evals_ts: +"\G\s \es\\vs\ (x',s'); wf_prog G; s\\(G,L); \prg=G,cls=C,lcl=L\\es\\Ts\ +\ (x',s')\\(G,L) \ (x'=None \ list_all2 (conf G s') vs Ts)" +apply (drule (3) eval_type_sound) +apply (unfold Let_def) +apply clarsimp +done + +theorem evar_ts: +"\G\s \v=\vf\ (x',s'); wf_prog G; s\\(G,L); \prg=G,cls=C,lcl=L\\v\=T\ \ + (x',s')\\(G,L) \ (x'=None \ G,L,s'\In2 v\In2 vf\\Inl T)" +apply (drule (3) eval_type_sound) +apply (unfold Let_def) +apply clarsimp +done + +theorem exec_ts: +"\G\s \s0\ s'; wf_prog G; s\\(G,L); \prg=G,cls=C,lcl=L\\s0\\\ \ s'\\(G,L)" +apply (drule (3) eval_type_sound) +apply (unfold Let_def) +apply clarsimp +done + +(* +theorem dyn_methods_understood: + "\s. \wf_prog G; \prg=G,cls=C,lcl=L\\{t,md,IntVir}e..mn({pTs'}ps)\-rT; + s\\(G,L); G\s \e-\a'\ Norm s'; a' \ Null\ \ + \a obj. a'=Addr a \ heap s' a = Some obj \ + cmethd G (obj_class obj) (mn, pTs') \ None" +apply (erule wt_elim_cases) +apply (drule max_spec2mheads) +apply (drule (3) eval_ts) +apply (clarsimp split del: split_if split_if_asm) +apply (drule (2) DynT_propI) +apply (simp (no_asm_simp)) +apply (tactic *) (* {* exhaust_cmethd_tac "the (cmethd G (target (invmode m e) s' a' md) (mn, pTs'))" 1 *} *)(*) +apply (drule (4) DynT_mheadsD [THEN conjunct1], rule HOL.refl) +apply (drule conf_RefTD) +apply clarsimp +done +*) + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/Value.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Value.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,51 @@ +(* Title: isabelle/Bali/Value.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* Java values *} + + + +theory Value = Type: + +typedecl loc (* locations, i.e. abstract references on objects *) +arities loc :: "type" + +datatype val + = Unit (* dummy result value of void methods *) + | Bool bool (* Boolean value *) + | Intg int (* integer value *) + | Null (* null reference *) + | Addr loc (* addresses, i.e. locations of objects *) + + +translations "val" <= (type) "Term.val" + "loc" <= (type) "Term.loc" + +consts the_Bool :: "val \ bool" +primrec "the_Bool (Bool b) = b" +consts the_Intg :: "val \ int" +primrec "the_Intg (Intg i) = i" +consts the_Addr :: "val \ loc" +primrec "the_Addr (Addr a) = a" + +types dyn_ty = "loc \ ty option" +consts + typeof :: "dyn_ty \ val \ ty option" + defpval :: "prim_ty \ val" (* default value for primitive types *) + default_val :: " ty \ val" (* default value for all types *) + +primrec "typeof dt Unit = Some (PrimT Void)" + "typeof dt (Bool b) = Some (PrimT Boolean)" + "typeof dt (Intg i) = Some (PrimT Integer)" + "typeof dt Null = Some NT" + "typeof dt (Addr a) = dt a" + +primrec "defpval Void = Unit" + "defpval Boolean = Bool False" + "defpval Integer = Intg 0" +primrec "default_val (PrimT pt) = defpval pt" + "default_val (RefT r ) = Null" + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/WellForm.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/WellForm.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,3004 @@ +(* Title: isabelle/Bali/WellForm.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) + +header {* Well-formedness of Java programs *} + +theory WellForm = WellType: + +text {* +For static checks on expressions and statements, see WellType.thy + +improvements over Java Specification 1.0 (cf. 8.4.6.3, 8.4.6.4, 9.4.1): +\begin{itemize} +\item a method implementing or overwriting another method may have a result + type that widens to the result type of the other method + (instead of identical type) +\item if a method hides another method (both methods have to be static!) + there are no restrictions to the result type + since the methods have to be static and there is no dynamic binding of + static methods +\item if an interface inherits more than one method with the same signature, the + methods need not have identical return types +\end{itemize} +simplifications: +\begin{itemize} +\item Object and standard exceptions are assumed to be declared like normal + classes +\end{itemize} +*} + +section "well-formed field declarations" + (* well-formed field declaration (common part for classes and interfaces), + cf. 8.3 and (9.3) *) + +constdefs + wf_fdecl :: "prog \ pname \ fdecl \ bool" + "wf_fdecl G P \ \(fn,f). is_acc_type G P (type f)" + +lemma wf_fdecl_def2: "\fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))" +apply (unfold wf_fdecl_def) +apply simp +done + + + +section "well-formed method declarations" + (*well-formed method declaration,cf. 8.4, 8.4.1, 8.4.3, 8.4.5, 14.3.2, (9.4)*) + (* cf. 14.15, 15.7.2, for scope issues cf. 8.4.1 and 14.3.2 *) + +text {* +A method head is wellformed if: +\begin{itemize} +\item the signature and the method head agree in the number of parameters +\item all types of the parameters are visible +\item the result type is visible +\item the parameter names are unique +\end{itemize} +*} +constdefs + wf_mhead :: "prog \ pname \ sig \ mhead \ bool" + "wf_mhead G P \ \ sig mh. length (parTs sig) = length (pars mh) \ + \ ( \T\set (parTs sig). is_acc_type G P T) \ + is_acc_type G P (resTy mh) \ + \ nodups (pars mh)" + + +text {* +A method declaration is wellformed if: +\begin{itemize} +\item the method head is wellformed +\item the names of the local variables are unique +\item the types of the local variables must be accessible +\item the local variables don't shadow the parameters +\item the class of the method is defined +\item the body statement is welltyped with respect to the + modified environment of local names, were the local variables, + the parameters the special result variable (Res) and This are assoziated + with there types. +\end{itemize} +*} + +constdefs + wf_mdecl :: "prog \ qtname \ mdecl \ bool" + "wf_mdecl G C \ + \(sig,m). + wf_mhead G (pid C) sig (mhead m) \ + unique (lcls (mbody m)) \ + (\(vn,T)\set (lcls (mbody m)). is_acc_type G (pid C) T) \ + (\pn\set (pars m). table_of (lcls (mbody m)) pn = None) \ + is_class G C \ + \prg=G,cls=C, + lcl=\ k. + (case k of + EName e + \ (case e of + VNam v + \(table_of (lcls (mbody m))((pars m)[\](parTs sig))) v + | Res \ Some (resTy m)) + | This \ if static m then None else Some (Class C)) + \\(stmt (mbody m))\\" + +lemma wf_mheadI: +"\length (parTs sig) = length (pars m); \T\set (parTs sig). is_acc_type G P T; + is_acc_type G P (resTy m); nodups (pars m)\ \ + wf_mhead G P sig m" +apply (unfold wf_mhead_def) +apply (simp (no_asm_simp)) +done + +lemma wf_mdeclI: "\ + wf_mhead G (pid C) sig (mhead m); unique (lcls (mbody m)); + (\pn\set (pars m). table_of (lcls (mbody m)) pn = None); + \(vn,T)\set (lcls (mbody m)). is_acc_type G (pid C) T; + is_class G C; + \prg=G,cls=C, + lcl=\ k. + (case k of + EName e + \ (case e of + VNam v + \ (table_of (lcls (mbody m))((pars m)[\](parTs sig))) v + | Res \ Some (resTy m)) + | This \ if static m then None else Some (Class C)) + \\(stmt (mbody m))\\ + \ \ + wf_mdecl G C (sig,m)" +apply (unfold wf_mdecl_def) +apply simp +done + + +lemma wf_mdeclD1: +"wf_mdecl G C (sig,m) \ + wf_mhead G (pid C) sig (mhead m) \ unique (lcls (mbody m)) \ + (\pn\set (pars m). table_of (lcls (mbody m)) pn = None) \ + (\(vn,T)\set (lcls (mbody m)). is_acc_type G (pid C) T)" +apply (unfold wf_mdecl_def) +apply auto +done + +lemma wf_mdecl_bodyD: +"wf_mdecl G C (sig,m) \ + (\T. \prg=G,cls=C, + lcl = \ k. + (case k of + EName e + \ (case e of + VNam v \ (table_of (lcls (mbody m))((pars m)[\](parTs sig))) v + | Res \ Some (resTy m)) + | This \ if static m then None else Some (Class C)) + \\Body C (stmt (mbody m))\-T \ G\T\(resTy m))" +apply (unfold wf_mdecl_def) +apply clarify +apply (rule_tac x="(resTy m)" in exI) +apply (unfold wf_mhead_def) +apply (auto simp add: wf_mhead_def is_acc_type_def intro: wt.Body ) +done + + +(* +lemma static_Object_methodsE [elim!]: + "\wf_mdecl G Object (sig, m);static m\ \ R" +apply (unfold wf_mdecl_def) +apply auto +done +*) + +lemma rT_is_acc_type: + "wf_mhead G P sig m \ is_acc_type G P (resTy m)" +apply (unfold wf_mhead_def) +apply auto +done + +section "well-formed interface declarations" + (* well-formed interface declaration, cf. 9.1, 9.1.2.1, 9.1.3, 9.4 *) + +text {* +A interface declaration is wellformed if: +\begin{itemize} +\item the interface hierarchy is wellstructured +\item there is no class with the same name +\item the method heads are wellformed and not static and have Public access +\item the methods are uniquely named +\item all superinterfaces are accessible +\item the result type of a method overriding a method of Object widens to the + result type of the overridden method. + Shadowing static methods is forbidden. +\item the result type of a method overriding a set of methods defined in the + superinterfaces widens to each of the corresponding result types +\end{itemize} +*} +constdefs + wf_idecl :: "prog \ idecl \ bool" + "wf_idecl G \ + \(I,i). + ws_idecl G I (isuperIfs i) \ + \is_class G I \ + (\(sig,mh)\set (imethods i). wf_mhead G (pid I) sig mh \ + \is_static mh \ + accmodi mh = Public) \ + unique (imethods i) \ + (\ J\set (isuperIfs i). is_acc_iface G (pid I) J) \ + (table_of (imethods i) + hiding (methd G Object) + under (\ new old. accmodi old \ Private) + entails (\new old. G\resTy new\resTy old \ + is_static new = is_static old)) \ + (o2s \ table_of (imethods i) + hidings Un_tables((\J.(imethds G J))`set (isuperIfs i)) + entails (\new old. G\resTy new\resTy old))" + +lemma wf_idecl_mhead: "\wf_idecl G (I,i); (sig,mh)\set (imethods i)\ \ + wf_mhead G (pid I) sig mh \ \is_static mh \ accmodi mh = Public" +apply (unfold wf_idecl_def) +apply auto +done + +lemma wf_idecl_hidings: +"wf_idecl G (I, i) \ + (\s. o2s (table_of (imethods i) s)) + hidings Un_tables ((\J. imethds G J) ` set (isuperIfs i)) + entails \new old. G\resTy new\resTy old" +apply (unfold wf_idecl_def o_def) +apply simp +done + +lemma wf_idecl_hiding: +"wf_idecl G (I, i) \ + (table_of (imethods i) + hiding (methd G Object) + under (\ new old. accmodi old \ Private) + entails (\new old. G\resTy new\resTy old \ + is_static new = is_static old))" +apply (unfold wf_idecl_def) +apply simp +done + +lemma wf_idecl_supD: +"\wf_idecl G (I,i); J \ set (isuperIfs i)\ + \ is_acc_iface G (pid I) J \ (J, I) \ (subint1 G)^+" +apply (unfold wf_idecl_def ws_idecl_def) +apply auto +done + +section "well-formed class declarations" + (* well-formed class declaration, cf. 8.1, 8.1.2.1, 8.1.2.2, 8.1.3, 8.1.4 and + class method declaration, cf. 8.4.3.3, 8.4.6.1, 8.4.6.2, 8.4.6.3, 8.4.6.4 *) + +text {* +A class declaration is wellformed if: +\begin{itemize} +\item there is no interface with the same name +\item all superinterfaces are accessible and for all methods implementing + an interface method the result type widens to the result type of + the interface method, the method is not static and offers at least + as much access + (this actually means that the method has Public access, since all + interface methods have public access) +\item all field declarations are wellformed and the field names are unique +\item all method declarations are wellformed and the method names are unique +\item the initialization statement is welltyped +\item the classhierarchy is wellstructured +\item Unless the class is Object: + \begin{itemize} + \item the superclass is accessible + \item for all methods overriding another method (of a superclass )the + result type widens to the result type of the overridden method, + the access modifier of the new method provides at least as much + access as the overwritten one. + \item for all methods hiding a method (of a superclass) the hidden + method must be static and offer at least as much access rights. + Remark: In contrast to the Java Language Specification we don't + restrict the result types of the method + (as in case of overriding), because there seems to be no reason, + since there is no dynamic binding of static methods. + (cf. 8.4.6.3 vs. 15.12.1). + Stricly speaking the restrictions on the access rights aren't + necessary to, since the static type and the access rights + together determine which method is to be called statically. + But if a class gains more then one static method with the + same signature due to inheritance, it is confusing when the + method selection depends on the access rights only: + e.g. + Class C declares static public method foo(). + Class D is subclass of C and declares static method foo() + with default package access. + D.foo() ? if this call is in the same package as D then + foo of class D is called, otherwise foo of class C. + \end{itemize} + +\end{itemize} +*} +(* to Table *) +constdefs entails:: "('a,'b) table \ ('b \ bool) \ bool" + ("_ entails _" 20) +"t entails P \ \k. \ x \ t k: P x" + +lemma entailsD: + "\t entails P; t k = Some x\ \ P x" +by (simp add: entails_def) + +lemma empty_entails[simp]: "empty entails P" +by (simp add: entails_def) + +constdefs + wf_cdecl :: "prog \ cdecl \ bool" +"wf_cdecl G \ + \(C,c). + \is_iface G C \ + (\I\set (superIfs c). is_acc_iface G (pid C) I \ + (\s. \ im \ imethds G I s. + (\ cm \ methd G C s: G\resTy cm\resTy im \ + \ is_static cm \ + accmodi im \ accmodi cm))) \ + (\f\set (cfields c). wf_fdecl G (pid C) f) \ unique (cfields c) \ + (\m\set (methods c). wf_mdecl G C m) \ unique (methods c) \ + \prg=G,cls=C,lcl=empty\\(init c)\\ \ ws_cdecl G C (super c) \ + (C \ Object \ + (is_acc_class G (pid C) (super c) \ + (table_of (map (\ (s,m). (s,C,m)) (methods c)) + entails (\ new. \ old sig. + (G,sig\new overrides\<^sub>S old + \ (G\resTy new\resTy old \ + accmodi old \ accmodi new \ + \is_static old)) \ + (G,sig\new hides old + \ (accmodi old \ accmodi new \ + is_static old)))) + ))" + +(* +constdefs + wf_cdecl :: "prog \ cdecl \ bool" +"wf_cdecl G \ + \(C,c). + \is_iface G C \ + (\I\set (superIfs c). is_acc_iface G (pid C) I \ + (\s. \ im \ imethds G I s. + (\ cm \ methd G C s: G\resTy (mthd cm)\resTy (mthd im) \ + \ is_static cm \ + accmodi im \ accmodi cm))) \ + (\f\set (cfields c). wf_fdecl G (pid C) f) \ unique (cfields c) \ + (\m\set (methods c). wf_mdecl G C m) \ unique (methods c) \ + \prg=G,cls=C,lcl=empty\\(init c)\\ \ ws_cdecl G C (super c) \ + (C \ Object \ + (is_acc_class G (pid C) (super c) \ + (table_of (map (\ (s,m). (s,C,m)) (methods c)) + hiding methd G (super c) + under (\ new old. G\new overrides old) + entails (\ new old. + (G\resTy (mthd new)\resTy (mthd old) \ + accmodi old \ accmodi new \ + \ is_static old))) \ + (table_of (map (\ (s,m). (s,C,m)) (methods c)) + hiding methd G (super c) + under (\ new old. G\new hides old) + entails (\ new old. is_static old \ + accmodi old \ accmodi new)) \ + (table_of (cfields c) hiding accfield G C (super c) + entails (\ newF oldF. accmodi oldF \ access newF))))" +*) + +lemma wf_cdecl_unique: +"wf_cdecl G (C,c) \ unique (cfields c) \ unique (methods c)" +apply (unfold wf_cdecl_def) +apply auto +done + +lemma wf_cdecl_fdecl: +"\wf_cdecl G (C,c); f\set (cfields c)\ \ wf_fdecl G (pid C) f" +apply (unfold wf_cdecl_def) +apply auto +done + +lemma wf_cdecl_mdecl: +"\wf_cdecl G (C,c); m\set (methods c)\ \ wf_mdecl G C m" +apply (unfold wf_cdecl_def) +apply auto +done + +lemma wf_cdecl_impD: +"\wf_cdecl G (C,c); I\set (superIfs c)\ +\ is_acc_iface G (pid C) I \ + (\s. \im \ imethds G I s. + (\cm \ methd G C s: G\resTy cm\resTy im \ \is_static cm \ + accmodi im \ accmodi cm))" +apply (unfold wf_cdecl_def) +apply auto +done + +lemma wf_cdecl_supD: +"\wf_cdecl G (C,c); C \ Object\ \ + is_acc_class G (pid C) (super c) \ (super c,C) \ (subcls1 G)^+ \ + (table_of (map (\ (s,m). (s,C,m)) (methods c)) + entails (\ new. \ old sig. + (G,sig\new overrides\<^sub>S old + \ (G\resTy new\resTy old \ + accmodi old \ accmodi new \ + \is_static old)) \ + (G,sig\new hides old + \ (accmodi old \ accmodi new \ + is_static old))))" +apply (unfold wf_cdecl_def ws_cdecl_def) +apply auto +done + + +lemma wf_cdecl_overrides_SomeD: +"\wf_cdecl G (C,c); C \ Object; table_of (methods c) sig = Some newM; + G,sig\(C,newM) overrides\<^sub>S old +\ \ G\resTy newM\resTy old \ + accmodi old \ accmodi newM \ + \ is_static old" +apply (drule (1) wf_cdecl_supD) +apply (clarify) +apply (drule entailsD) +apply (blast intro: table_of_map_SomeI) +apply (drule_tac x="old" in spec) +apply (auto dest: overrides_eq_sigD simp add: msig_def) +done + +lemma wf_cdecl_hides_SomeD: +"\wf_cdecl G (C,c); C \ Object; table_of (methods c) sig = Some newM; + G,sig\(C,newM) hides old +\ \ accmodi old \ access newM \ + is_static old" +apply (drule (1) wf_cdecl_supD) +apply (clarify) +apply (drule entailsD) +apply (blast intro: table_of_map_SomeI) +apply (drule_tac x="old" in spec) +apply (auto dest: hides_eq_sigD simp add: msig_def) +done + +lemma wf_cdecl_wt_init: + "wf_cdecl G (C, c) \ \prg=G,cls=C,lcl=empty\\init c\\" +apply (unfold wf_cdecl_def) +apply auto +done + + +section "well-formed programs" + (* well-formed program, cf. 8.1, 9.1 *) + +text {* +A program declaration is wellformed if: +\begin{itemize} +\item the class ObjectC of Object is defined +\item every method of has an access modifier distinct from Package. This is + necessary since every interface automatically inherits from Object. + We must know, that every time a Object method is "overriden" by an + interface method this is also overriden by the class implementing the + the interface (see @{text "implement_dynmethd and class_mheadsD"}) +\item all standard Exceptions are defined +\item all defined interfaces are wellformed +\item all defined classes are wellformed +\end{itemize} +*} +constdefs + wf_prog :: "prog \ bool" + "wf_prog G \ let is = ifaces G; cs = classes G in + ObjectC \ set cs \ + (\ m\set Object_mdecls. accmodi m \ Package) \ + (\xn. SXcptC xn \ set cs) \ + (\i\set is. wf_idecl G i) \ unique is \ + (\c\set cs. wf_cdecl G c) \ unique cs" + +lemma wf_prog_idecl: "\iface G I = Some i; wf_prog G\ \ wf_idecl G (I,i)" +apply (unfold wf_prog_def Let_def) +apply simp +apply (fast dest: map_of_SomeD) +done + +lemma wf_prog_cdecl: "\class G C = Some c; wf_prog G\ \ wf_cdecl G (C,c)" +apply (unfold wf_prog_def Let_def) +apply simp +apply (fast dest: map_of_SomeD) +done + +lemma wf_prog_Object_mdecls: +"wf_prog G \ (\ m\set Object_mdecls. accmodi m \ Package)" +apply (unfold wf_prog_def Let_def) +apply simp +done + +lemma wf_prog_acc_superD: + "\wf_prog G; class G C = Some c; C \ Object \ + \ is_acc_class G (pid C) (super c)" +by (auto dest: wf_prog_cdecl wf_cdecl_supD) + +lemma wf_ws_prog [elim!,simp]: "wf_prog G \ ws_prog G" +apply (unfold wf_prog_def Let_def) +apply (rule ws_progI) +apply (simp_all (no_asm)) +apply (auto simp add: is_acc_class_def is_acc_iface_def + dest!: wf_idecl_supD wf_cdecl_supD )+ +done + +lemma class_Object [simp]: +"wf_prog G \ + class G Object = Some \access=Public,cfields=[],methods=Object_mdecls, + init=Skip,super=arbitrary,superIfs=[]\" +apply (unfold wf_prog_def Let_def ObjectC_def) +apply (fast dest!: map_of_SomeI) +done + +lemma methd_Object[simp]: "wf_prog G \ methd G Object = + table_of (map (\(s,m). (s, Object, m)) Object_mdecls)" +apply (subst methd_rec) +apply (auto simp add: Let_def) +done + +lemma wf_prog_Object_methd: +"\wf_prog G; methd G Object sig = Some m\ \ accmodi m \ Package" +by (auto dest!: wf_prog_Object_mdecls) (auto dest!: map_of_SomeD) + +lemma wf_prog_Object_is_public[intro]: + "wf_prog G \ is_public G Object" +by (auto simp add: is_public_def dest: class_Object) + +lemma class_SXcpt [simp]: +"wf_prog G \ + class G (SXcpt xn) = Some \access=Public,cfields=[],methods=SXcpt_mdecls, + init=Skip, + super=if xn = Throwable then Object + else SXcpt Throwable, + superIfs=[]\" +apply (unfold wf_prog_def Let_def SXcptC_def) +apply (fast dest!: map_of_SomeI) +done + +lemma wf_ObjectC [simp]: + "wf_cdecl G ObjectC = (\is_iface G Object \ Ball (set Object_mdecls) + (wf_mdecl G Object) \ unique Object_mdecls)" +apply (unfold wf_cdecl_def ws_cdecl_def ObjectC_def) +apply (simp (no_asm)) +done + +lemma Object_is_class [simp,elim!]: "wf_prog G \ is_class G Object" +apply (simp (no_asm_simp)) +done + +lemma Object_is_acc_class [simp,elim!]: "wf_prog G \ is_acc_class G S Object" +apply (simp (no_asm_simp) add: is_acc_class_def is_public_def + accessible_in_RefT_simp) +done + +lemma SXcpt_is_class [simp,elim!]: "wf_prog G \ is_class G (SXcpt xn)" +apply (simp (no_asm_simp)) +done + +lemma SXcpt_is_acc_class [simp,elim!]: +"wf_prog G \ is_acc_class G S (SXcpt xn)" +apply (simp (no_asm_simp) add: is_acc_class_def is_public_def + accessible_in_RefT_simp) +done + +lemma fields_Object [simp]: "wf_prog G \ DeclConcepts.fields G Object = []" +by (force intro: fields_emptyI) + +lemma accfield_Object [simp]: + "wf_prog G \ accfield G S Object = empty" +apply (unfold accfield_def) +apply (simp (no_asm_simp) add: Let_def) +done + +lemma fields_Throwable [simp]: + "wf_prog G \ DeclConcepts.fields G (SXcpt Throwable) = []" +by (force intro: fields_emptyI) + +lemma fields_SXcpt [simp]: "wf_prog G \ DeclConcepts.fields G (SXcpt xn) = []" +apply (case_tac "xn = Throwable") +apply (simp (no_asm_simp)) +by (force intro: fields_emptyI) + +lemmas widen_trans = ws_widen_trans [OF _ _ wf_ws_prog, elim] +lemma widen_trans2 [elim]: "\G\U\T; G\S\U; wf_prog G\ \ G\S\T" +apply (erule (2) widen_trans) +done + +lemma Xcpt_subcls_Throwable [simp]: +"wf_prog G \ G\SXcpt xn\\<^sub>C SXcpt Throwable" +apply (rule SXcpt_subcls_Throwable_lemma) +apply auto +done + +lemma unique_fields: + "\is_class G C; wf_prog G\ \ unique (DeclConcepts.fields G C)" +apply (erule ws_unique_fields) +apply (erule wf_ws_prog) +apply (erule (1) wf_prog_cdecl [THEN wf_cdecl_unique [THEN conjunct1]]) +done + +lemma fields_mono: +"\table_of (DeclConcepts.fields G C) fn = Some f; G\D\\<^sub>C C; + is_class G D; wf_prog G\ + \ table_of (DeclConcepts.fields G D) fn = Some f" +apply (rule map_of_SomeI) +apply (erule (1) unique_fields) +apply (erule (1) map_of_SomeD [THEN fields_mono_lemma]) +apply (erule wf_ws_prog) +done + + +lemma fields_is_type [elim]: +"\table_of (DeclConcepts.fields G C) m = Some f; wf_prog G; is_class G C\ \ + is_type G (type f)" +apply (frule wf_ws_prog) +apply (force dest: fields_declC [THEN conjunct1] + wf_prog_cdecl [THEN wf_cdecl_fdecl] + simp add: wf_fdecl_def2 is_acc_type_def) +done + +lemma imethds_wf_mhead [rule_format (no_asm)]: +"\m \ imethds G I sig; wf_prog G; is_iface G I\ \ + wf_mhead G (pid (decliface m)) sig (mthd m) \ + \ is_static m \ accmodi m = Public" +apply (frule wf_ws_prog) +apply (drule (2) imethds_declI [THEN conjunct1]) +apply clarify +apply (frule_tac I="(decliface m)" in wf_prog_idecl,assumption) +apply (drule wf_idecl_mhead) +apply (erule map_of_SomeD) +apply (cases m, simp) +done + +lemma methd_wf_mdecl: + "\methd G C sig = Some m; wf_prog G; class G C = Some y\ \ + G\C\\<^sub>C (declclass m) \ is_class G (declclass m) \ + wf_mdecl G (declclass m) (sig,(mthd m))" +apply (frule wf_ws_prog) +apply (drule (1) methd_declC) +apply fast +apply clarsimp +apply (frule (1) wf_prog_cdecl, erule wf_cdecl_mdecl, erule map_of_SomeD) +done + +(* +This lemma doesn't hold! +lemma methd_rT_is_acc_type: +"\wf_prog G;methd G C C sig = Some (D,m); + class G C = Some y\ +\ is_acc_type G (pid C) (resTy m)" +The result Type is only visible in the scope of defining class D +"is_vis_type G (pid D) (resTy m)" but not necessarily in scope of class C! +(The same is true for the type of pramaters of a method) +*) + + +lemma methd_rT_is_type: +"\wf_prog G;methd G C sig = Some m; + class G C = Some y\ +\ is_type G (resTy m)" +apply (drule (2) methd_wf_mdecl) +apply clarify +apply (drule wf_mdeclD1) +apply clarify +apply (drule rT_is_acc_type) +apply (cases m, simp add: is_acc_type_def) +done + +lemma accmethd_rT_is_type: +"\wf_prog G;accmethd G S C sig = Some m; + class G C = Some y\ +\ is_type G (resTy m)" +by (auto simp add: accmethd_def + intro: methd_rT_is_type) + +lemma methd_Object_SomeD: +"\wf_prog G;methd G Object sig = Some m\ + \ declclass m = Object" +by (auto dest: class_Object simp add: methd_rec ) + +lemma wf_imethdsD: + "\im \ imethds G I sig;wf_prog G; is_iface G I\ + \ \is_static im \ accmodi im = Public" +proof - + assume asm: "wf_prog G" "is_iface G I" "im \ imethds G I sig" + have "wf_prog G \ + (\ i im. iface G I = Some i \ im \ imethds G I sig + \ \is_static im \ accmodi im = Public)" (is "?P G I") + proof (rule iface_rec.induct,intro allI impI) + fix G I i im + assume hyp: "\ J i. J \ set (isuperIfs i) \ ws_prog G \ iface G I = Some i + \ ?P G J" + assume wf: "wf_prog G" and if_I: "iface G I = Some i" and + im: "im \ imethds G I sig" + show "\is_static im \ accmodi im = Public" + proof - + let ?inherited = "Un_tables (imethds G ` set (isuperIfs i))" + let ?new = "(o2s \ table_of (map (\(s, mh). (s, I, mh)) (imethods i)))" + from if_I wf im have imethds:"im \ (?inherited \\ ?new) sig" + by (simp add: imethds_rec) + from wf if_I have + wf_supI: "\ J. J \ set (isuperIfs i) \ (\ j. iface G J = Some j)" + by (blast dest: wf_prog_idecl wf_idecl_supD is_acc_ifaceD) + from wf if_I have + "\ im \ set (imethods i). \ is_static im \ accmodi im = Public" + by (auto dest!: wf_prog_idecl wf_idecl_mhead) + then have new_ok: "\ im. table_of (imethods i) sig = Some im + \ \ is_static im \ accmodi im = Public" + by (auto dest!: table_of_Some_in_set) + show ?thesis + proof (cases "?new sig = {}") + case True + from True wf wf_supI if_I imethds hyp + show ?thesis by (auto simp del: split_paired_All) + next + case False + from False wf wf_supI if_I imethds new_ok hyp + show ?thesis by (auto dest: wf_idecl_hidings hidings_entailsD) + qed + qed + qed + with asm show ?thesis by (auto simp del: split_paired_All) +qed + +lemma wf_prog_hidesD: + (assumes hides: "G \new hides old" and wf: "wf_prog G") + "accmodi old \ accmodi new \ + is_static old" +proof - + from hides + obtain c where + clsNew: "class G (declclass new) = Some c" and + neqObj: "declclass new \ Object" + by (auto dest: hidesD declared_in_classD) + with hides obtain newM oldM where + newM: "table_of (methods c) (msig new) = Some newM" and + new: "new = (declclass new,(msig new),newM)" and + old: "old = (declclass old,(msig old),oldM)" and + "msig new = msig old" + by (cases new,cases old) + (auto dest: hidesD + simp add: cdeclaredmethd_def declared_in_def) + with hides + have hides': + "G,(msig new)\(declclass new,newM) hides (declclass old,oldM)" + by auto + from clsNew wf + have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl) + note wf_cdecl_hides_SomeD [OF this neqObj newM hides'] + with new old + show ?thesis + by (cases new, cases old) auto +qed + +text {* Compare this lemma about static +overriding @{term "G \new overrides\<^sub>S old"} with the definition of +dynamic overriding @{term "G \new overrides old"}. +Conforming result types and restrictions on the access modifiers of the old +and the new method are not part of the predicate for static overriding. But +they are enshured in a wellfromed program. Dynamic overriding has +no restrictions on the access modifiers but enforces confrom result types +as precondition. But with some efford we can guarantee the access modifier +restriction for dynamic overriding, too. See lemma +@{text wf_prog_dyn_override_prop}. +*} +lemma wf_prog_stat_overridesD: + (assumes stat_override: "G \new overrides\<^sub>S old" and wf: "wf_prog G") + "G\resTy new\resTy old \ + accmodi old \ accmodi new \ + \ is_static old" +proof - + from stat_override + obtain c where + clsNew: "class G (declclass new) = Some c" and + neqObj: "declclass new \ Object" + by (auto dest: stat_overrides_commonD declared_in_classD) + with stat_override obtain newM oldM where + newM: "table_of (methods c) (msig new) = Some newM" and + new: "new = (declclass new,(msig new),newM)" and + old: "old = (declclass old,(msig old),oldM)" and + "msig new = msig old" + by (cases new,cases old) + (auto dest: stat_overrides_commonD + simp add: cdeclaredmethd_def declared_in_def) + with stat_override + have stat_override': + "G,(msig new)\(declclass new,newM) overrides\<^sub>S (declclass old,oldM)" + by auto + from clsNew wf + have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl) + note wf_cdecl_overrides_SomeD [OF this neqObj newM stat_override'] + with new old + show ?thesis + by (cases new, cases old) auto +qed + +lemma static_to_dynamic_overriding: + (assumes stat_override: "G\new overrides\<^sub>S old" and wf : "wf_prog G") + "G\new overrides old" +proof - + from stat_override + show ?thesis (is "?Overrides new old") + proof (induct) + case (Direct new old superNew) + then have stat_override:"G\new overrides\<^sub>S old" + by (rule stat_overridesR.Direct) + from stat_override wf + have resTy_widen: "G\resTy new\resTy old" and + not_static_old: "\ is_static old" + by (auto dest: wf_prog_stat_overridesD) + have not_private_new: "accmodi new \ Private" + proof - + from stat_override + have "accmodi old \ Private" + by (rule no_Private_stat_override) + moreover + from stat_override wf + have "accmodi old \ accmodi new" + by (auto dest: wf_prog_stat_overridesD) + ultimately + show ?thesis + by (auto dest: acc_modi_bottom) + qed + with Direct resTy_widen not_static_old + show "?Overrides new old" + by (auto intro: overridesR.Direct) + next + case (Indirect inter new old) + then show "?Overrides new old" + by (blast intro: overridesR.Indirect) + qed +qed + +lemma non_Package_instance_method_inheritance: +(assumes old_inheritable: "G\Method old inheritable_in (pid C)" and + accmodi_old: "accmodi old \ Package" and + instance_method: "\ is_static old" and + subcls: "G\C \\<^sub>C declclass old" and + old_declared: "G\Method old declared_in (declclass old)" and + wf: "wf_prog G" +) "G\Method old member_of C \ + (\ new. G\ new overrides\<^sub>S old \ G\Method new member_of C)" +proof - + from wf have ws: "ws_prog G" by auto + from old_declared have iscls_declC_old: "is_class G (declclass old)" + by (auto simp add: declared_in_def cdeclaredmethd_def) + from subcls have iscls_C: "is_class G C" + by (blast dest: subcls_is_class) + from iscls_C ws old_inheritable subcls + show ?thesis (is "?P C old") + proof (induct rule: ws_class_induct') + case Object + assume "G\Object\\<^sub>C declclass old" + then show "?P Object old" + by blast + next + case (Subcls C c) + assume cls_C: "class G C = Some c" and + neq_C_Obj: "C \ Object" and + hyp: "\G \Method old inheritable_in pid (super c); + G\super c\\<^sub>C declclass old\ \ ?P (super c) old" and + inheritable: "G \Method old inheritable_in pid C" and + subclsC: "G\C\\<^sub>C declclass old" + from cls_C neq_C_Obj + have super: "G\C \\<^sub>C\<^sub>1 super c" + by (rule subcls1I) + from wf cls_C neq_C_Obj + have accessible_super: "G\(Class (super c)) accessible_in (pid C)" + by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD) + { + fix old + assume member_super: "G\Method old member_of (super c)" + assume inheritable: "G \Method old inheritable_in pid C" + assume instance_method: "\ is_static old" + from member_super + have old_declared: "G\Method old declared_in (declclass old)" + by (cases old) (auto dest: member_of_declC) + have "?P C old" + proof (cases "G\mid (msig old) undeclared_in C") + case True + with inheritable super accessible_super member_super + have "G\Method old member_of C" + by (cases old) (auto intro: members.Inherited) + then show ?thesis + by auto + next + case False + then obtain new_member where + "G\new_member declared_in C" and + "mid (msig old) = memberid new_member" + by (auto dest: not_undeclared_declared) + then obtain new where + new: "G\Method new declared_in C" and + eq_sig: "msig old = msig new" and + declC_new: "declclass new = C" + by (cases new_member) auto + then have member_new: "G\Method new member_of C" + by (cases new) (auto intro: members.Immediate) + from declC_new super member_super + have subcls_new_old: "G\declclass new \\<^sub>C declclass old" + by (auto dest!: member_of_subclseq_declC + dest: r_into_trancl intro: trancl_rtrancl_trancl) + show ?thesis + proof (cases "is_static new") + case False + with eq_sig declC_new new old_declared inheritable + super member_super subcls_new_old + have "G\new overrides\<^sub>S old" + by (auto intro!: stat_overridesR.Direct) + with member_new show ?thesis + by blast + next + case True + with eq_sig declC_new subcls_new_old new old_declared inheritable + have "G\new hides old" + by (auto intro: hidesI) + with wf + have "is_static old" + by (blast dest: wf_prog_hidesD) + with instance_method + show ?thesis + by (contradiction) + qed + qed + } note hyp_member_super = this + from subclsC cls_C + have "G\(super c)\\<^sub>C declclass old" + by (rule subcls_superD) + then + show "?P C old" + proof (cases rule: subclseq_cases) + case Eq + assume "super c = declclass old" + with old_declared + have "G\Method old member_of (super c)" + by (cases old) (auto intro: members.Immediate) + with inheritable instance_method + show ?thesis + by (blast dest: hyp_member_super) + next + case Subcls + assume "G\super c\\<^sub>C declclass old" + moreover + from inheritable accmodi_old + have "G \Method old inheritable_in pid (super c)" + by (cases "accmodi old") (auto simp add: inheritable_in_def) + ultimately + have "?P (super c) old" + by (blast dest: hyp) + then show ?thesis + proof + assume "G \Method old member_of super c" + with inheritable instance_method + show ?thesis + by (blast dest: hyp_member_super) + next + assume "\new. G \ new overrides\<^sub>S old \ G \Method new member_of super c" + then obtain super_new where + super_new_override: "G \ super_new overrides\<^sub>S old" and + super_new_member: "G \Method super_new member_of super c" + by blast + from super_new_override wf + have "accmodi old \ accmodi super_new" + by (auto dest: wf_prog_stat_overridesD) + with inheritable accmodi_old + have "G \Method super_new inheritable_in pid C" + by (auto simp add: inheritable_in_def + split: acc_modi.splits + dest: acc_modi_le_Dests) + moreover + from super_new_override + have "\ is_static super_new" + by (auto dest: stat_overrides_commonD) + moreover + note super_new_member + ultimately have "?P C super_new" + by (auto dest: hyp_member_super) + then show ?thesis + proof + assume "G \Method super_new member_of C" + with super_new_override + show ?thesis + by blast + next + assume "\new. G \ new overrides\<^sub>S super_new \ + G \Method new member_of C" + with super_new_override show ?thesis + by (blast intro: stat_overridesR.Indirect) + qed + qed + qed + qed +qed + +lemma non_Package_instance_method_inheritance_cases [consumes 6, + case_names Inheritance Overriding]: +(assumes old_inheritable: "G\Method old inheritable_in (pid C)" and + accmodi_old: "accmodi old \ Package" and + instance_method: "\ is_static old" and + subcls: "G\C \\<^sub>C declclass old" and + old_declared: "G\Method old declared_in (declclass old)" and + wf: "wf_prog G" and + inheritance: "G\Method old member_of C \ P" and + overriding: "\ new. + \G\ new overrides\<^sub>S old;G\Method new member_of C\ + \ P" +) "P" +proof - + from old_inheritable accmodi_old instance_method subcls old_declared wf + inheritance overriding + show ?thesis + by (auto dest: non_Package_instance_method_inheritance) +qed + +lemma dynamic_to_static_overriding: +(assumes dyn_override: "G\ new overrides old" and + accmodi_old: "accmodi old \ Package" and + wf: "wf_prog G" +) "G\ new overrides\<^sub>S old" +proof - + from dyn_override accmodi_old + show ?thesis (is "?Overrides new old") + proof (induct rule: overridesR.induct) + case (Direct new old) + assume new_declared: "G\Method new declared_in declclass new" + assume eq_sig_new_old: "msig new = msig old" + assume subcls_new_old: "G\declclass new \\<^sub>C declclass old" + assume "G \Method old inheritable_in pid (declclass new)" and + "accmodi old \ Package" and + "\ is_static old" and + "G\declclass new\\<^sub>C declclass old" and + "G\Method old declared_in declclass old" + from this wf + show "?Overrides new old" + proof (cases rule: non_Package_instance_method_inheritance_cases) + case Inheritance + assume "G \Method old member_of declclass new" + then have "G\mid (msig old) undeclared_in declclass new" + proof cases + case Immediate + with subcls_new_old wf show ?thesis + by (auto dest: subcls_irrefl) + next + case Inherited + then show ?thesis + by (cases old) auto + qed + with eq_sig_new_old new_declared + show ?thesis + by (cases old,cases new) (auto dest!: declared_not_undeclared) + next + case (Overriding new') + assume stat_override_new': "G \ new' overrides\<^sub>S old" + then have "msig new' = msig old" + by (auto dest: stat_overrides_commonD) + with eq_sig_new_old have eq_sig_new_new': "msig new=msig new'" + by simp + assume "G \Method new' member_of declclass new" + then show ?thesis + proof (cases) + case Immediate + then have declC_new: "declclass new' = declclass new" + by auto + from Immediate + have "G\Method new' declared_in declclass new" + by (cases new') auto + with new_declared eq_sig_new_new' declC_new + have "new=new'" + by (cases new, cases new') (auto dest: unique_declared_in) + with stat_override_new' + show ?thesis + by simp + next + case Inherited + then have "G\mid (msig new') undeclared_in declclass new" + by (cases new') (auto) + with eq_sig_new_new' new_declared + show ?thesis + by (cases new,cases new') (auto dest!: declared_not_undeclared) + qed + qed + next + case (Indirect inter new old) + assume accmodi_old: "accmodi old \ Package" + assume "accmodi old \ Package \ G \ inter overrides\<^sub>S old" + with accmodi_old + have stat_override_inter_old: "G \ inter overrides\<^sub>S old" + by blast + moreover + assume hyp_inter: "accmodi inter \ Package \ G \ new overrides\<^sub>S inter" + moreover + have "accmodi inter \ Package" + proof - + from stat_override_inter_old wf + have "accmodi old \ accmodi inter" + by (auto dest: wf_prog_stat_overridesD) + with stat_override_inter_old accmodi_old + show ?thesis + by (auto dest!: no_Private_stat_override + split: acc_modi.splits + dest: acc_modi_le_Dests) + qed + ultimately show "?Overrides new old" + by (blast intro: stat_overridesR.Indirect) + qed +qed + +lemma wf_prog_dyn_override_prop: + (assumes dyn_override: "G \ new overrides old" and + wf: "wf_prog G" + ) "accmodi old \ accmodi new" +proof (cases "accmodi old = Package") + case True + note old_Package = this + show ?thesis + proof (cases "accmodi old \ accmodi new") + case True then show ?thesis . + next + case False + with old_Package + have "accmodi new = Private" + by (cases "accmodi new") (auto simp add: le_acc_def less_acc_def) + with dyn_override + show ?thesis + by (auto dest: overrides_commonD) + qed +next + case False + with dyn_override wf + have "G \ new overrides\<^sub>S old" + by (blast intro: dynamic_to_static_overriding) + with wf + show ?thesis + by (blast dest: wf_prog_stat_overridesD) +qed + +lemma overrides_Package_old: +(assumes dyn_override: "G \ new overrides old" and + accmodi_new: "accmodi new = Package" and + wf: "wf_prog G " +) "accmodi old = Package" +proof (cases "accmodi old") + case Private + with dyn_override show ?thesis + by (simp add: no_Private_override) +next + case Package + then show ?thesis . +next + case Protected + with dyn_override wf + have "G \ new overrides\<^sub>S old" + by (auto intro: dynamic_to_static_overriding) + with wf + have "accmodi old \ accmodi new" + by (auto dest: wf_prog_stat_overridesD) + with Protected accmodi_new + show ?thesis + by (simp add: less_acc_def le_acc_def) +next + case Public + with dyn_override wf + have "G \ new overrides\<^sub>S old" + by (auto intro: dynamic_to_static_overriding) + with wf + have "accmodi old \ accmodi new" + by (auto dest: wf_prog_stat_overridesD) + with Public accmodi_new + show ?thesis + by (simp add: less_acc_def le_acc_def) +qed + +lemma dyn_override_Package: + (assumes dyn_override: "G \ new overrides old" and + accmodi_old: "accmodi old = Package" and + accmodi_new: "accmodi new = Package" and + wf: "wf_prog G" + ) "pid (declclass old) = pid (declclass new)" +proof - + from dyn_override accmodi_old accmodi_new + show ?thesis (is "?EqPid old new") + proof (induct rule: overridesR.induct) + case (Direct new old) + assume "accmodi old = Package" + "G \Method old inheritable_in pid (declclass new)" + then show "pid (declclass old) = pid (declclass new)" + by (auto simp add: inheritable_in_def) + next + case (Indirect inter new old) + assume accmodi_old: "accmodi old = Package" and + accmodi_new: "accmodi new = Package" + assume "G \ new overrides inter" + with accmodi_new wf + have "accmodi inter = Package" + by (auto intro: overrides_Package_old) + with Indirect + show "pid (declclass old) = pid (declclass new)" + by auto + qed +qed + +lemma dyn_override_Package_escape: + (assumes dyn_override: "G \ new overrides old" and + accmodi_old: "accmodi old = Package" and + outside_pack: "pid (declclass old) \ pid (declclass new)" and + wf: "wf_prog G" + ) "\ inter. G \ new overrides inter \ G \ inter overrides old \ + pid (declclass old) = pid (declclass inter) \ + Protected \ accmodi inter" +proof - + from dyn_override accmodi_old outside_pack + show ?thesis (is "?P new old") + proof (induct rule: overridesR.induct) + case (Direct new old) + assume accmodi_old: "accmodi old = Package" + assume outside_pack: "pid (declclass old) \ pid (declclass new)" + assume "G \Method old inheritable_in pid (declclass new)" + with accmodi_old + have "pid (declclass old) = pid (declclass new)" + by (simp add: inheritable_in_def) + with outside_pack + show "?P new old" + by (contradiction) + next + case (Indirect inter new old) + assume accmodi_old: "accmodi old = Package" + assume outside_pack: "pid (declclass old) \ pid (declclass new)" + assume override_new_inter: "G \ new overrides inter" + assume override_inter_old: "G \ inter overrides old" + assume hyp_new_inter: "\accmodi inter = Package; + pid (declclass inter) \ pid (declclass new)\ + \ ?P new inter" + assume hyp_inter_old: "\accmodi old = Package; + pid (declclass old) \ pid (declclass inter)\ + \ ?P inter old" + show "?P new old" + proof (cases "pid (declclass old) = pid (declclass inter)") + case True + note same_pack_old_inter = this + show ?thesis + proof (cases "pid (declclass inter) = pid (declclass new)") + case True + with same_pack_old_inter outside_pack + show ?thesis + by auto + next + case False + note diff_pack_inter_new = this + show ?thesis + proof (cases "accmodi inter = Package") + case True + with diff_pack_inter_new hyp_new_inter + obtain newinter where + over_new_newinter: "G \ new overrides newinter" and + over_newinter_inter: "G \ newinter overrides inter" and + eq_pid: "pid (declclass inter) = pid (declclass newinter)" and + accmodi_newinter: "Protected \ accmodi newinter" + by auto + from over_newinter_inter override_inter_old + have "G\newinter overrides old" + by (rule overridesR.Indirect) + moreover + from eq_pid same_pack_old_inter + have "pid (declclass old) = pid (declclass newinter)" + by simp + moreover + note over_new_newinter accmodi_newinter + ultimately show ?thesis + by blast + next + case False + with override_new_inter + have "Protected \ accmodi inter" + by (cases "accmodi inter") (auto dest: no_Private_override) + with override_new_inter override_inter_old same_pack_old_inter + show ?thesis + by blast + qed + qed + next + case False + with accmodi_old hyp_inter_old + obtain newinter where + over_inter_newinter: "G \ inter overrides newinter" and + over_newinter_old: "G \ newinter overrides old" and + eq_pid: "pid (declclass old) = pid (declclass newinter)" and + accmodi_newinter: "Protected \ accmodi newinter" + by auto + from override_new_inter over_inter_newinter + have "G \ new overrides newinter" + by (rule overridesR.Indirect) + with eq_pid over_newinter_old accmodi_newinter + show ?thesis + by blast + qed + qed +qed + +declare split_paired_All [simp del] split_paired_Ex [simp del] +ML_setup {* +simpset_ref() := simpset() delloop "split_all_tac"; +claset_ref () := claset () delSWrapper "split_all_tac" +*} + +lemma declclass_widen[rule_format]: + "wf_prog G + \ (\c m. class G C = Some c \ methd G C sig = Some m + \ G\C \\<^sub>C declclass m)" (is "?P G C") +proof (rule class_rec.induct,intro allI impI) + fix G C c m + assume Hyp: "\c. C \ Object \ ws_prog G \ class G C = Some c + \ ?P G (super c)" + assume wf: "wf_prog G" and cls_C: "class G C = Some c" and + m: "methd G C sig = Some m" + show "G\C\\<^sub>C declclass m" + proof (cases "C=Object") + case True + with wf m show ?thesis by (simp add: methd_Object_SomeD) + next + let ?filter="filter_tab (\sig m. G\C inherits method sig m)" + let ?table = "table_of (map (\(s, m). (s, C, m)) (methods c))" + case False + with cls_C wf m + have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m " + by (simp add: methd_rec) + show ?thesis + proof (cases "?table sig") + case None + from this methd_C have "?filter (methd G (super c)) sig = Some m" + by simp + moreover + from wf cls_C False obtain sup where "class G (super c) = Some sup" + by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class) + moreover note wf False cls_C Hyp + ultimately have "G\super c \\<^sub>C declclass m" by auto + moreover from cls_C False have "G\C \\<^sub>C\<^sub>1 super c" by (rule subcls1I) + ultimately show ?thesis by - (rule rtrancl_into_rtrancl2) + next + case Some + from this wf False cls_C methd_C show ?thesis by auto + qed + qed +qed + +(* +lemma declclass_widen[rule_format]: + "wf_prog G + \ (\c m. class G C = Some c \ methd G C sig = Some m + \ G\C \\<^sub>C declclass m)" (is "?P G C") +apply (rule class_rec.induct) +apply (rule impI)+ +apply (case_tac "C=Object") +apply (force simp add: methd_Object_SomeD) + +apply (rule allI)+ +apply (rule impI) +apply (simp (no_asm_simp) add: methd_rec) +apply (case_tac "table_of (map (\(s, m). (s, C, m)) (methods c)) sig") +apply (simp add: override_def) +apply (frule (1) subcls1I) +apply (drule (1) wf_prog_cdecl) +apply (drule (1) wf_cdecl_supD) +apply clarify +apply (drule is_acc_class_is_class) +apply clarify +apply (blast dest: rtrancl_into_rtrancl2) + +apply auto +done +*) + +(* +lemma accessible_public_inheritance_lemma1: + "\wf_prog G; class G C = Some c; C \ Object; accmodi m = Public; + G\m accessible_through_inheritance_in (super c)\ + \ G\m accessible_through_inheritance_in C" +apply (frule (1) subcls1I) +apply (rule accessible_through_inheritance.Indirect) +apply (blast) +apply (erule accessible_through_inheritance_subclsD) +apply (blast dest: wf_prog_acc_superD is_acc_classD) +apply assumption +apply (force dest: wf_prog_acc_superD is_acc_classD + simp add: accessible_for_inheritance_in_def) +done + +lemma accessible_public_inheritance_lemma[rule_format]: + "\wf_prog G;C \ Object; class G C = Some c; + accmodi m = Public + \ \ methd G (super c) sig = Some m + \ G\m accessible_through_inheritance_in C" +apply (frule (2) wf_prog_acc_superD [THEN is_acc_classD]) +apply (erule conjE) +apply (simp only: not_None_eq) +apply (erule exE) +apply (case_tac "(super c) = Object") +apply (rule impI) +apply (rule accessible_through_inheritance.Direct) +apply force +apply (force simp add: accessible_for_inheritance_in_def) + +apply (frule wf_ws_prog) +apply (simp add: methd_rec) +apply (case_tac "table_of (map (\(s, m). (s, super c, m)) (methods y)) sig") +apply simp +apply (clarify) +apply (rule_tac D="super c" in accessible_through_inheritance.Indirect) +apply (blast dest: subcls1I) +apply (blast) +apply simp +apply assumption +apply (simp add: accessible_for_inheritance_in_def) + +apply clarsimp +apply (rule accessible_through_inheritance.Direct) +apply (auto dest: subcls1I simp add: accessible_for_inheritance_in_def) +done + +lemma accessible_public_inheritance: + "\wf_prog G; class G D = Some d; G\C \\<^sub>C D; methd G D sig = Some m; + accmodi m = Public\ + \ G \ m accessible_through_inheritance_in C" +apply (erule converse_trancl_induct) +apply (blast dest: subcls1D intro: accessible_public_inheritance_lemma) + +apply (frule subcls1D) +apply clarify +apply (frule (2) wf_prog_acc_superD [THEN is_acc_classD]) +apply clarify +apply (rule_tac D="super c" in accessible_through_inheritance.Indirect) +apply (auto intro:trancl_into_trancl2 + accessible_through_inheritance_subclsD + simp add: accessible_for_inheritance_in_def) +done +*) + + +lemma declclass_methd_Object: + "\wf_prog G; methd G Object sig = Some m\ \ declclass m = Object" +by auto + + +lemma methd_declaredD: + "\wf_prog G; is_class G C;methd G C sig = Some m\ + \ G\(mdecl (sig,mthd m)) declared_in (declclass m)" +proof - + assume wf: "wf_prog G" + then have ws: "ws_prog G" .. + assume clsC: "is_class G C" + from clsC ws + show "methd G C sig = Some m + \ G\(mdecl (sig,mthd m)) declared_in (declclass m)" + (is "PROP ?P C") + proof (induct ?P C rule: ws_class_induct') + case Object + assume "methd G Object sig = Some m" + with wf show ?thesis + by - (rule method_declared_inI, auto) + next + case Subcls + fix C c + assume clsC: "class G C = Some c" + and m: "methd G C sig = Some m" + and hyp: "methd G (super c) sig = Some m \ ?thesis" + let ?newMethods = "table_of (map (\(s, m). (s, C, m)) (methods c))" + show ?thesis + proof (cases "?newMethods sig") + case None + from None ws clsC m hyp + show ?thesis by (auto intro: method_declared_inI simp add: methd_rec) + next + case Some + from Some ws clsC m + show ?thesis by (auto intro: method_declared_inI simp add: methd_rec) + qed + qed +qed + + +lemma methd_rec_Some_cases [consumes 4, case_names NewMethod InheritedMethod]: +(assumes methd_C: "methd G C sig = Some m" and + ws: "ws_prog G" and + clsC: "class G C = Some c" and + neq_C_Obj: "C\Object" ) +"\table_of (map (\(s, m). (s, C, m)) (methods c)) sig = Some m \ P; + \G\C inherits (method sig m); methd G (super c) sig = Some m\ \ P + \ \ P" +proof - + let ?inherited = "filter_tab (\sig m. G\C inherits method sig m) + (methd G (super c))" + let ?new = "table_of (map (\(s, m). (s, C, m)) (methods c))" + from ws clsC neq_C_Obj methd_C + have methd_unfold: "(?inherited ++ ?new) sig = Some m" + by (simp add: methd_rec) + assume NewMethod: "?new sig = Some m \ P" + assume InheritedMethod: "\G\C inherits (method sig m); + methd G (super c) sig = Some m\ \ P" + show P + proof (cases "?new sig") + case None + with methd_unfold have "?inherited sig = Some m" + by (auto) + with InheritedMethod show P by blast + next + case Some + with methd_unfold have "?new sig = Some m" + by auto + with NewMethod show P by blast + qed +qed + + +lemma methd_member_of: + (assumes wf: "wf_prog G") + "\is_class G C; methd G C sig = Some m\ \ G\Methd sig m member_of C" + (is "?Class C \ ?Method C \ ?MemberOf C") +proof - + from wf have ws: "ws_prog G" .. + assume defC: "is_class G C" + from defC ws + show "?Class C \ ?Method C \ ?MemberOf C" + proof (induct rule: ws_class_induct') + case Object + with wf have declC: "declclass m = Object" + by (blast intro: declclass_methd_Object) + with Object wf have "G\Methd sig m declared_in Object" + by (auto dest: methd_declaredD simp del: methd_Object) + with declC + show "?MemberOf Object" + by (auto intro!: members.Immediate + simp del: methd_Object) + next + case (Subcls C c) + assume clsC: "class G C = Some c" and + neq_C_Obj: "C \ Object" + assume methd: "?Method C" + from methd ws clsC neq_C_Obj + show "?MemberOf C" + proof (cases rule: methd_rec_Some_cases) + case NewMethod + with clsC show ?thesis + by (auto dest: method_declared_inI intro!: members.Immediate) + next + case InheritedMethod + then show "?thesis" + by (blast dest: inherits_member) + qed + qed +qed + +lemma current_methd: + "\table_of (methods c) sig = Some new; + ws_prog G; class G C = Some c; C \ Object; + methd G (super c) sig = Some old\ + \ methd G C sig = Some (C,new)" +by (auto simp add: methd_rec + intro: filter_tab_SomeI override_find_right table_of_map_SomeI) + +lemma wf_prog_staticD: + (assumes wf: "wf_prog G" and + clsC: "class G C = Some c" and + neq_C_Obj: "C \ Object" and + old: "methd G (super c) sig = Some old" and + accmodi_old: "Protected \ accmodi old" and + new: "table_of (methods c) sig = Some new" + ) "is_static new = is_static old" +proof - + from clsC wf + have wf_cdecl: "wf_cdecl G (C,c)" by (rule wf_prog_cdecl) + from wf clsC neq_C_Obj + have is_cls_super: "is_class G (super c)" + by (blast dest: wf_prog_acc_superD is_acc_classD) + from wf is_cls_super old + have old_member_of: "G\Methd sig old member_of (super c)" + by (rule methd_member_of) + from old wf is_cls_super + have old_declared: "G\Methd sig old declared_in (declclass old)" + by (auto dest: methd_declared_in_declclass) + from new clsC + have new_declared: "G\Methd sig (C,new) declared_in C" + by (auto intro: method_declared_inI) + note trancl_rtrancl_tranc = trancl_rtrancl_trancl [trans] (* ### in Basis *) + from clsC neq_C_Obj + have subcls1_C_super: "G\C \\<^sub>C\<^sub>1 super c" + by (rule subcls1I) + then have "G\C \\<^sub>C super c" .. + also from old wf is_cls_super + have "G\super c \\<^sub>C (declclass old)" by (auto dest: methd_declC) + finally have subcls_C_old: "G\C \\<^sub>C (declclass old)" . + from accmodi_old + have inheritable: "G\Methd sig old inheritable_in pid C" + by (auto simp add: inheritable_in_def + dest: acc_modi_le_Dests) + show ?thesis + proof (cases "is_static new") + case True + with subcls_C_old new_declared old_declared inheritable + have "G,sig\(C,new) hides old" + by (auto intro: hidesI) + with True wf_cdecl neq_C_Obj new + show ?thesis + by (auto dest: wf_cdecl_hides_SomeD) + next + case False + with subcls_C_old new_declared old_declared inheritable subcls1_C_super + old_member_of + have "G,sig\(C,new) overrides\<^sub>S old" + by (auto intro: stat_overridesR.Direct) + with False wf_cdecl neq_C_Obj new + show ?thesis + by (auto dest: wf_cdecl_overrides_SomeD) + qed +qed + +lemma inheritable_instance_methd: + (assumes subclseq_C_D: "G\C \\<^sub>C D" and + is_cls_D: "is_class G D" and + wf: "wf_prog G" and + old: "methd G D sig = Some old" and + accmodi_old: "Protected \ accmodi old" and + not_static_old: "\ is_static old" + ) + "\new. methd G C sig = Some new \ + (new = old \ G,sig\new overrides\<^sub>S old)" + (is "(\new. (?Constraint C new old))") +proof - + from subclseq_C_D is_cls_D + have is_cls_C: "is_class G C" by (rule subcls_is_class2) + from wf + have ws: "ws_prog G" .. + from is_cls_C ws subclseq_C_D + show "\new. ?Constraint C new old" + proof (induct rule: ws_class_induct') + case (Object co) + then have eq_D_Obj: "D=Object" by auto + with old + have "?Constraint Object old old" + by auto + with eq_D_Obj + show "\ new. ?Constraint Object new old" by auto + next + case (Subcls C c) + assume hyp: "G\super c\\<^sub>C D \ \new. ?Constraint (super c) new old" + assume clsC: "class G C = Some c" + assume neq_C_Obj: "C\Object" + from clsC wf + have wf_cdecl: "wf_cdecl G (C,c)" + by (rule wf_prog_cdecl) + from ws clsC neq_C_Obj + have is_cls_super: "is_class G (super c)" + by (auto dest: ws_prog_cdeclD) + from clsC wf neq_C_Obj + have superAccessible: "G\(Class (super c)) accessible_in (pid C)" and + subcls1_C_super: "G\C \\<^sub>C\<^sub>1 super c" + by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD + intro: subcls1I) + show "\new. ?Constraint C new old" + proof (cases "G\super c\\<^sub>C D") + case False + from False Subcls + have eq_C_D: "C=D" + by (auto dest: subclseq_superD) + with old + have "?Constraint C old old" + by auto + with eq_C_D + show "\ new. ?Constraint C new old" by auto + next + case True + with hyp obtain super_method + where super: "?Constraint (super c) super_method old" by blast + from super not_static_old + have not_static_super: "\ is_static super_method" + by (auto dest!: stat_overrides_commonD) + from super old wf accmodi_old + have accmodi_super_method: "Protected \ accmodi super_method" + by (auto dest!: wf_prog_stat_overridesD + intro: order_trans) + from super accmodi_old wf + have inheritable: "G\Methd sig super_method inheritable_in (pid C)" + by (auto dest!: wf_prog_stat_overridesD + acc_modi_le_Dests + simp add: inheritable_in_def) + from super wf is_cls_super + have member: "G\Methd sig super_method member_of (super c)" + by (auto intro: methd_member_of) + from member + have decl_super_method: + "G\Methd sig super_method declared_in (declclass super_method)" + by (auto dest: member_of_declC) + from super subcls1_C_super ws is_cls_super + have subcls_C_super: "G\C \\<^sub>C (declclass super_method)" + by (auto intro: rtrancl_into_trancl2 dest: methd_declC) + show "\ new. ?Constraint C new old" + proof (cases "methd G C sig") + case None + have "methd G (super c) sig = None" + proof - + from clsC ws None + have no_new: "table_of (methods c) sig = None" + by (auto simp add: methd_rec) + with clsC + have undeclared: "G\mid sig undeclared_in C" + by (auto simp add: undeclared_in_def cdeclaredmethd_def) + with inheritable member superAccessible subcls1_C_super + have inherits: "G\C inherits (method sig super_method)" + by (auto simp add: inherits_def) + with clsC ws no_new super neq_C_Obj + have "methd G C sig = Some super_method" + by (auto simp add: methd_rec override_def + intro: filter_tab_SomeI) + with None show ?thesis + by simp + qed + with super show ?thesis by auto + next + case (Some new) + from this ws clsC neq_C_Obj + show ?thesis + proof (cases rule: methd_rec_Some_cases) + case InheritedMethod + with super Some show ?thesis + by auto + next + case NewMethod + assume new: "table_of (map (\(s, m). (s, C, m)) (methods c)) sig + = Some new" + from new + have declcls_new: "declclass new = C" + by auto + from wf clsC neq_C_Obj super new not_static_super accmodi_super_method + have not_static_new: "\ is_static new" + by (auto dest: wf_prog_staticD) + from clsC new + have decl_new: "G\Methd sig new declared_in C" + by (auto simp add: declared_in_def cdeclaredmethd_def) + from not_static_new decl_new decl_super_method + member subcls1_C_super inheritable declcls_new subcls_C_super + have "G,sig\ new overrides\<^sub>S super_method" + by (auto intro: stat_overridesR.Direct) + with super Some + show ?thesis + by (auto intro: stat_overridesR.Indirect) + qed + qed + qed + qed +qed + +lemma inheritable_instance_methd_cases [consumes 6 + , case_names Inheritance Overriding]: + (assumes subclseq_C_D: "G\C \\<^sub>C D" and + is_cls_D: "is_class G D" and + wf: "wf_prog G" and + old: "methd G D sig = Some old" and + accmodi_old: "Protected \ accmodi old" and + not_static_old: "\ is_static old" and + inheritance: "methd G C sig = Some old \ P" and + overriding: "\ new. \methd G C sig = Some new; + G,sig\new overrides\<^sub>S old\ \ P" + + ) "P" +proof - +from subclseq_C_D is_cls_D wf old accmodi_old not_static_old + inheritance overriding +show ?thesis + by (auto dest: inheritable_instance_methd) +qed + +lemma inheritable_instance_methd_props: + (assumes subclseq_C_D: "G\C \\<^sub>C D" and + is_cls_D: "is_class G D" and + wf: "wf_prog G" and + old: "methd G D sig = Some old" and + accmodi_old: "Protected \ accmodi old" and + not_static_old: "\ is_static old" + ) + "\new. methd G C sig = Some new \ + \ is_static new \ G\resTy new\resTy old \ accmodi old \accmodi new" + (is "(\new. (?Constraint C new old))") +proof - + from subclseq_C_D is_cls_D wf old accmodi_old not_static_old + show ?thesis + proof (cases rule: inheritable_instance_methd_cases) + case Inheritance + with not_static_old accmodi_old show ?thesis by auto + next + case (Overriding new) + then have "\ is_static new" by (auto dest: stat_overrides_commonD) + with Overriding not_static_old accmodi_old wf + show ?thesis + by (auto dest!: wf_prog_stat_overridesD + intro: order_trans) + qed +qed + +(* ### Probleme: Die tollen methd_subcls_cases Lemma wird warscheinlich + kaum gebraucht: +Redundanz: stat_overrides.Direct old declared in declclass old folgt aus + member of + Problem: Predikate wie overrides, sind global üper die Hierarchie hinweg + definiert, aber oft barucht man eben zusätlich Induktion + : von super c auf C; Dann ist aber auss dem Kontext + die Unterscheidung in die 5 fälle overkill, + da man dann warscheinlich meistens eh in einem speziellen + Fall kommt (durch die Hypothesen) +*) + +(* local lemma *) +ML {* bind_thm("bexI'",permute_prems 0 1 bexI) *} +ML {* bind_thm("ballE'",permute_prems 1 1 ballE) *} +lemma subint_widen_imethds: + "\G\I\I J; wf_prog G; is_iface G J; jm \ imethds G J sig\ \ + \ im \ imethds G I sig. is_static im = is_static jm \ + accmodi im = accmodi jm \ + G\resTy im\resTy jm" +proof - + assume irel: "G\I\I J" and + wf: "wf_prog G" and + is_iface: "is_iface G J" + from irel show "jm \ imethds G J sig \ ?thesis" + (is "PROP ?P I" is "PROP ?Prem J \ ?Concl I") + proof (induct ?P I rule: converse_rtrancl_induct) + case Id + assume "jm \ imethds G J sig" + then show "?Concl J" by (blast elim: bexI') + next + case Step + fix I SI + assume subint1_I_SI: "G\I \I1 SI" and + subint_SI_J: "G\SI \I J" and + hyp: "PROP ?P SI" and + jm: "jm \ imethds G J sig" + from subint1_I_SI + obtain i where + ifI: "iface G I = Some i" and + SI: "SI \ set (isuperIfs i)" + by (blast dest: subint1D) + + let ?newMethods + = "(o2s \ table_of (map (\(sig, mh). (sig, I, mh)) (imethods i)))" + show "?Concl I" + proof (cases "?newMethods sig = {}") + case True + with ifI SI hyp wf jm + show "?thesis" + by (auto simp add: imethds_rec) + next + case False + from ifI wf False + have imethds: "imethds G I sig = ?newMethods sig" + by (simp add: imethds_rec) + from False + obtain im where + imdef: "im \ ?newMethods sig" + by (blast) + with imethds + have im: "im \ imethds G I sig" + by (blast) + with im wf ifI + obtain + imStatic: "\ is_static im" and + imPublic: "accmodi im = Public" + by (auto dest!: imethds_wf_mhead) + from ifI wf + have wf_I: "wf_idecl G (I,i)" + by (rule wf_prog_idecl) + with SI wf + obtain si where + ifSI: "iface G SI = Some si" and + wf_SI: "wf_idecl G (SI,si)" + by (auto dest!: wf_idecl_supD is_acc_ifaceD + dest: wf_prog_idecl) + from jm hyp + obtain sim::"qtname \ mhead" where + sim: "sim \ imethds G SI sig" and + eq_static_sim_jm: "is_static sim = is_static jm" and + eq_access_sim_jm: "accmodi sim = accmodi jm" and + resTy_widen_sim_jm: "G\resTy sim\resTy jm" + by blast + with wf_I SI imdef sim + have "G\resTy im\resTy sim" + by (auto dest!: wf_idecl_hidings hidings_entailsD) + with wf resTy_widen_sim_jm + have resTy_widen_im_jm: "G\resTy im\resTy jm" + by (blast intro: widen_trans) + from sim wf ifSI + obtain + simStatic: "\ is_static sim" and + simPublic: "accmodi sim = Public" + by (auto dest!: imethds_wf_mhead) + from im + imStatic simStatic eq_static_sim_jm + imPublic simPublic eq_access_sim_jm + resTy_widen_im_jm + show ?thesis + by auto + qed + qed +qed + +(* Tactical version *) +(* +lemma subint_widen_imethds: "\G\I\I J; wf_prog G; is_iface G J\ \ + \ jm \ imethds G J sig. + \ im \ imethds G I sig. static (mthd im)=static (mthd jm) \ + access (mthd im)= access (mthd jm) \ + G\resTy (mthd im)\resTy (mthd jm)" +apply (erule converse_rtrancl_induct) +apply (clarsimp elim!: bexI') +apply (frule subint1D) +apply clarify +apply (erule ballE') +apply fast +apply (erule_tac V = "?x \ imethds G J sig" in thin_rl) +apply clarsimp +apply (subst imethds_rec, assumption, erule wf_ws_prog) +apply (unfold overrides_t_def) +apply (drule (1) wf_prog_idecl) +apply (frule (3) imethds_wf_mhead [OF _ _ wf_idecl_supD [THEN conjunct1 + [THEN is_acc_ifaceD [THEN conjunct1]]]]) +apply (case_tac "(o2s \ table_of (map (\(s, mh). (s, y, mh)) (imethods i))) + sig ={}") +apply force + +apply (simp only:) +apply (simp) +apply clarify +apply (frule wf_idecl_hidings [THEN hidings_entailsD]) +apply blast +apply blast +apply (rule bexI') +apply simp +apply (drule table_of_map_SomeI [of _ "sig"]) +apply simp + +apply (frule wf_idecl_mhead [of _ _ _ "sig"]) +apply (rule table_of_Some_in_set) +apply assumption +apply auto +done +*) + + +(* local lemma *) +lemma implmt1_methd: + "\sig. \G\C\1I; wf_prog G; im \ imethds G I sig\ \ + \cm \methd G C sig: \ is_static cm \ \ is_static im \ + G\resTy cm\resTy im \ + accmodi im = Public \ accmodi cm = Public" +apply (drule implmt1D) +apply clarify +apply (drule (2) wf_prog_cdecl [THEN wf_cdecl_impD]) +apply (frule (1) imethds_wf_mhead) +apply (simp add: is_acc_iface_def) +apply (force) +done + + +(* local lemma *) +lemma implmt_methd [rule_format (no_asm)]: +"\wf_prog G; G\C\I\ \ is_iface G I \ + (\ im \imethds G I sig. + \ cm\methd G C sig: \is_static cm \ \ is_static im \ + G\resTy cm\resTy im \ + accmodi im = Public \ accmodi cm = Public)" +apply (frule implmt_is_class) +apply (erule implmt.induct) +apply safe +apply (drule (2) implmt1_methd) +apply fast +apply (drule (1) subint_widen_imethds) +apply simp +apply assumption +apply clarify +apply (drule (2) implmt1_methd) +apply (force) +apply (frule subcls1D) +apply (drule (1) bspec) +apply clarify +apply (drule (3) r_into_rtrancl [THEN inheritable_instance_methd_props, + OF _ implmt_is_class]) +apply auto +done + + +declare split_paired_All [simp] split_paired_Ex [simp] +ML_setup {* +claset_ref() := claset() addSbefore ("split_all_tac", split_all_tac); +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac) +*} +lemma mheadsD [rule_format (no_asm)]: +"emh \ mheads G S t sig \ wf_prog G \ + (\C D m. t = ClassT C \ declrefT emh = ClassT D \ + accmethd G S C sig = Some m \ + (declclass m = D) \ mhead (mthd m) = (mhd emh)) \ + (\I. t = IfaceT I \ ((\im. im \ accimethds G (pid S) I sig \ + mthd im = mhd emh) \ + (\m. G\Iface I accessible_in (pid S) \ accmethd G S Object sig = Some m \ + accmodi m \ Private \ + declrefT emh = ClassT Object \ mhead (mthd m) = mhd emh))) \ + (\T m. t = ArrayT T \ G\Array T accessible_in (pid S) \ + accmethd G S Object sig = Some m \ accmodi m \ Private \ + declrefT emh = ClassT Object \ mhead (mthd m) = mhd emh)" +apply (rule_tac "ref_ty1"="t" in ref_ty_ty.induct [THEN conjunct1]) +apply auto +apply (auto simp add: cmheads_def accObjectmheads_def Objectmheads_def) +apply (auto dest!: accmethd_SomeD) +done + +lemma mheads_cases [consumes 2, case_names Class_methd + Iface_methd Iface_Object_methd Array_Object_methd]: +"\emh \ mheads G S t sig; wf_prog G; + \ C D m. \t = ClassT C;declrefT emh = ClassT D; accmethd G S C sig = Some m; + (declclass m = D); mhead (mthd m) = (mhd emh)\ \ P emh; + \ I im. \t = IfaceT I; im \ accimethds G (pid S) I sig; mthd im = mhd emh\ + \ P emh; + \ I m. \t = IfaceT I; G\Iface I accessible_in (pid S); + accmethd G S Object sig = Some m; accmodi m \ Private; + declrefT emh = ClassT Object; mhead (mthd m) = mhd emh\ \ P emh; + \ T m. \t = ArrayT T;G\Array T accessible_in (pid S); + accmethd G S Object sig = Some m; accmodi m \ Private; + declrefT emh = ClassT Object; mhead (mthd m) = mhd emh\ \ P emh +\ \ P emh" +by (blast dest!: mheadsD) + +lemma declclassD[rule_format]: + "\wf_prog G;class G C = Some c; methd G C sig = Some m; + class G (declclass m) = Some d\ + \ table_of (methods d) sig = Some (mthd m)" +proof - + assume wf: "wf_prog G" + then have ws: "ws_prog G" .. + assume clsC: "class G C = Some c" + from clsC ws + show "\ m d. \methd G C sig = Some m; class G (declclass m) = Some d\ + \ table_of (methods d) sig = Some (mthd m)" + (is "PROP ?P C") + proof (induct ?P C rule: ws_class_induct) + case Object + fix m d + assume "methd G Object sig = Some m" + "class G (declclass m) = Some d" + with wf show "?thesis m d" by auto + next + case Subcls + fix C c m d + assume hyp: "PROP ?P (super c)" + and m: "methd G C sig = Some m" + and declC: "class G (declclass m) = Some d" + and clsC: "class G C = Some c" + and nObj: "C \ Object" + let ?newMethods = "table_of (map (\(s, m). (s, C, m)) (methods c)) sig" + show "?thesis m d" + proof (cases "?newMethods") + case None + from None clsC nObj ws m declC hyp + show "?thesis" by (auto simp add: methd_rec) + next + case Some + from Some clsC nObj ws m declC hyp + show "?thesis" + by (auto simp add: methd_rec + dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class) + qed + qed +qed + +(* Tactical version *) +(* +lemma declclassD[rule_format]: + "wf_prog G \ + (\ c d m. class G C = Some c \ methd G C sig = Some m \ + class G (declclass m) = Some d + \ table_of (methods d) sig = Some (mthd m))" +apply (rule class_rec.induct) +apply (rule impI) +apply (rule allI)+ +apply (rule impI) +apply (case_tac "C=Object") +apply (force simp add: methd_rec) + +apply (subst methd_rec) +apply (blast dest: wf_ws_prog)+ +apply (case_tac "table_of (map (\(s, m). (s, C, m)) (methods c)) sig") +apply (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class) +done +*) + +lemma dynmethd_Object: + (assumes statM: "methd G Object sig = Some statM" and + private: "accmodi statM = Private" and + is_cls_C: "is_class G C" and + wf: "wf_prog G" + ) "dynmethd G Object C sig = Some statM" +proof - + from is_cls_C wf + have subclseq: "G\C \\<^sub>C Object" + by (auto intro: subcls_ObjectI) + from wf have ws: "ws_prog G" + by simp + from wf + have is_cls_Obj: "is_class G Object" + by simp + from statM subclseq is_cls_Obj ws private + show ?thesis + proof (cases rule: dynmethd_cases) + case Static then show ?thesis . + next + case Overrides + with private show ?thesis + by (auto dest: no_Private_override) + qed +qed + +lemma wf_imethds_hiding_objmethdsD: + (assumes old: "methd G Object sig = Some old" and + is_if_I: "is_iface G I" and + wf: "wf_prog G" and + not_private: "accmodi old \ Private" and + new: "new \ imethds G I sig" + ) "G\resTy new\resTy old \ is_static new = is_static old" (is "?P new") +proof - + from wf have ws: "ws_prog G" by simp + { + fix I i new + assume ifI: "iface G I = Some i" + assume new: "table_of (imethods i) sig = Some new" + from ifI new not_private wf old + have "?P (I,new)" + by (auto dest!: wf_prog_idecl wf_idecl_hiding cond_hiding_entailsD + simp del: methd_Object) + } note hyp_newmethod = this + from is_if_I ws new + show ?thesis + proof (induct rule: ws_interface_induct) + case (Step I i) + assume ifI: "iface G I = Some i" + assume new: "new \ imethds G I sig" + from Step + have hyp: "\ J \ set (isuperIfs i). (new \ imethds G J sig \ ?P new)" + by auto + from new ifI ws + show "?P new" + proof (cases rule: imethds_cases) + case NewMethod + with ifI hyp_newmethod + show ?thesis + by auto + next + case (InheritedMethod J) + assume "J \ set (isuperIfs i)" + "new \ imethds G J sig" + with hyp + show "?thesis" + by auto + qed + qed +qed + +text {* +Which dynamic classes are valid to look up a member of a distinct static type? +We have to distinct class members (named static members in Java) +from instance members. Class members are global to all Objects of a class, +instance members are local to a single Object instance. If a member is +equipped with the static modifier it is a class member, else it is an +instance member. +The following table gives an overview of the current framework. We assume +to have a reference with static type statT and a dynamic class dynC. Between +both of these types the widening relation holds +@{term "G\Class dynC\ statT"}. Unfortunately this ordinary widening relation +isn't enough to describe the valid lookup classes, since we must cope the +special cases of arrays and interfaces,too. If we statically expect an array or +inteface we may lookup a field or a method in Object which isn't covered in +the widening relation. +\begin{verbatim} +statT field instance method static (class) method +------------------------------------------------------------------------ + NullT / / / + Iface / dynC Object + Class dynC dynC dynC + Array / Object Object +\end{verbatim} +In most cases we con lookup the member in the dynamic class. But as an +interface can't declare new static methods, nor an array can define new +methods at all, we have to lookup methods in the base class Object. + +The limitation to classes in the field column is artificial and comes out +of the typing rule for the field access (see rule @{text "FVar"} in the +welltyping relation @{term "wt"} in theory WellType). +I stems out of the fact, that Object +indeed has no non private fields. So interfaces and arrays can actually +have no fields at all and a field access would be senseless. (In Java +interfaces are allowed to declare new fields but in current Bali not!). +So there is no principal reason why we should not allow Objects to declare +non private fields. Then we would get the following column: +\begin{verbatim} + statT field +----------------- + NullT / + Iface Object + Class dynC + Array Object +\end{verbatim} +*} +consts valid_lookup_cls:: "prog \ ref_ty \ qtname \ bool \ bool" + ("_,_ \ _ valid'_lookup'_cls'_for _" [61,61,61,61] 60) +primrec +"G,NullT \ dynC valid_lookup_cls_for static_membr = False" +"G,IfaceT I \ dynC valid_lookup_cls_for static_membr + = (if static_membr + then dynC=Object + else G\Class dynC\ Iface I)" +"G,ClassT C \ dynC valid_lookup_cls_for static_membr = G\Class dynC\ Class C" +"G,ArrayT T \ dynC valid_lookup_cls_for static_membr = (dynC=Object)" + +lemma valid_lookup_cls_is_class: + (assumes dynC: "G,statT \ dynC valid_lookup_cls_for static_membr" and + ty_statT: "isrtype G statT" and + wf: "wf_prog G" + ) "is_class G dynC" +proof (cases statT) + case NullT + with dynC ty_statT show ?thesis + by (auto dest: widen_NT2) +next + case (IfaceT I) + with dynC wf show ?thesis + by (auto dest: implmt_is_class) +next + case (ClassT C) + with dynC ty_statT show ?thesis + by (auto dest: subcls_is_class2) +next + case (ArrayT T) + with dynC wf show ?thesis + by (auto) +qed + +declare split_paired_All [simp del] split_paired_Ex [simp del] +ML_setup {* +simpset_ref() := simpset() delloop "split_all_tac"; +claset_ref () := claset () delSWrapper "split_all_tac" +*} + +lemma dynamic_mheadsD: +"\emh \ mheads G S statT sig; + G,statT \ dynC valid_lookup_cls_for (is_static emh); + isrtype G statT; wf_prog G + \ \ \m \ dynlookup G statT dynC sig: + is_static m=is_static emh \ G\resTy m\resTy emh" +proof - + assume emh: "emh \ mheads G S statT sig" + and wf: "wf_prog G" + and dynC_Prop: "G,statT \ dynC valid_lookup_cls_for (is_static emh)" + and istype: "isrtype G statT" + from dynC_Prop istype wf + obtain y where + dynC: "class G dynC = Some y" + by (auto dest: valid_lookup_cls_is_class) + from emh wf show ?thesis + proof (cases rule: mheads_cases) + case Class_methd + fix statC statDeclC sm + assume statC: "statT = ClassT statC" + assume "accmethd G S statC sig = Some sm" + then have sm: "methd G statC sig = Some sm" + by (blast dest: accmethd_SomeD) + assume eq_mheads: "mhead (mthd sm) = mhd emh" + from statC + have dynlookup: "dynlookup G statT dynC sig = dynmethd G statC dynC sig" + by (simp add: dynlookup_def) + from wf statC istype dynC_Prop sm + obtain dm where + "dynmethd G statC dynC sig = Some dm" + "is_static dm = is_static sm" + "G\resTy dm\resTy sm" + by (auto dest!: ws_dynmethd accmethd_SomeD) + with dynlookup eq_mheads + show ?thesis + by (cases emh type: *) (auto) + next + case Iface_methd + fix I im + assume statI: "statT = IfaceT I" and + eq_mheads: "mthd im = mhd emh" and + "im \ accimethds G (pid S) I sig" + then have im: "im \ imethds G I sig" + by (blast dest: accimethdsD) + with istype statI eq_mheads wf + have not_static_emh: "\ is_static emh" + by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead) + from statI im + have dynlookup: "dynlookup G statT dynC sig = methd G dynC sig" + by (auto simp add: dynlookup_def dynimethd_def) + from wf dynC_Prop statI istype im not_static_emh + obtain dm where + "methd G dynC sig = Some dm" + "is_static dm = is_static im" + "G\resTy (mthd dm)\resTy (mthd im)" + by (auto dest: implmt_methd) + with dynlookup eq_mheads + show ?thesis + by (cases emh type: *) (auto) + next + case Iface_Object_methd + fix I sm + assume statI: "statT = IfaceT I" and + sm: "accmethd G S Object sig = Some sm" and + eq_mheads: "mhead (mthd sm) = mhd emh" and + nPriv: "accmodi sm \ Private" + show ?thesis + proof (cases "imethds G I sig = {}") + case True + with statI + have dynlookup: "dynlookup G statT dynC sig = dynmethd G Object dynC sig" + by (simp add: dynlookup_def dynimethd_def) + from wf dynC + have subclsObj: "G\dynC \\<^sub>C Object" + by (auto intro: subcls_ObjectI) + from wf dynC dynC_Prop istype sm subclsObj + obtain dm where + "dynmethd G Object dynC sig = Some dm" + "is_static dm = is_static sm" + "G\resTy (mthd dm)\resTy (mthd sm)" + by (auto dest!: ws_dynmethd accmethd_SomeD + intro: class_Object [OF wf]) + with dynlookup eq_mheads + show ?thesis + by (cases emh type: *) (auto) + next + case False + with statI + have dynlookup: "dynlookup G statT dynC sig = methd G dynC sig" + by (simp add: dynlookup_def dynimethd_def) + from istype statI + have "is_iface G I" + by auto + with wf sm nPriv False + obtain im where + im: "im \ imethds G I sig" and + eq_stat: "is_static im = is_static sm" and + resProp: "G\resTy (mthd im)\resTy (mthd sm)" + by (auto dest: wf_imethds_hiding_objmethdsD accmethd_SomeD) + from im wf statI istype eq_stat eq_mheads + have not_static_sm: "\ is_static emh" + by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead) + from im wf dynC_Prop dynC istype statI not_static_sm + obtain dm where + "methd G dynC sig = Some dm" + "is_static dm = is_static im" + "G\resTy (mthd dm)\resTy (mthd im)" + by (auto dest: implmt_methd) + with wf eq_stat resProp dynlookup eq_mheads + show ?thesis + by (cases emh type: *) (auto intro: widen_trans) + qed + next + case Array_Object_methd + fix T sm + assume statArr: "statT = ArrayT T" and + sm: "accmethd G S Object sig = Some sm" and + eq_mheads: "mhead (mthd sm) = mhd emh" + from statArr dynC_Prop wf + have dynlookup: "dynlookup G statT dynC sig = methd G Object sig" + by (auto simp add: dynlookup_def dynmethd_C_C) + with sm eq_mheads sm + show ?thesis + by (cases emh type: *) (auto dest: accmethd_SomeD) + qed +qed + +(* Tactical version *) +(* +lemma dynamic_mheadsD: " + \emh \ mheads G S statT sig; wf_prog G; class G dynC = Some y; + if (\T. statT=ArrayT T) then dynC=Object else G\Class dynC\RefT statT; + isrtype G statT\ \ + \m \ dynlookup G statT dynC sig: + static (mthd m)=static (mhd emh) \ G\resTy (mthd m)\resTy (mhd emh)" +apply (drule mheadsD) +apply safe + -- reftype statT is a class +apply (case_tac "\T. ClassT C = ArrayT T") +apply (simp) + +apply (clarsimp simp add: dynlookup_def ) +apply (frule_tac statC="C" and dynC="dynC" and sig="sig" + in ws_dynmethd) +apply assumption+ +apply (case_tac "emh") +apply (force dest: accmethd_SomeD) + + -- reftype statT is a interface, method defined in interface +apply (clarsimp simp add: dynlookup_def) +apply (drule (1) implmt_methd) +apply blast +apply blast +apply (clarify) +apply (unfold dynimethd_def) +apply (rule_tac x="cm" in bexI) +apply (case_tac "emh") +apply force + +apply force + + -- reftype statT is a interface, method defined in Object +apply (simp add: dynlookup_def) +apply (simp only: dynimethd_def) +apply (case_tac "imethds G I sig = {}") +apply simp +apply (frule_tac statC="Object" and dynC="dynC" and sig="sig" + in ws_dynmethd) +apply (blast intro: subcls_ObjectI wf_ws_prog) +apply (blast dest: class_Object) +apply (case_tac "emh") +apply (force dest: accmethd_SomeD) + +apply simp +apply (subgoal_tac "\ im. im \ imethds G I sig") +prefer 2 apply blast +apply clarify +apply (frule (1) implmt_methd) +apply simp +apply blast +apply (clarify dest!: accmethd_SomeD) +apply (frule (4) iface_overrides_Object) +apply clarify +apply (case_tac emh) +apply force + + -- reftype statT is a array +apply (simp add: dynlookup_def) +apply (case_tac emh) +apply (force dest: accmethd_SomeD simp add: dynmethd_def) +done +*) + +(* ### auf ws_class_induct umstellen *) +lemma methd_declclass: +"\class G C = Some c; wf_prog G; methd G C sig = Some m\ + \ methd G (declclass m) sig = Some m" +proof - + assume asm: "class G C = Some c" "wf_prog G" "methd G C sig = Some m" + have "wf_prog G \ + (\ c m. class G C = Some c \ methd G C sig = Some m + \ methd G (declclass m) sig = Some m)" (is "?P G C") + proof (rule class_rec.induct,intro allI impI) + fix G C c m + assume hyp: "\c. C \ Object \ ws_prog G \ class G C = Some c \ + ?P G (super c)" + assume wf: "wf_prog G" and cls_C: "class G C = Some c" and + m: "methd G C sig = Some m" + show "methd G (declclass m) sig = Some m" + proof (cases "C=Object") + case True + with wf m show ?thesis by (auto intro: table_of_map_SomeI) + next + let ?filter="filter_tab (\sig m. G\C inherits method sig m)" + let ?table = "table_of (map (\(s, m). (s, C, m)) (methods c))" + case False + with cls_C wf m + have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m " + by (simp add: methd_rec) + show ?thesis + proof (cases "?table sig") + case None + from this methd_C have "?filter (methd G (super c)) sig = Some m" + by simp + moreover + from wf cls_C False obtain sup where "class G (super c) = Some sup" + by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class) + moreover note wf False cls_C hyp + ultimately show ?thesis by auto + next + case Some + from this methd_C m show ?thesis by auto + qed + qed + qed + with asm show ?thesis by auto +qed + +lemma dynmethd_declclass: + "\dynmethd G statC dynC sig = Some m; + wf_prog G; is_class G statC + \ \ methd G (declclass m) sig = Some m" +by (auto dest: dynmethd_declC) + +lemma dynlookup_declC: + "\dynlookup G statT dynC sig = Some m; wf_prog G; + is_class G dynC;isrtype G statT + \ \ G\dynC \\<^sub>C (declclass m) \ is_class G (declclass m)" +by (cases "statT") + (auto simp add: dynlookup_def dynimethd_def + dest: methd_declC dynmethd_declC) + +lemma dynlookup_Array_declclassD [simp]: +"\dynlookup G (ArrayT T) Object sig = Some dm;wf_prog G\ + \ declclass dm = Object" +proof - + assume dynL: "dynlookup G (ArrayT T) Object sig = Some dm" + assume wf: "wf_prog G" + from wf have ws: "ws_prog G" by auto + from wf have is_cls_Obj: "is_class G Object" by auto + from dynL wf + show ?thesis + by (auto simp add: dynlookup_def dynmethd_C_C [OF is_cls_Obj ws] + dest: methd_Object_SomeD) +qed + +declare split_paired_All [simp del] split_paired_Ex [simp del] +ML_setup {* +simpset_ref() := simpset() delloop "split_all_tac"; +claset_ref () := claset () delSWrapper "split_all_tac" +*} + +lemma wt_is_type: "E,dt\v\T \ wf_prog (prg E) \ + dt=empty_dt \ (case T of + Inl T \ is_type (prg E) T + | Inr Ts \ Ball (set Ts) (is_type (prg E)))" +apply (unfold empty_dt_def) +apply (erule wt.induct) +apply (auto split del: split_if_asm simp del: snd_conv + simp add: is_acc_class_def is_acc_type_def) +apply (erule typeof_empty_is_type) +apply (frule (1) wf_prog_cdecl [THEN wf_cdecl_supD], rotate_tac -1, + force simp del: snd_conv, clarsimp simp add: is_acc_class_def) +apply (drule (1) max_spec2mheads [THEN conjunct1, THEN mheadsD]) +apply (drule_tac [2] accfield_fields) +apply (frule class_Object) +apply (auto dest: accmethd_rT_is_type + imethds_wf_mhead [THEN conjunct1, THEN rT_is_acc_type] + dest!:accimethdsD + simp del: class_Object + simp add: is_acc_type_def + ) +done +declare split_paired_All [simp] split_paired_Ex [simp] +ML_setup {* +claset_ref() := claset() addSbefore ("split_all_tac", split_all_tac); +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac) +*} + +lemma ty_expr_is_type: +"\E\e\-T; wf_prog (prg E)\ \ is_type (prg E) T" +by (auto dest!: wt_is_type) +lemma ty_var_is_type: +"\E\v\=T; wf_prog (prg E)\ \ is_type (prg E) T" +by (auto dest!: wt_is_type) +lemma ty_exprs_is_type: +"\E\es\\Ts; wf_prog (prg E)\ \ Ball (set Ts) (is_type (prg E))" +by (auto dest!: wt_is_type) + + +lemma static_mheadsD: + "\ emh \ mheads G S t sig; wf_prog G; E\e\-RefT t; prg E=G ; + invmode (mhd emh) e \ IntVir + \ \ \m. ( (\ C. t = ClassT C \ accmethd G S C sig = Some m) + \ (\ C. t \ ClassT C \ accmethd G S Object sig = Some m )) \ + declrefT emh = ClassT (declclass m) \ mhead (mthd m) = (mhd emh)" +apply (subgoal_tac "is_static emh \ e = Super") +defer apply (force simp add: invmode_def) +apply (frule ty_expr_is_type) +apply simp +apply (case_tac "is_static emh") +apply (frule (1) mheadsD) +apply clarsimp +apply safe +apply blast +apply (auto dest!: imethds_wf_mhead + accmethd_SomeD + accimethdsD + simp add: accObjectmheads_def Objectmheads_def) + +apply (erule wt_elim_cases) +apply (force simp add: cmheads_def) +done + +lemma wt_MethdI: +"\methd G C sig = Some m; wf_prog G; + class G C = Some c\ \ + \T. \prg=G,cls=(declclass m), + lcl=\ k. + (case k of + EName e + \ (case e of + VNam v + \ (table_of (lcls (mbody (mthd m))) + ((pars (mthd m))[\](parTs sig))) v + | Res \ Some (resTy (mthd m))) + | This \ if is_static m then None else Some (Class (declclass m))) + \\ Methd C sig\-T \ G\T\resTy m" +apply (frule (2) methd_wf_mdecl, clarify) +apply (force dest!: wf_mdecl_bodyD intro!: wt.Methd) +done + +subsection "accessibility concerns" + +lemma mheads_type_accessible: + "\emh \ mheads G S T sig; wf_prog G\ + \ G\RefT T accessible_in (pid S)" +by (erule mheads_cases) + (auto dest: accmethd_SomeD accessible_from_commonD accimethdsD) + +lemma static_to_dynamic_accessible_from: +"\G\m of C accessible_from accC;wf_prog G\ + \ G\m in C dyn_accessible_from accC" +proof (induct rule: accessible_fromR.induct) +qed (auto intro: dyn_accessible_fromR.intros + member_of_to_member_in + static_to_dynamic_overriding) + +lemma static_to_dynamic_accessible_from: + (assumes stat_acc: "G\m of statC accessible_from accC" and + subclseq: "G\dynC \\<^sub>C statC" and + wf: "wf_prog G" + ) "G\m in dynC dyn_accessible_from accC" +proof - + from stat_acc subclseq + show ?thesis (is "?Dyn_accessible m") + proof (induct rule: accessible_fromR.induct) + case (immediate statC m) + then show "?Dyn_accessible m" + by (blast intro: dyn_accessible_fromR.immediate + member_inI + permits_acc_inheritance) + next + case (overriding _ _ m) + with wf show "?Dyn_accessible m" + by (blast intro: dyn_accessible_fromR.overriding + member_inI + static_to_dynamic_overriding + rtrancl_trancl_trancl + static_to_dynamic_accessible_from) + qed +qed + +lemma dynmethd_member_in: + (assumes m: "dynmethd G statC dynC sig = Some m" and + iscls_statC: "is_class G statC" and + wf: "wf_prog G" + ) "G\Methd sig m member_in dynC" +proof - + from m + have subclseq: "G\dynC \\<^sub>C statC" + by (auto simp add: dynmethd_def) + from subclseq iscls_statC + have iscls_dynC: "is_class G dynC" + by (rule subcls_is_class2) + from iscls_dynC iscls_statC wf m + have "G\dynC \\<^sub>C (declclass m) \ is_class G (declclass m) \ + methd G (declclass m) sig = Some m" + by - (drule dynmethd_declC, auto) + with wf + show ?thesis + by (auto intro: member_inI dest: methd_member_of) +qed + +lemma dynmethd_access_prop: + (assumes statM: "methd G statC sig = Some statM" and + stat_acc: "G\Methd sig statM of statC accessible_from accC" and + dynM: "dynmethd G statC dynC sig = Some dynM" and + wf: "wf_prog G" + ) "G\Methd sig dynM in dynC dyn_accessible_from accC" +proof - + from wf have ws: "ws_prog G" .. + from dynM + have subclseq: "G\dynC \\<^sub>C statC" + by (auto simp add: dynmethd_def) + from stat_acc + have is_cls_statC: "is_class G statC" + by (auto dest: accessible_from_commonD member_of_is_classD) + with subclseq + have is_cls_dynC: "is_class G dynC" + by (rule subcls_is_class2) + from is_cls_statC statM wf + have member_statC: "G\Methd sig statM member_of statC" + by (auto intro: methd_member_of) + from stat_acc + have statC_acc: "G\Class statC accessible_in (pid accC)" + by (auto dest: accessible_from_commonD) + from statM subclseq is_cls_statC ws + show ?thesis + proof (cases rule: dynmethd_cases) + case Static + assume dynmethd: "dynmethd G statC dynC sig = Some statM" + with dynM have eq_dynM_statM: "dynM=statM" + by simp + with stat_acc subclseq wf + show ?thesis + by (auto intro: static_to_dynamic_accessible_from) + next + case (Overrides newM) + assume dynmethd: "dynmethd G statC dynC sig = Some newM" + assume override: "G,sig\newM overrides statM" + assume neq: "newM\statM" + from dynmethd dynM + have eq_dynM_newM: "dynM=newM" + by simp + from dynmethd eq_dynM_newM wf is_cls_statC + have "G\Methd sig dynM member_in dynC" + by (auto intro: dynmethd_member_in) + moreover + from subclseq + have "G\dynC\\<^sub>C statC" + proof (cases rule: subclseq_cases) + case Eq + assume "dynC=statC" + moreover + from is_cls_statC obtain c + where "class G statC = Some c" + by auto + moreover + note statM ws dynmethd + ultimately + have "newM=statM" + by (auto simp add: dynmethd_C_C) + with neq show ?thesis + by (contradiction) + next + case Subcls show ?thesis . + qed + moreover + from stat_acc wf + have "G\Methd sig statM in statC dyn_accessible_from accC" + by (blast intro: static_to_dynamic_accessible_from) + moreover + note override eq_dynM_newM + ultimately show ?thesis + by (cases dynM,cases statM) (auto intro: dyn_accessible_fromR.overriding) + qed +qed + +lemma implmt_methd_access: +(fixes accC::qtname + assumes iface_methd: "imethds G I sig \ {}" and + implements: "G\dynC\I" and + isif_I: "is_iface G I" and + wf: "wf_prog G" + ) "\ dynM. methd G dynC sig = Some dynM \ + G\Methd sig dynM in dynC dyn_accessible_from accC" +proof - + from implements + have iscls_dynC: "is_class G dynC" by (rule implmt_is_class) + from iface_methd + obtain im + where "im \ imethds G I sig" + by auto + with wf implements isif_I + obtain dynM + where dynM: "methd G dynC sig = Some dynM" and + pub: "accmodi dynM = Public" + by (blast dest: implmt_methd) + with iscls_dynC wf + have "G\Methd sig dynM in dynC dyn_accessible_from accC" + by (auto intro!: dyn_accessible_fromR.immediate + intro: methd_member_of member_of_to_member_in + simp add: permits_acc_def) + with dynM + show ?thesis + by blast +qed + +corollary implmt_dynimethd_access: +(fixes accC::qtname + assumes iface_methd: "imethds G I sig \ {}" and + implements: "G\dynC\I" and + isif_I: "is_iface G I" and + wf: "wf_prog G" + ) "\ dynM. dynimethd G I dynC sig = Some dynM \ + G\Methd sig dynM in dynC dyn_accessible_from accC" +proof - + from iface_methd + have "dynimethd G I dynC sig = methd G dynC sig" + by (simp add: dynimethd_def) + with iface_methd implements isif_I wf + show ?thesis + by (simp only:) + (blast intro: implmt_methd_access) +qed + +lemma dynlookup_access_prop: + (assumes emh: "emh \ mheads G accC statT sig" and + dynM: "dynlookup G statT dynC sig = Some dynM" and + dynC_prop: "G,statT \ dynC valid_lookup_cls_for is_static emh" and + isT_statT: "isrtype G statT" and + wf: "wf_prog G" + ) "G \Methd sig dynM in dynC dyn_accessible_from accC" +proof - + from emh wf + have statT_acc: "G\RefT statT accessible_in (pid accC)" + by (rule mheads_type_accessible) + from dynC_prop isT_statT wf + have iscls_dynC: "is_class G dynC" + by (rule valid_lookup_cls_is_class) + from emh dynC_prop isT_statT wf dynM + have eq_static: "is_static emh = is_static dynM" + by (auto dest: dynamic_mheadsD) + from emh wf show ?thesis + proof (cases rule: mheads_cases) + case (Class_methd statC _ statM) + assume statT: "statT = ClassT statC" + assume "accmethd G accC statC sig = Some statM" + then have statM: "methd G statC sig = Some statM" and + stat_acc: "G\Methd sig statM of statC accessible_from accC" + by (auto dest: accmethd_SomeD) + from dynM statT + have dynM': "dynmethd G statC dynC sig = Some dynM" + by (simp add: dynlookup_def) + from statM stat_acc wf dynM' + show ?thesis + by (auto dest!: dynmethd_access_prop) + next + case (Iface_methd I im) + then have iface_methd: "imethds G I sig \ {}" and + statT_acc: "G\RefT statT accessible_in (pid accC)" + by (auto dest: accimethdsD) + assume statT: "statT = IfaceT I" + assume im: "im \ accimethds G (pid accC) I sig" + assume eq_mhds: "mthd im = mhd emh" + from dynM statT + have dynM': "dynimethd G I dynC sig = Some dynM" + by (simp add: dynlookup_def) + from isT_statT statT + have isif_I: "is_iface G I" + by simp + show ?thesis + proof (cases "is_static emh") + case False + with statT dynC_prop + have widen_dynC: "G\Class dynC \ RefT statT" + by simp + from statT widen_dynC + have implmnt: "G\dynC\I" + by auto + from eq_static False + have not_static_dynM: "\ is_static dynM" + by simp + from iface_methd implmnt isif_I wf dynM' + show ?thesis + by - (drule implmt_dynimethd_access, auto) + next + case True + assume "is_static emh" + moreover + from wf isT_statT statT im + have "\ is_static im" + by (auto dest: accimethdsD wf_prog_idecl imethds_wf_mhead) + moreover note eq_mhds + ultimately show ?thesis + by (cases emh) auto + qed + next + case (Iface_Object_methd I statM) + assume statT: "statT = IfaceT I" + assume "accmethd G accC Object sig = Some statM" + then have statM: "methd G Object sig = Some statM" and + stat_acc: "G\Methd sig statM of Object accessible_from accC" + by (auto dest: accmethd_SomeD) + assume not_Private_statM: "accmodi statM \ Private" + assume eq_mhds: "mhead (mthd statM) = mhd emh" + from iscls_dynC wf + have widen_dynC_Obj: "G\dynC \\<^sub>C Object" + by (auto intro: subcls_ObjectI) + show ?thesis + proof (cases "imethds G I sig = {}") + case True + from dynM statT True + have dynM': "dynmethd G Object dynC sig = Some dynM" + by (simp add: dynlookup_def dynimethd_def) + from statT + have "G\RefT statT \Class Object" + by auto + with statM statT_acc stat_acc widen_dynC_Obj statT isT_statT + wf dynM' eq_static dynC_prop + show ?thesis + by - (drule dynmethd_access_prop,force+) + next + case False + then obtain im where + im: "im \ imethds G I sig" + by auto + have not_static_emh: "\ is_static emh" + proof - + from im statM statT isT_statT wf not_Private_statM + have "is_static im = is_static statM" + by (auto dest: wf_imethds_hiding_objmethdsD) + with wf isT_statT statT im + have "\ is_static statM" + by (auto dest: wf_prog_idecl imethds_wf_mhead) + with eq_mhds + show ?thesis + by (cases emh) auto + qed + with statT dynC_prop + have implmnt: "G\dynC\I" + by simp + with isT_statT statT + have isif_I: "is_iface G I" + by simp + from dynM statT + have dynM': "dynimethd G I dynC sig = Some dynM" + by (simp add: dynlookup_def) + from False implmnt isif_I wf dynM' + show ?thesis + by - (drule implmt_dynimethd_access, auto) + qed + next + case (Array_Object_methd T statM) + assume statT: "statT = ArrayT T" + assume "accmethd G accC Object sig = Some statM" + then have statM: "methd G Object sig = Some statM" and + stat_acc: "G\Methd sig statM of Object accessible_from accC" + by (auto dest: accmethd_SomeD) + from statT dynC_prop + have dynC_Obj: "dynC = Object" + by simp + then + have widen_dynC_Obj: "G\Class dynC \ Class Object" + by simp + from dynM statT + have dynM': "dynmethd G Object dynC sig = Some dynM" + by (simp add: dynlookup_def) + from statM statT_acc stat_acc dynM' wf widen_dynC_Obj + statT isT_statT + show ?thesis + by - (drule dynmethd_access_prop, simp+) + qed +qed + +lemma dynlookup_access: + (assumes emh: "emh \ mheads G accC statT sig" and + dynC_prop: "G,statT \ dynC valid_lookup_cls_for (is_static emh) " and + isT_statT: "isrtype G statT" and + wf: "wf_prog G" + ) "\ dynM. dynlookup G statT dynC sig = Some dynM \ + G\Methd sig dynM in dynC dyn_accessible_from accC" +proof - + from dynC_prop isT_statT wf + have is_cls_dynC: "is_class G dynC" + by (auto dest: valid_lookup_cls_is_class) + with emh wf dynC_prop isT_statT + obtain dynM where + "dynlookup G statT dynC sig = Some dynM" + by - (drule dynamic_mheadsD,auto) + with emh dynC_prop isT_statT wf + show ?thesis + by (blast intro: dynlookup_access_prop) +qed + +lemma stat_overrides_Package_old: +(assumes stat_override: "G \ new overrides\<^sub>S old" and + accmodi_new: "accmodi new = Package" and + wf: "wf_prog G " +) "accmodi old = Package" +proof - + from stat_override wf + have "accmodi old \ accmodi new" + by (auto dest: wf_prog_stat_overridesD) + with stat_override accmodi_new show ?thesis + by (cases "accmodi old") (auto dest: no_Private_stat_override + dest: acc_modi_le_Dests) +qed + +text {* @{text dyn_accessible_Package} only works with the @{text wf_prog} assumption. +Without it. it is easy to leaf the Package! +*} +lemma dyn_accessible_Package: + "\G \ m in C dyn_accessible_from accC; accmodi m = Package; + wf_prog G\ + \ pid accC = pid (declclass m)" +proof - + assume wf: "wf_prog G " + assume accessible: "G \ m in C dyn_accessible_from accC" + then show "accmodi m = Package + \ pid accC = pid (declclass m)" + (is "?Pack m \ ?P m") + proof (induct rule: dyn_accessible_fromR.induct) + case (immediate C m) + assume "G\m member_in C" + "G \ m in C permits_acc_to accC" + "accmodi m = Package" + then show "?P m" + by (auto simp add: permits_acc_def) + next + case (overriding declC C new newm old Sup) + assume member_new: "G \ new member_in C" and + new: "new = (declC, mdecl newm)" and + override: "G \ (declC, newm) overrides old" and + subcls_C_Sup: "G\C \\<^sub>C Sup" and + acc_old: "G \ methdMembr old in Sup dyn_accessible_from accC" and + hyp: "?Pack (methdMembr old) \ ?P (methdMembr old)" and + accmodi_new: "accmodi new = Package" + from override accmodi_new new wf + have accmodi_old: "accmodi old = Package" + by (auto dest: overrides_Package_old) + with hyp + have P_sup: "?P (methdMembr old)" + by (simp) + from wf override new accmodi_old accmodi_new + have eq_pid_new_old: "pid (declclass new) = pid (declclass old)" + by (auto dest: dyn_override_Package) + with eq_pid_new_old P_sup show "?P new" + by auto + qed +qed + +end \ No newline at end of file diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/WellType.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/WellType.thy Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,598 @@ +(* Title: isabelle/Bali/WellType.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1997 Technische Universitaet Muenchen +*) +header {* Well-typedness of Java programs *} + +theory WellType = DeclConcepts: + +text {* +improvements over Java Specification 1.0: +\begin{itemize} +\item methods of Object can be called upon references of interface or array type +\end{itemize} +simplifications: +\begin{itemize} +\item the type rules include all static checks on statements and expressions, + e.g. definedness of names (of parameters, locals, fields, methods) +\end{itemize} +design issues: +\begin{itemize} +\item unified type judgment for statements, variables, expressions, + expression lists +\item statements are typed like expressions with dummy type Void +\item the typing rules take an extra argument that is capable of determining + the dynamic type of objects. Therefore, they can be used for both + checking static types and determining runtime types in transition semantics. +\end{itemize} +*} + +types lenv + = "(lname, ty) table" (* local variables, including This and Result *) + +record env = + prg:: "prog" (* program *) + cls:: "qtname" (* current package and class name *) + lcl:: "lenv" (* local environment *) + +translations + "lenv" <= (type) "(lname, ty) table" + "lenv" <= (type) "lname \ ty option" + "env" <= (type) "\prg::prog,cls::qtname,lcl::lenv\" + "env" <= (type) "\prg::prog,cls::qtname,lcl::lenv,\::'a\" + + + +syntax + pkg :: "env \ pname" (* select the current package from an environment *) +translations + "pkg e" == "pid (cls e)" + +section "Static overloading: maximally specific methods " + +types + emhead = "ref_ty \ mhead" + +(* Some mnemotic selectors for emhead *) +constdefs + "declrefT" :: "emhead \ ref_ty" + "declrefT \ fst" + + "mhd" :: "emhead \ mhead" + "mhd \ snd" + +lemma declrefT_simp[simp]:"declrefT (r,m) = r" +by (simp add: declrefT_def) + +lemma mhd_simp[simp]:"mhd (r,m) = m" +by (simp add: mhd_def) + +lemma static_mhd_simp[simp]: "static (mhd m) = is_static m" +by (cases m) (simp add: member_is_static_simp mhd_def) + +lemma mhd_resTy_simp [simp]: "resTy (mhd m) = resTy m" +by (cases m) simp + +lemma mhd_is_static_simp [simp]: "is_static (mhd m) = is_static m" +by (cases m) simp + +lemma mhd_accmodi_simp [simp]: "accmodi (mhd m) = accmodi m" +by (cases m) simp + +consts + cmheads :: "prog \ qtname \ qtname \ sig \ emhead set" + Objectmheads :: "prog \ qtname \ sig \ emhead set" + accObjectmheads:: "prog \ qtname \ ref_ty \ sig \ emhead set" + mheads :: "prog \ qtname \ ref_ty \ sig \ emhead set" +defs + cmheads_def: +"cmheads G S C + \ \sig. (\(Cls,mthd). (ClassT Cls,(mhead mthd))) ` o2s (accmethd G S C sig)" + Objectmheads_def: +"Objectmheads G S + \ \sig. (\(Cls,mthd). (ClassT Cls,(mhead mthd))) + ` o2s (filter_tab (\sig m. accmodi m \ Private) (accmethd G S Object) sig)" + accObjectmheads_def: +"accObjectmheads G S T + \ if G\RefT T accessible_in (pid S) + then Objectmheads G S + else \sig. {}" +primrec +"mheads G S NullT = (\sig. {})" +"mheads G S (IfaceT I) = (\sig. (\(I,h).(IfaceT I,h)) + ` accimethds G (pid S) I sig \ + accObjectmheads G S (IfaceT I) sig)" +"mheads G S (ClassT C) = cmheads G S C" +"mheads G S (ArrayT T) = accObjectmheads G S (ArrayT T)" + + +(* more detailed than necessary for type-safety, see below. *) +constdefs + (* applicable methods, cf. 15.11.2.1 *) + appl_methds :: "prog \ qtname \ ref_ty \ sig \ (emhead \ ty list)\ set" + "appl_methds G S rt \ \ sig. + {(mh,pTs') |mh pTs'. mh \ mheads G S rt \name=name sig,parTs=pTs'\ \ + G\(parTs sig)[\]pTs'}" + + (* more specific methods, cf. 15.11.2.2 *) + more_spec :: "prog \ emhead \ ty list \ emhead \ ty list \ bool" + "more_spec G \ \(mh,pTs). \(mh',pTs'). G\pTs[\]pTs'" +(*more_spec G \\((d,h),pTs). \((d',h'),pTs'). G\RefT d\RefT d'\G\pTs[\]pTs'*) + + (* maximally specific methods, cf. 15.11.2.2 *) + max_spec :: "prog \ qtname \ ref_ty \ sig \ (emhead \ ty list)\ set" + + "max_spec G S rt sig \{m. m \appl_methds G S rt sig \ + (\m'\appl_methds G S rt sig. more_spec G m' m \ m'=m)}" +(* +rules (* all properties of more_spec, appl_methods and max_spec we actually need + these can easily be proven from the above definitions *) + +max_spec2mheads "max_spec G S rt (mn, pTs) = insert (mh, pTs') A \ + mh\mheads G S rt (mn, pTs') \ G\pTs[\]pTs'" +*) + +lemma max_spec2appl_meths: + "x \ max_spec G S T sig \ x \ appl_methds G S T sig" +by (auto simp: max_spec_def) + +lemma appl_methsD: "(mh,pTs')\appl_methds G S T \name=mn,parTs=pTs\ \ + mh \ mheads G S T \name=mn,parTs=pTs'\ \ G\pTs[\]pTs'" +by (auto simp: appl_methds_def) + +lemma max_spec2mheads: +"max_spec G S rt \name=mn,parTs=pTs\ = insert (mh, pTs') A + \ mh \ mheads G S rt \name=mn,parTs=pTs'\ \ G\pTs[\]pTs'" +apply (auto dest: equalityD2 subsetD max_spec2appl_meths appl_methsD) +done + + +constdefs + empty_dt :: "dyn_ty" + "empty_dt \ \a. None" + + invmode :: "('a::type)member_scheme \ expr \ inv_mode" +"invmode m e \ if static m then Static else if e=Super then SuperM else IntVir" + +lemma invmode_nonstatic [simp]: + "invmode \access=a,static=False,\=x\ (Acc (LVar e)) = IntVir" +apply (unfold invmode_def) +apply (simp (no_asm)) +done + + +lemma invmode_Static_eq [simp]: "(invmode m e = Static) = static m" +apply (unfold invmode_def) +apply (simp (no_asm)) +done + + +lemma invmode_IntVir_eq: "(invmode m e = IntVir) = (\(static m) \ e\Super)" +apply (unfold invmode_def) +apply (simp (no_asm)) +done + +lemma Null_staticD: + "a'=Null \ (static m) \ invmode m e = IntVir \ a' \ Null" +apply (clarsimp simp add: invmode_IntVir_eq) +done + + +types tys = "ty + ty list" +translations + "tys" <= (type) "ty + ty list" + +consts + wt :: "(env\ \ dyn_ty\ \ term \ tys) set" +(*wt :: " env \ dyn_ty \ (term \ tys) set" not feasible because of + \ changing env in Try stmt *) + +syntax +wt :: "env \ dyn_ty \ [term,tys] \ bool" ("_,_|=_::_" [51,51,51,51] 50) +wt_stmt :: "env \ dyn_ty \ stmt \ bool" ("_,_|=_:<>" [51,51,51 ] 50) +ty_expr :: "env \ dyn_ty \ [expr ,ty ] \ bool" ("_,_|=_:-_" [51,51,51,51] 50) +ty_var :: "env \ dyn_ty \ [var ,ty ] \ bool" ("_,_|=_:=_" [51,51,51,51] 50) +ty_exprs:: "env \ dyn_ty \ [expr list, + \ \ ty list] \ bool" ("_,_|=_:#_" [51,51,51,51] 50) + +syntax (xsymbols) +wt :: "env \ dyn_ty \ [term,tys] \ bool" ("_,_\_\_" [51,51,51,51] 50) +wt_stmt :: "env \ dyn_ty \ stmt \ bool" ("_,_\_\\" [51,51,51 ] 50) +ty_expr :: "env \ dyn_ty \ [expr ,ty ] \ bool" ("_,_\_\-_" [51,51,51,51] 50) +ty_var :: "env \ dyn_ty \ [var ,ty ] \ bool" ("_,_\_\=_" [51,51,51,51] 50) +ty_exprs:: "env \ dyn_ty \ [expr list, + \ \ ty list] \ bool"("_,_\_\\_" [51,51,51,51] 50) + +translations + "E,dt\t\T" == "(E,dt,t,T) \ wt" + "E,dt\s\\" == "E,dt\In1r s\Inl (PrimT Void)" + "E,dt\e\-T" == "E,dt\In1l e\Inl T" + "E,dt\e\=T" == "E,dt\In2 e\Inl T" + "E,dt\e\\T" == "E,dt\In3 e\Inr T" + +syntax (* for purely static typing *) + wt_ :: "env \ [term,tys] \ bool" ("_|-_::_" [51,51,51] 50) + wt_stmt_ :: "env \ stmt \ bool" ("_|-_:<>" [51,51 ] 50) + ty_expr_ :: "env \ [expr ,ty ] \ bool" ("_|-_:-_" [51,51,51] 50) + ty_var_ :: "env \ [var ,ty ] \ bool" ("_|-_:=_" [51,51,51] 50) + ty_exprs_:: "env \ [expr list, + \ ty list] \ bool" ("_|-_:#_" [51,51,51] 50) + +syntax (xsymbols) + wt_ :: "env \ [term,tys] \ bool" ("_\_\_" [51,51,51] 50) + wt_stmt_ :: "env \ stmt \ bool" ("_\_\\" [51,51 ] 50) + ty_expr_ :: "env \ [expr ,ty ] \ bool" ("_\_\-_" [51,51,51] 50) + ty_var_ :: "env \ [var ,ty ] \ bool" ("_\_\=_" [51,51,51] 50) + ty_exprs_ :: "env \ [expr list, + \ ty list] \ bool" ("_\_\\_" [51,51,51] 50) + +translations + "E\t\ T" == "E,empty_dt\t\ T" + "E\s\\" == "E,empty_dt\s\\" + "E\e\-T" == "E,empty_dt\e\-T" + "E\e\=T" == "E,empty_dt\e\=T" + "E\e\\T" == "E,empty_dt\e\\T" + + +inductive wt intros + + +(* well-typed statements *) + + Skip: "E,dt\Skip\\" + + Expr: "\E,dt\e\-T\ \ + E,dt\Expr e\\" + (* cf. 14.6 *) + Lab: "E,dt\c\\ \ + E,dt\l\ c\\" + + Comp: "\E,dt\c1\\; + E,dt\c2\\\ \ + E,dt\c1;; c2\\" + + (* cf. 14.8 *) + If: "\E,dt\e\-PrimT Boolean; + E,dt\c1\\; + E,dt\c2\\\ \ + E,dt\If(e) c1 Else c2\\" + + (* cf. 14.10 *) + Loop: "\E,dt\e\-PrimT Boolean; + E,dt\c\\\ \ + E,dt\l\ While(e) c\\" + (* cf. 14.13, 14.15, 14.16 *) + Do: "E,dt\Do jump\\" + + (* cf. 14.16 *) + Throw: "\E,dt\e\-Class tn; + prg E\tn\\<^sub>C SXcpt Throwable\ \ + E,dt\Throw e\\" + (* cf. 14.18 *) + Try: "\E,dt\c1\\; prg E\tn\\<^sub>C SXcpt Throwable; + lcl E (VName vn)=None; E \lcl := lcl E(VName vn\Class tn)\,dt\c2\\\ + \ + E,dt\Try c1 Catch(tn vn) c2\\" + + (* cf. 14.18 *) + Fin: "\E,dt\c1\\; E,dt\c2\\\ \ + E,dt\c1 Finally c2\\" + + Init: "\is_class (prg E) C\ \ + E,dt\Init C\\" + (* Init is created on the fly during evaluation (see Eval.thy). The class + * isn't necessarily accessible from the points Init is called. Therefor + * we only demand is_class and not is_acc_class here + *) + +(* well-typed expressions *) + + (* cf. 15.8 *) + NewC: "\is_acc_class (prg E) (pkg E) C\ \ + E,dt\NewC C\-Class C" + (* cf. 15.9 *) + NewA: "\is_acc_type (prg E) (pkg E) T; + E,dt\i\-PrimT Integer\ \ + E,dt\New T[i]\-T.[]" + + (* cf. 15.15 *) + Cast: "\E,dt\e\-T; is_acc_type (prg E) (pkg E) T'; + prg E\T\? T'\ \ + E,dt\Cast T' e\-T'" + + (* cf. 15.19.2 *) + Inst: "\E,dt\e\-RefT T; is_acc_type (prg E) (pkg E) (RefT T'); + prg E\RefT T\? RefT T'\ \ + E,dt\e InstOf T'\-PrimT Boolean" + + (* cf. 15.7.1 *) + Lit: "\typeof dt x = Some T\ \ + E,dt\Lit x\-T" + + (* cf. 15.10.2, 15.11.1 *) + Super: "\lcl E This = Some (Class C); C \ Object; + class (prg E) C = Some c\ \ + E,dt\Super\-Class (super c)" + + (* cf. 15.13.1, 15.10.1, 15.12 *) + Acc: "\E,dt\va\=T\ \ + E,dt\Acc va\-T" + + (* cf. 15.25, 15.25.1 *) + Ass: "\E,dt\va\=T; va \ LVar This; + E,dt\v \-T'; + prg E\T'\T\ \ + E,dt\va:=v\-T'" + + (* cf. 15.24 *) + Cond: "\E,dt\e0\-PrimT Boolean; + E,dt\e1\-T1; E,dt\e2\-T2; + prg E\T1\T2 \ T = T2 \ prg E\T2\T1 \ T = T1\ \ + E,dt\e0 ? e1 : e2\-T" + + (* cf. 15.11.1, 15.11.2, 15.11.3 *) + Call: "\E,dt\e\-RefT statT; + E,dt\ps\\pTs; + max_spec (prg E) (cls E) statT \name=mn,parTs=pTs\ + = {((statDeclT,m),pTs')} + \ \ + E,dt\{statT,invmode m e}e\mn({pTs'}ps)\-(resTy m)" + + Methd: "\is_class (prg E) C; + methd (prg E) C sig = Some m; + E,dt\Body (declclass m) (stmt (mbody (mthd m)))\-T\ \ + E,dt\Methd C sig\-T" + (* The class C is the dynamic class of the method call (cf. Eval.thy). + * It hasn't got to be directly accessible from the current package (pkg E). + * Only the static class must be accessible (enshured indirectly by Call). + *) + + Body: "\is_class (prg E) D; + E,dt\blk\\; + (lcl E) Result = Some T; + is_type (prg E) T\ \ + E,dt\Body D blk\-T" + (* the class D implementing the method must not directly be accessible + * from the current package (pkg E), but can also be indirectly accessible + * due to inheritance (enshured in Call) + * The result type hasn't got to be accessible in Java! (If it is not + * accessible you can only assign it to Object) + *) + +(* well-typed variables *) + + (* cf. 15.13.1 *) + LVar: "\lcl E vn = Some T; is_acc_type (prg E) (pkg E) T\ \ + E,dt\LVar vn\=T" + (* cf. 15.10.1 *) + FVar: "\E,dt\e\-Class C; + accfield (prg E) (cls E) C fn = Some (fd,f)\ \ + E,dt\{fd,static f}e..fn\=(type f)" + (* cf. 15.12 *) + AVar: "\E,dt\e\-T.[]; + E,dt\i\-PrimT Integer\ \ + E,dt\e.[i]\=T" + + +(* well-typed expression lists *) + + (* cf. 15.11.??? *) + Nil: "E,dt\[]\\[]" + + (* cf. 15.11.??? *) + Cons: "\E,dt\e \-T; + E,dt\es\\Ts\ \ + E,dt\e#es\\T#Ts" + + +declare not_None_eq [simp del] +declare split_if [split del] split_if_asm [split del] +declare split_paired_All [simp del] split_paired_Ex [simp del] +ML_setup {* +simpset_ref() := simpset() delloop "split_all_tac" +*} +inductive_cases wt_stmt_cases: "E,dt\c\\" +inductive_cases wt_elim_cases: + "E,dt\In2 (LVar vn) \T" + "E,dt\In2 ({fd,s}e..fn) \T" + "E,dt\In2 (e.[i]) \T" + "E,dt\In1l (NewC C) \T" + "E,dt\In1l (New T'[i]) \T" + "E,dt\In1l (Cast T' e) \T" + "E,dt\In1l (e InstOf T') \T" + "E,dt\In1l (Lit x) \T" + "E,dt\In1l (Super) \T" + "E,dt\In1l (Acc va) \T" + "E,dt\In1l (Ass va v) \T" + "E,dt\In1l (e0 ? e1 : e2) \T" + "E,dt\In1l ({statT,mode}e\mn({pT'}p))\T" + "E,dt\In1l (Methd C sig) \T" + "E,dt\In1l (Body D blk) \T" + "E,dt\In3 ([]) \Ts" + "E,dt\In3 (e#es) \Ts" + "E,dt\In1r Skip \x" + "E,dt\In1r (Expr e) \x" + "E,dt\In1r (c1;; c2) \x" + "E,dt\In1r (l\ c) \x" + "E,dt\In1r (If(e) c1 Else c2) \x" + "E,dt\In1r (l\ While(e) c) \x" + "E,dt\In1r (Do jump) \x" + "E,dt\In1r (Throw e) \x" + "E,dt\In1r (Try c1 Catch(tn vn) c2)\x" + "E,dt\In1r (c1 Finally c2) \x" + "E,dt\In1r (Init C) \x" +declare not_None_eq [simp] +declare split_if [split] split_if_asm [split] +declare split_paired_All [simp] split_paired_Ex [simp] +ML_setup {* +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac) +*} + +lemma is_acc_class_is_accessible: + "is_acc_class G P C \ G\(Class C) accessible_in P" +by (auto simp add: is_acc_class_def) + +lemma is_acc_iface_is_iface: "is_acc_iface G P I \ is_iface G I" +by (auto simp add: is_acc_iface_def) + +lemma is_acc_iface_is_accessible: + "is_acc_iface G P I \ G\(Iface I) accessible_in P" +by (auto simp add: is_acc_iface_def) + +lemma is_acc_type_is_type: "is_acc_type G P T \ is_type G T" +by (auto simp add: is_acc_type_def) + +lemma is_acc_iface_is_accessible: "is_acc_type G P T \ G\T accessible_in P" +by (auto simp add: is_acc_type_def) + +lemma wt_Methd_is_methd: + "E\In1l (Methd C sig)\T \ is_methd (prg E) C sig" +apply (erule_tac wt_elim_cases) +apply clarsimp +apply (erule is_methdI, assumption) +done + +(* Special versions of some typing rules, better suited to pattern match the + * conclusion (no selectors in the conclusion\) + *) + +lemma wt_Super: +"\lcl E This = Some (Class C); C \ Object; + class (prg E) C = Some c;D=super c\ + \ E,dt\Super\-Class D" +by (auto elim: wt.Super) + +lemma wt_Call: +"\E,dt\e\-RefT statT; E,dt\ps\\pTs; + max_spec (prg E) (cls E) statT \name=mn,parTs=pTs\ + = {((statDeclC,m),pTs')};rT=(resTy m); + mode = invmode m e\ \ E,dt\{statT,mode}e\mn({pTs'}ps)\-rT" +by (auto elim: wt.Call) + + + +lemma invocationTypeExpr_noClassD: +"\ E\e\-RefT statT\ + \ (\ statC. statT \ ClassT statC) \ invmode m e \ SuperM" +proof - + assume wt: "E\e\-RefT statT" + show ?thesis + proof (cases "e=Super") + case True + with wt obtain "C" where "statT = ClassT C" by (blast elim: wt_elim_cases) + then show ?thesis by blast + next + case False then show ?thesis + by (auto simp add: invmode_def split: split_if_asm) + qed +qed + +lemma wt_Super: +"\lcl E This = Some (Class C); C \ Object; class (prg E) C = Some c; D=super c\ +\ E,dt\Super\-Class D" +by (auto elim: wt.Super) + + +lemma wt_FVar: +"\E,dt\e\-Class C; accfield (prg E) (cls E) C fn = Some (fd,f); + sf=static f; fT=(type f)\ \ E,dt\{fd,sf}e..fn\=fT" +by (auto elim: wt.FVar) + +lemma wt_init [iff]: "E,dt\Init C\\ = is_class (prg E) C" +by (auto elim: wt_elim_cases intro: "wt.Init") + +declare wt.Skip [iff] + +lemma wt_StatRef: + "is_acc_type (prg E) (pkg E) (RefT rt) \ E\StatRef rt\-RefT rt" +apply (rule wt.Cast) +apply (rule wt.Lit) +apply (simp (no_asm)) +apply (simp (no_asm_simp)) +apply (rule cast.widen) +apply (simp (no_asm)) +done + +lemma wt_Inj_elim: + "\E. E,dt\t\U \ case t of + In1 ec \ (case ec of + Inl e \ \T. U=Inl T + | Inr s \ U=Inl (PrimT Void)) + | In2 e \ (\T. U=Inl T) + | In3 e \ (\T. U=Inr T)" +apply (erule wt.induct) +apply auto +done + +ML {* +fun wt_fun name inj rhs = +let + val lhs = "E,dt\" ^ inj ^ " t\U" + val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") + (K [Auto_tac, ALLGOALS (ftac (thm "wt_Inj_elim")) THEN Auto_tac]) + fun is_Inj (Const (inj,_) $ _) = true + | is_Inj _ = false + fun pred (t as (_ $ (Const ("Pair",_) $ + _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $ + x))) $ _ )) = is_Inj x +in + make_simproc name lhs pred (thm name) +end + +val wt_expr_proc = wt_fun "wt_expr_eq" "In1l" "\T. U=Inl T \ E,dt\t\-T" +val wt_var_proc = wt_fun "wt_var_eq" "In2" "\T. U=Inl T \ E,dt\t\=T" +val wt_exprs_proc = wt_fun "wt_exprs_eq" "In3" "\Ts. U=Inr Ts \ E,dt\t\\Ts" +val wt_stmt_proc = wt_fun "wt_stmt_eq" "In1r" "U=Inl(PrimT Void)\E,dt\t\\" +*} + +ML {* +Addsimprocs [wt_expr_proc,wt_var_proc,wt_exprs_proc,wt_stmt_proc] +*} + +lemma Inj_eq_lemma [simp]: + "(\T. (\T'. T = Inj T' \ P T') \ Q T) = (\T'. P T' \ Q (Inj T'))" +by auto + +(* unused *) +lemma single_valued_tys_lemma [rule_format (no_asm)]: + "\S T. G\S\T \ G\T\S \ S = T \ E,dt\t\T \ + G = prg E \ (\T'. E,dt\t\T' \ T = T')" +apply (cases "E", erule wt.induct) +apply (safe del: disjE) +apply (simp_all (no_asm_use) split del: split_if_asm) +apply (safe del: disjE) +(* 17 subgoals *) +apply (tactic {* ALLGOALS (fn i => if i = 9 then EVERY'[thin_tac "?E,dt\e0\-PrimT Boolean", thin_tac "?E,dt\e1\-?T1", thin_tac "?E,dt\e2\-?T2"] i else thin_tac "All ?P" i) *}) +(*apply (safe del: disjE elim!: wt_elim_cases)*) +apply (tactic {*ALLGOALS (eresolve_tac (thms "wt_elim_cases"))*}) +apply (simp_all (no_asm_use) split del: split_if_asm) +apply (erule_tac [10] V = "All ?P" in thin_rl) (* Call *) +apply ((blast del: equalityCE dest: sym [THEN trans])+) +done + +(* unused *) +lemma single_valued_tys: +"ws_prog (prg E) \ single_valued {(t,T). E,dt\t\T}" +apply (unfold single_valued_def) +apply clarsimp +apply (rule single_valued_tys_lemma) +apply (auto intro!: widen_antisym) +done + +lemma typeof_empty_is_type [rule_format (no_asm)]: + "typeof (\a. None) v = Some T \ is_type G T" +apply (rule val.induct) +apply auto +done + +(* unused *) +lemma typeof_is_type [rule_format (no_asm)]: + "(\a. v \ Addr a) \ (\T. typeof dt v = Some T \ is_type G T)" +apply (rule val.induct) +prefer 5 +apply fast +apply (simp_all (no_asm)) +done + +end diff -r de505273c971 -r 00d4a435777f src/HOL/Bali/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/document/root.tex Mon Jan 28 17:00:19 2002 +0100 @@ -0,0 +1,238 @@ + +\documentclass[11pt,a4paper]{book} +\usepackage{isabelle,isabellesym} + +% further packages required for unusual symbols (see also isabellesym.sty) +\usepackage{latexsym} +%\usepackage{amssymb} +%\usepackage[english]{babel} +%\usepackage[latin1]{inputenc} +%\usepackage[only,bigsqcap]{stmaryrd} +%\usepackage{wasysym} +%\usepackage{eufrak} +%\usepackage{textcomp} +%\usepackage{marvosym} + +% this should be the last package used +\usepackage{pdfsetup} + +% proper setup for best-style documents +\urlstyle{rm} +\isabellestyle{it} + +\pagestyle{myheadings} + +\addtolength{\hoffset}{-1,5cm} +\addtolength{\textwidth}{4cm} +\addtolength{\voffset}{-2cm} +\addtolength{\textheight}{4cm} + +%subsection instead of section to make the toc readable +\renewcommand{\thesubsection}{\arabic{subsection}} +\renewcommand{\isamarkupheader}[1]{\newpage\markright{Theory~\isabellecontext}\subsection{#1}} +\renewcommand{\isamarkupsection}[1]{\subsubsection{#1}} + +%remove spaces from the isabelle environment (trivlist makes them too large) +\renewenvironment{isabelle} +{\begin{isabellebody}} +{\end{isabellebody}} + +\begin{document} + +\title{Java Source and Bytecode Formalizations in Isabelle: Bali\\ + {\large -- VerifiCard Project Deliverables -- }} +\author{Gerwin Klein \and Tobias Nipkow \and David von Oheimb \and + Leonor Prensa Nieto \and Norbert Schirmer \and Martin Strecker} +\maketitle + +\tableofcontents + +\parindent 0pt\parskip 0.5ex +\chapter{Overview} +These theories, called Bali, model and analyse different aspects of the +JavaCard \textbf{source language}. +The basis is an abstract model of the JavaCard source language. +On it, a type system, an operational semantics and an axiomatic semantics +(Hoare logic) are built. The execution of a wellformed program (with respect to +the type system) according to the operational semantics is proved to be +typesafe. The axiomatic semantics is proved to be sound and relative complete +with respect to the operational semantics. + +We have modelled large parts of the original JavaCard source language. It models +features such as: +\begin{itemize} +\item The basic ``primitive types'' of Java +\item Classes and related concepts +\item Class fields and methods +\item Instance fields and methods +\item Interfaces and related concepts +\item Arrays +\item Static initialisation +\item Static overloading of fields and methods +\item Inheritance, overriding and hiding of methods, dynamic binding +\item All cases of abrupt termination + \begin{itemize} + \item Exception throwing and handling + \item \texttt{break}, \texttt{continue} and \texttt{return} + \end{itemize} +\item Packages +\item Access Modifiers (\texttt{private}, \texttt{protected}, \texttt{public}) +\end{itemize} + +The following features are missing in Bali wrt.{} JavaCard: +\begin{itemize} +\item Some primitive types (\texttt{byte, short}) +\item Most numeric operations, syntactic variants of statements + (\texttt{do}-loop, \texttt{for}-loop) +\item A ``definite assignment'' check +\end{itemize} + +In addition, features are missing that are not part of the JavaCard +language, such as multithreading and garbage collection. No attempt +has been made to model peculiarities of JavaCard such as the applet +firewall or the transaction mechanism. + + +Overview of the theories: +\begin{description} +\item [Basis.thy] +Some basic definitions and settings not specific to JavaCard but missing in HOL. + +\item [Table.thy] +Definition and some properties of a lookup table to map various names +(like class names or method names) to some content (like classes or methods). + +\item[Name.thy] +Definition of various names (class names, variable names, package names,...) + +\item[Value.thy] +JavaCard expression values (Boolean, Integer, Addresses,...) + +\item[Type.thy] +JavaCard types. Primitive types (Boolean, Integer,...) and reference types +(Classes, Interfaces, Arrays,...) + +\item[Term.thy] +JavaCard terms. Variables, expressions and statements. + +\item[Decl.thy] +Class, interface and program declarations. Recursion operators for the +class and the interface hierarchy. + +\item[TypeRel.thy] +Various relations on types like the subclass-, subinterface-, widening-, +narrowing- and casting-relation. + +\item[DeclConcepts.thy] +Advanced concepts on the class and interface hierarchy like inheritance, +overriding, hiding, accessibility of types and members according to the access +modifiers, method lookup. + +\item[WellType.thy] +Typesystem on the JavaCard term level. + +\item[WellForm.thy] +Typesystem on the JavaCard class, interface and program level. + +\item[State.thy] +The program state (like object store) for the execution of JavaCard. +Abrupt completion (exceptions, break, continue, return) is modelled as flag +inside the state. + +\item[Eval.thy] +Operational (big step) semantics for JavaCard. + +\item[Example.thy] +An concrete example of a JavaCard program to validate the typesystem and the +operational semantics. + +\item[Conform.thy] +Conformance predicate for states. When does an execution state conform to the +static types of the program given by the typesystem. + +\item[TypeSafe.thy] +Typesafety proof of the execution of JavaCard. ''Welltyped programs don't go +wrong'' or more technical: The execution of a welltyped JavaCard program +preserves the conformance of execution states. + +\item[Evaln.thy] +Copy of the operational semantics given in Eval.thy expanded with an annotation +for the maximal recursive depth. The semantics is not altered. The annotation +is needed for the soundness proof of the axiomatic semantics. + +\item[AxSem.thy] +An axiomatic semantics (Hoare logic) for JavaCard. + +\item[AxSound.thy] +The soundness proof of the axiomatic semantics with respect to the operational +semantics. + +\item[AxCompl.thy] +The proof of (relative) completeness of the axiomatic semantics with respect +to the operational semantics. + +\item[AxExample.thy] +An concrete example of the axiomatic semantics at work, applied to prove +some properties of the JavaCard example given in Example.thy. +\end{description} + +% include generated text of all theories +\chapter{Basis.thy} +\input{../generated/Basis.tex} +\chapter{Table.thy} +\input{../generated/Table.tex} + +\chapter{Name.thy} +\input{../generated/Name.tex} +\chapter{Value.thy} +\input{../generated/Value.tex} +\chapter{Type.thy} +\input{../generated/Type.tex} +\chapter{Term.thy} +\input{../generated/Term.tex} +\chapter{Decl.thy} +\input{../generated/Decl.tex} +\chapter{TypeRel.thy} +\input{../generated/TypeRel.tex} +\chapter{DeclConcepts.thy} +\input{../generated/DeclConcepts.tex} + +\chapter{WellType.thy} +\input{../generated/WellType.tex} +\chapter{WellForm.thy} +\input{../generated/WellForm.tex} + +\chapter{State.thy} +\input{../generated/State.tex} +\chapter{Eval.thy} +\input{../generated/Eval.tex} + +\chapter{Example.thy} +\input{../generated/Example.tex} + +\chapter{Conform.thy} +\input{../generated/Conform.tex} +\chapter{TypeSafe.thy} +\input{../generated/TypeSafe.tex} + +\chapter{Evaln.thy} +\input{../generated/Evaln.tex} +\chapter{AxSem.thy} +\input{../generated/AxSem.tex} +\chapter{AxSound.thy} +\input{../generated/AxSound.tex} +\chapter{AxCompl.thy} +\input{../generated/AxCompl.tex} +\chapter{AxExample.thy} +\input{../generated/AxExample.tex} + + + + + + + +%\bibliographystyle{abbrv} +%\bibliography{root} + +\end{document}