Isabelle/Bali sources;
authorschirmer
Mon, 28 Jan 2002 17:00:19 +0100
changeset 1285400d4a435777f
parent 12853 de505273c971
child 12855 21225338f8db
Isabelle/Bali sources;
src/HOL/Bali/AxCompl.thy
src/HOL/Bali/AxExample.thy
src/HOL/Bali/AxSem.thy
src/HOL/Bali/AxSound.thy
src/HOL/Bali/Basis.thy
src/HOL/Bali/Conform.thy
src/HOL/Bali/Decl.thy
src/HOL/Bali/DeclConcepts.thy
src/HOL/Bali/Eval.thy
src/HOL/Bali/Evaln.thy
src/HOL/Bali/Example.thy
src/HOL/Bali/Name.thy
src/HOL/Bali/ROOT.ML
src/HOL/Bali/State.thy
src/HOL/Bali/Table.thy
src/HOL/Bali/Term.thy
src/HOL/Bali/Trans.thy
src/HOL/Bali/Type.thy
src/HOL/Bali/TypeRel.thy
src/HOL/Bali/TypeSafe.thy
src/HOL/Bali/Value.thy
src/HOL/Bali/WellForm.thy
src/HOL/Bali/WellType.thy
src/HOL/Bali/document/root.tex
     1.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.2 +++ b/src/HOL/Bali/AxCompl.thy	Mon Jan 28 17:00:19 2002 +0100
     1.3 @@ -0,0 +1,645 @@
     1.4 +(*  Title:      isabelle/Bali/AxCompl.thy
     1.5 +    ID:         $Id$
     1.6 +    Author:     David von Oheimb
     1.7 +    Copyright   1999 Technische Universitaet Muenchen
     1.8 +*)
     1.9 +
    1.10 +header {*
    1.11 +Completeness proof for Axiomatic semantics of Java expressions and statements
    1.12 +*}
    1.13 +
    1.14 +theory AxCompl = AxSem:
    1.15 +
    1.16 +text {*
    1.17 +design issues:
    1.18 +\begin{itemize}
    1.19 +\item proof structured by Most General Formulas (-> Thomas Kleymann)
    1.20 +\end{itemize}
    1.21 +*}
    1.22 +section "set of not yet initialzed classes"
    1.23 +
    1.24 +constdefs
    1.25 +
    1.26 +  nyinitcls :: "prog \<Rightarrow> state \<Rightarrow> qtname set"
    1.27 + "nyinitcls G s \<equiv> {C. is_class G C \<and> \<not> initd C s}"
    1.28 +
    1.29 +lemma nyinitcls_subset_class: "nyinitcls G s \<subseteq> {C. is_class G C}"
    1.30 +apply (unfold nyinitcls_def)
    1.31 +apply fast
    1.32 +done
    1.33 +
    1.34 +lemmas finite_nyinitcls [simp] =
    1.35 +   finite_is_class [THEN nyinitcls_subset_class [THEN finite_subset], standard]
    1.36 +
    1.37 +lemma card_nyinitcls_bound: "card (nyinitcls G s) \<le> card {C. is_class G C}"
    1.38 +apply (rule nyinitcls_subset_class [THEN finite_is_class [THEN card_mono]])
    1.39 +done
    1.40 +
    1.41 +lemma nyinitcls_set_locals_cong [simp]: 
    1.42 +  "nyinitcls G (x,set_locals l s) = nyinitcls G (x,s)"
    1.43 +apply (unfold nyinitcls_def)
    1.44 +apply (simp (no_asm))
    1.45 +done
    1.46 +
    1.47 +lemma nyinitcls_abrupt_cong [simp]: "nyinitcls G (f x, y) = nyinitcls G (x, y)"
    1.48 +apply (unfold nyinitcls_def)
    1.49 +apply (simp (no_asm))
    1.50 +done
    1.51 +
    1.52 +lemma nyinitcls_abupd_cong [simp]:"!!s. nyinitcls G (abupd f s) = nyinitcls G s"
    1.53 +apply (unfold nyinitcls_def)
    1.54 +apply (simp (no_asm_simp) only: split_tupled_all)
    1.55 +apply (simp (no_asm))
    1.56 +done
    1.57 +
    1.58 +lemma card_nyinitcls_abrupt_congE [elim!]: 
    1.59 +        "card (nyinitcls G (x, s)) \<le> n \<Longrightarrow> card (nyinitcls G (y, s)) \<le> n"
    1.60 +apply (unfold nyinitcls_def)
    1.61 +apply auto
    1.62 +done
    1.63 +
    1.64 +lemma nyinitcls_new_xcpt_var [simp]: 
    1.65 +"nyinitcls G (new_xcpt_var vn s) = nyinitcls G s"
    1.66 +apply (unfold nyinitcls_def)
    1.67 +apply (induct_tac "s")
    1.68 +apply (simp (no_asm))
    1.69 +done
    1.70 +
    1.71 +lemma nyinitcls_init_lvars [simp]: 
    1.72 +  "nyinitcls G ((init_lvars G C sig mode a' pvs) s) = nyinitcls G s"
    1.73 +apply (induct_tac "s")
    1.74 +apply (simp (no_asm) add: init_lvars_def2 split add: split_if)
    1.75 +done
    1.76 +
    1.77 +lemma nyinitcls_emptyD: "\<lbrakk>nyinitcls G s = {}; is_class G C\<rbrakk> \<Longrightarrow> initd C s"
    1.78 +apply (unfold nyinitcls_def)
    1.79 +apply fast
    1.80 +done
    1.81 +
    1.82 +lemma card_Suc_lemma: "\<lbrakk>card (insert a A) \<le> Suc n; a\<notin>A; finite A\<rbrakk> \<Longrightarrow> card A \<le> n"
    1.83 +apply (rotate_tac 1)
    1.84 +apply clarsimp
    1.85 +done
    1.86 +
    1.87 +lemma nyinitcls_le_SucD: 
    1.88 +"\<lbrakk>card (nyinitcls G (x,s)) \<le> Suc n; \<not>inited C (globs s); class G C=Some y\<rbrakk> \<Longrightarrow> 
    1.89 +  card (nyinitcls G (x,init_class_obj G C s)) \<le> n"
    1.90 +apply (subgoal_tac 
    1.91 +        "nyinitcls G (x,s) = insert C (nyinitcls G (x,init_class_obj G C s))")
    1.92 +apply  clarsimp
    1.93 +apply  (erule thin_rl)
    1.94 +apply  (rule card_Suc_lemma [OF _ _ finite_nyinitcls])
    1.95 +apply   (auto dest!: not_initedD elim!: 
    1.96 +              simp add: nyinitcls_def inited_def split add: split_if_asm)
    1.97 +done
    1.98 +
    1.99 +ML {* bind_thm("inited_gext'",permute_prems 0 1 (thm "inited_gext")) *}
   1.100 +
   1.101 +lemma nyinitcls_gext: "snd s\<le>|snd s' \<Longrightarrow> nyinitcls G s' \<subseteq> nyinitcls G s"
   1.102 +apply (unfold nyinitcls_def)
   1.103 +apply (force dest!: inited_gext')
   1.104 +done
   1.105 +
   1.106 +lemma card_nyinitcls_gext: 
   1.107 +  "\<lbrakk>snd s\<le>|snd s'; card (nyinitcls G s) \<le> n\<rbrakk>\<Longrightarrow> card (nyinitcls G s') \<le> n"
   1.108 +apply (rule le_trans)
   1.109 +apply  (rule card_mono)
   1.110 +apply   (rule finite_nyinitcls)
   1.111 +apply  (erule nyinitcls_gext)
   1.112 +apply assumption
   1.113 +done
   1.114 +
   1.115 +
   1.116 +section "init-le"
   1.117 +
   1.118 +constdefs
   1.119 +  init_le :: "prog \<Rightarrow> nat \<Rightarrow> state \<Rightarrow> bool"            ("_\<turnstile>init\<le>_"  [51,51] 50)
   1.120 + "G\<turnstile>init\<le>n \<equiv> \<lambda>s. card (nyinitcls G s) \<le> n"
   1.121 +  
   1.122 +lemma init_le_def2 [simp]: "(G\<turnstile>init\<le>n) s = (card (nyinitcls G s)\<le>n)"
   1.123 +apply (unfold init_le_def)
   1.124 +apply auto
   1.125 +done
   1.126 +
   1.127 +lemma All_init_leD: "\<forall>n::nat. G,A\<turnstile>{P \<and>. G\<turnstile>init\<le>n} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
   1.128 +apply (drule spec)
   1.129 +apply (erule conseq1)
   1.130 +apply clarsimp
   1.131 +apply (rule card_nyinitcls_bound)
   1.132 +done
   1.133 +
   1.134 +section "Most General Triples and Formulas"
   1.135 +
   1.136 +constdefs
   1.137 +
   1.138 +  remember_init_state :: "state assn"                ("\<doteq>")
   1.139 +  "\<doteq> \<equiv> \<lambda>Y s Z. s = Z"
   1.140 +
   1.141 +lemma remember_init_state_def2 [simp]: "\<doteq> Y = op ="
   1.142 +apply (unfold remember_init_state_def)
   1.143 +apply (simp (no_asm))
   1.144 +done
   1.145 +
   1.146 +consts
   1.147 +  
   1.148 +  MGF ::"[state assn, term, prog] \<Rightarrow> state triple"   ("{_} _\<succ> {_\<rightarrow>}"[3,65,3]62)
   1.149 +  MGFn::"[nat       , term, prog] \<Rightarrow> state triple" ("{=:_} _\<succ> {_\<rightarrow>}"[3,65,3]62)
   1.150 +
   1.151 +defs
   1.152 +  
   1.153 +
   1.154 +  MGF_def:
   1.155 +  "{P} t\<succ> {G\<rightarrow>} \<equiv> {P} t\<succ> {\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y,s')}"
   1.156 +
   1.157 +  MGFn_def:
   1.158 +  "{=:n} t\<succ> {G\<rightarrow>} \<equiv> {\<doteq> \<and>. G\<turnstile>init\<le>n} t\<succ> {G\<rightarrow>}"
   1.159 +
   1.160 +(* unused *)
   1.161 +lemma MGF_valid: "G,{}\<Turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
   1.162 +apply (unfold MGF_def)
   1.163 +apply (force dest!: evaln_eval simp add: ax_valids_def triple_valid_def2)
   1.164 +done
   1.165 +
   1.166 +lemma MGF_res_eq_lemma [simp]: 
   1.167 +  "(\<forall>Y' Y s. Y = Y' \<and> P s \<longrightarrow> Q s) = (\<forall>s. P s \<longrightarrow> Q s)"
   1.168 +apply auto
   1.169 +done
   1.170 +
   1.171 +lemma MGFn_def2: 
   1.172 +"G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>} = G,A\<turnstile>{\<doteq> \<and>. G\<turnstile>init\<le>n} 
   1.173 +                    t\<succ> {\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y,s')}"
   1.174 +apply (unfold MGFn_def MGF_def)
   1.175 +apply fast
   1.176 +done
   1.177 +
   1.178 +lemma MGF_MGFn_iff: "G,A\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} = (\<forall>n. G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>})"
   1.179 +apply (simp (no_asm_use) add: MGFn_def2 MGF_def)
   1.180 +apply safe
   1.181 +apply  (erule_tac [2] All_init_leD)
   1.182 +apply (erule conseq1)
   1.183 +apply clarsimp
   1.184 +done
   1.185 +
   1.186 +lemma MGFnD: 
   1.187 +"G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>} \<Longrightarrow>  
   1.188 + G,A\<turnstile>{(\<lambda>Y' s' s. s' = s           \<and> P s) \<and>. G\<turnstile>init\<le>n}  
   1.189 + t\<succ>  {(\<lambda>Y' s' s. G\<turnstile>s\<midarrow>t\<succ>\<rightarrow>(Y',s') \<and> P s) \<and>. G\<turnstile>init\<le>n}"
   1.190 +apply (unfold init_le_def)
   1.191 +apply (simp (no_asm_use) add: MGFn_def2)
   1.192 +apply (erule conseq12)
   1.193 +apply clarsimp
   1.194 +apply (erule (1) eval_gext [THEN card_nyinitcls_gext])
   1.195 +done
   1.196 +lemmas MGFnD' = MGFnD [of _ _ _ _ "\<lambda>x. True"] 
   1.197 +
   1.198 +lemma MGFNormalI: "G,A\<turnstile>{Normal \<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow>  
   1.199 +  G,(A::state triple set)\<turnstile>{\<doteq>::state assn} t\<succ> {G\<rightarrow>}"
   1.200 +apply (unfold MGF_def)
   1.201 +apply (rule ax_Normal_cases)
   1.202 +apply  (erule conseq1)
   1.203 +apply  clarsimp
   1.204 +apply (rule ax_derivs.Abrupt [THEN conseq1])
   1.205 +apply (clarsimp simp add: Let_def)
   1.206 +done
   1.207 +
   1.208 +lemma MGFNormalD: "G,A\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow> G,A\<turnstile>{Normal \<doteq>} t\<succ> {G\<rightarrow>}"
   1.209 +apply (unfold MGF_def)
   1.210 +apply (erule conseq1)
   1.211 +apply clarsimp
   1.212 +done
   1.213 +
   1.214 +lemma MGFn_NormalI: 
   1.215 +"G,(A::state triple set)\<turnstile>{Normal((\<lambda>Y' s' s. s'=s \<and> normal s) \<and>. G\<turnstile>init\<le>n)}t\<succ> 
   1.216 + {\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y,s')} \<Longrightarrow> G,A\<turnstile>{=:n}t\<succ>{G\<rightarrow>}"
   1.217 +apply (simp (no_asm_use) add: MGFn_def2)
   1.218 +apply (rule ax_Normal_cases)
   1.219 +apply  (erule conseq1)
   1.220 +apply  clarsimp
   1.221 +apply (rule ax_derivs.Abrupt [THEN conseq1])
   1.222 +apply (clarsimp simp add: Let_def)
   1.223 +done
   1.224 +
   1.225 +lemma MGFn_free_wt: 
   1.226 +  "(\<exists>T L C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) 
   1.227 +    \<longrightarrow> G,(A::state triple set)\<turnstile>{=:n} t\<succ> {G\<rightarrow>} 
   1.228 +   \<Longrightarrow> G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>}"
   1.229 +apply (rule MGFn_NormalI)
   1.230 +apply (rule ax_free_wt)
   1.231 +apply (auto elim: conseq12 simp add: MGFn_def MGF_def)
   1.232 +done
   1.233 +
   1.234 +
   1.235 +section "main lemmas"
   1.236 +
   1.237 +declare fun_upd_apply [simp del]
   1.238 +declare splitI2 [rule del] (*prevents ugly renaming of state variables*)
   1.239 +
   1.240 +ML_setup {* 
   1.241 +Delsimprocs [eval_expr_proc, eval_var_proc, eval_exprs_proc, eval_stmt_proc]
   1.242 +*} (*prevents modifying rhs of MGF*)
   1.243 +ML {*
   1.244 +val eval_css = (claset() delrules [thm "eval.Abrupt"] addSIs (thms "eval.intros") 
   1.245 +                delrules[thm "eval.Expr", thm "eval.Init", thm "eval.Try"] 
   1.246 +                addIs   [thm "eval.Expr", thm "eval.Init"]
   1.247 +                addSEs[thm "eval.Try"] delrules[equalityCE],
   1.248 +                simpset() addsimps [split_paired_all,Let_def]
   1.249 + addsimprocs [eval_expr_proc,eval_var_proc,eval_exprs_proc,eval_stmt_proc]);
   1.250 +val eval_Force_tac = force_tac eval_css;
   1.251 +
   1.252 +val wt_prepare_tac = EVERY'[
   1.253 +    rtac (thm "MGFn_free_wt"),
   1.254 +    clarsimp_tac (claset() addSEs (thms "wt_elim_cases"), simpset())]
   1.255 +val compl_prepare_tac = EVERY'[rtac (thm "MGFn_NormalI"), Simp_tac]
   1.256 +val forw_hyp_tac = EVERY'[etac (thm "MGFnD'" RS thm "conseq12"), Clarsimp_tac]
   1.257 +val forw_hyp_eval_Force_tac = 
   1.258 +         EVERY'[TRY o rtac allI, forw_hyp_tac, eval_Force_tac]
   1.259 +*}
   1.260 +
   1.261 +lemma MGFn_Init: "\<forall>m. Suc m\<le>n \<longrightarrow> (\<forall>t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>}) \<Longrightarrow>  
   1.262 +  G,(A::state triple set)\<turnstile>{=:n} In1r (Init C)\<succ> {G\<rightarrow>}"
   1.263 +apply (tactic "wt_prepare_tac 1")
   1.264 +(* requires is_class G C two times for nyinitcls *)
   1.265 +apply (tactic "compl_prepare_tac 1")
   1.266 +apply (rule_tac C = "initd C" in ax_cases)
   1.267 +apply  (rule ax_derivs.Done [THEN conseq1])
   1.268 +apply  (clarsimp intro!: init_done)
   1.269 +apply (rule_tac y = n in nat.exhaust, clarsimp)
   1.270 +apply  (rule ax_impossible [THEN conseq1])
   1.271 +apply  (force dest!: nyinitcls_emptyD)
   1.272 +apply clarsimp
   1.273 +apply (drule_tac x = "nat" in spec)
   1.274 +apply clarsimp
   1.275 +apply (rule_tac Q = " (\<lambda>Y s' (x,s) . G\<turnstile> (x,init_class_obj G C s) \<midarrow> (if C=Object then Skip else Init (super (the (class G C))))\<rightarrow> s' \<and> x=None \<and> \<not>inited C (globs s)) \<and>. G\<turnstile>init\<le>nat" in ax_derivs.Init)
   1.276 +apply   simp
   1.277 +apply  (rule_tac P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s \<and> normal s \<and> \<not> initd C s) \<and>. G\<turnstile>init\<le>nat) " in conseq1)
   1.278 +prefer 2
   1.279 +apply   (force elim!: nyinitcls_le_SucD)
   1.280 +apply  (simp split add: split_if, rule conjI, clarify)
   1.281 +apply   (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1")
   1.282 +apply  clarify
   1.283 +apply  (drule spec)
   1.284 +apply  (erule MGFnD' [THEN conseq12])
   1.285 +apply  (tactic "force_tac (claset(), simpset() addsimprocs[eval_stmt_proc]) 1")
   1.286 +apply (rule allI)
   1.287 +apply (drule spec)
   1.288 +apply (erule MGFnD' [THEN conseq12])
   1.289 +apply clarsimp
   1.290 +apply (tactic {* pair_tac "sa" 1 *})
   1.291 +apply (tactic"clarsimp_tac (claset(), simpset() addsimprocs[eval_stmt_proc]) 1")
   1.292 +apply (rule eval_Init, force+)
   1.293 +done
   1.294 +lemmas MGFn_InitD = MGFn_Init [THEN MGFnD, THEN ax_NormalD]
   1.295 +
   1.296 +lemma MGFn_Call: 
   1.297 +"\<lbrakk>\<forall>C sig. G,(A::state triple set)\<turnstile>{=:n} In1l (Methd C sig)\<succ> {G\<rightarrow>};  
   1.298 +  G,A\<turnstile>{=:n} In1l e\<succ> {G\<rightarrow>}; G,A\<turnstile>{=:n} In3 ps\<succ> {G\<rightarrow>}\<rbrakk> \<Longrightarrow>  
   1.299 +  G,A\<turnstile>{=:n} In1l ({statT,mode}e\<cdot>mn({pTs'}ps))\<succ> {G\<rightarrow>}"
   1.300 +apply (tactic "wt_prepare_tac 1") (* required for equating mode = invmode m e *)
   1.301 +apply (tactic "compl_prepare_tac 1")
   1.302 +apply (rule_tac R = "\<lambda>a'. (\<lambda>Y (x2,s2) (x,s) . x = None \<and> (\<exists>s1 pvs. G\<turnstile>Norm s \<midarrow>e-\<succ>a'\<rightarrow> s1 \<and> Y = In3 pvs \<and> G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>pvs\<rightarrow> (x2,s2))) \<and>. G\<turnstile>init\<le>n" in ax_derivs.Call)
   1.303 +apply  (erule MGFnD [THEN ax_NormalD])
   1.304 +apply safe
   1.305 +apply  (erule_tac V = "All ?P" in thin_rl, tactic "forw_hyp_eval_Force_tac 1")
   1.306 +apply (drule spec, drule spec)
   1.307 +apply (erule MGFnD' [THEN conseq12])
   1.308 +apply (tactic "clarsimp_tac eval_css 1")
   1.309 +apply (erule (1) eval_Call)
   1.310 +apply   (rule HOL.refl)
   1.311 +apply  (simp (no_asm_simp))+
   1.312 +done
   1.313 +
   1.314 +lemma MGF_altern: "G,A\<turnstile>{Normal (\<doteq> \<and>. p)} t\<succ> {G\<rightarrow>} =  
   1.315 + G,A\<turnstile>{Normal ((\<lambda>Y s Z. \<forall>w s'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<longrightarrow> (w,s') = Z) \<and>. p)} 
   1.316 +  t\<succ> {\<lambda>Y s Z. (Y,s) = Z}"
   1.317 +apply (unfold MGF_def)
   1.318 +apply (auto del: conjI elim!: conseq12)
   1.319 +apply (case_tac "\<exists>w s. G\<turnstile>Norm sa \<midarrow>t\<succ>\<rightarrow> (w,s) ")
   1.320 +apply  (fast dest: unique_eval)
   1.321 +apply clarsimp
   1.322 +apply (erule thin_rl)
   1.323 +apply (erule thin_rl)
   1.324 +apply (drule split_paired_All [THEN subst])
   1.325 +apply (clarsimp elim!: state_not_single)
   1.326 +done
   1.327 +
   1.328 +
   1.329 +lemma MGFn_Loop: 
   1.330 +"\<lbrakk>G,(A::state triple set)\<turnstile>{=:n} In1l expr\<succ> {G\<rightarrow>};G,A\<turnstile>{=:n} In1r stmnt\<succ> {G\<rightarrow>} \<rbrakk> 
   1.331 +\<Longrightarrow> 
   1.332 +  G,A\<turnstile>{=:n} In1r (l\<bullet> While(expr) stmnt)\<succ> {G\<rightarrow>}"
   1.333 +apply (rule MGFn_NormalI, simp)
   1.334 +apply (rule_tac p2 = "\<lambda>s. card (nyinitcls G s) \<le> n" in 
   1.335 +          MGF_altern [unfolded MGF_def, THEN iffD2, THEN conseq1])
   1.336 +prefer 2
   1.337 +apply  clarsimp
   1.338 +apply (rule_tac P' = 
   1.339 +"((\<lambda>Y s Z. \<forall>w s'. G\<turnstile>s \<midarrow>In1r (l\<bullet>  While(expr) stmnt) \<succ>\<rightarrow> (w,s') \<longrightarrow> (w,s') = Z) 
   1.340 +  \<and>. (\<lambda>s. card (nyinitcls G s) \<le> n))" in conseq12)
   1.341 +prefer 2
   1.342 +apply  clarsimp
   1.343 +apply  (tactic "smp_tac 1 1", erule_tac V = "All ?P" in thin_rl)
   1.344 +apply  (rule_tac [2] P' = " (\<lambda>b s (Y',s') . (\<exists>s0. G\<turnstile>s0 \<midarrow>In1l expr\<succ>\<rightarrow> (b,s)) \<and> (if normal s \<and> the_Bool (the_In1 b) then (\<forall>s'' w s0. G\<turnstile>s \<midarrow>stmnt\<rightarrow> s'' \<and> G\<turnstile>(abupd (absorb (Cont l)) s'') \<midarrow>In1r (l\<bullet> While(expr) stmnt) \<succ>\<rightarrow> (w,s0) \<longrightarrow> (w,s0) = (Y',s')) else (\<diamondsuit>,s) = (Y',s'))) \<and>. G\<turnstile>init\<le>n" in polymorphic_Loop)
   1.345 +apply   (force dest!: eval.Loop split add: split_if_asm)
   1.346 +prefer 2
   1.347 +apply  (erule MGFnD' [THEN conseq12])
   1.348 +apply  clarsimp
   1.349 +apply  (erule_tac V = "card (nyinitcls G s') \<le> n" in thin_rl)
   1.350 +apply  (tactic "eval_Force_tac 1")
   1.351 +apply (erule MGFnD' [THEN conseq12] , clarsimp)
   1.352 +apply (rule conjI, erule exI)
   1.353 +apply (tactic "clarsimp_tac eval_css 1")
   1.354 +apply (case_tac "a")
   1.355 +prefer 2
   1.356 +apply  (clarsimp)
   1.357 +apply (clarsimp split add: split_if)
   1.358 +apply (rule conjI, (tactic {* force_tac (claset() addSDs [thm "eval.Loop"],
   1.359 +  simpset() addsimps [split_paired_all] addsimprocs [eval_stmt_proc]) 1*})+)
   1.360 +done
   1.361 +
   1.362 +lemma MGFn_lemma [rule_format (no_asm)]: 
   1.363 + "\<forall>n C sig. G,(A::state triple set)\<turnstile>{=:n} In1l (Methd C sig)\<succ> {G\<rightarrow>} \<Longrightarrow>  
   1.364 +  \<forall>t. G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>}"
   1.365 +apply (rule full_nat_induct)
   1.366 +apply (rule allI)
   1.367 +apply (drule_tac x = n in spec)
   1.368 +apply (drule_tac psi = "All ?P" in asm_rl)
   1.369 +apply (subgoal_tac "\<forall>v e c es. G,A\<turnstile>{=:n} In2 v\<succ> {G\<rightarrow>} \<and> G,A\<turnstile>{=:n} In1l e\<succ> {G\<rightarrow>} \<and> G,A\<turnstile>{=:n} In1r c\<succ> {G\<rightarrow>} \<and> G,A\<turnstile>{=:n} In3 es\<succ> {G\<rightarrow>}")
   1.370 +apply  (tactic "Clarify_tac 2")
   1.371 +apply  (induct_tac "t")
   1.372 +apply    (induct_tac "a")
   1.373 +apply     fast+
   1.374 +apply (rule var_expr_stmt.induct)
   1.375 +(* 28 subgoals *)
   1.376 +prefer 14 apply fast (* Methd *)
   1.377 +prefer 13 apply (erule (2) MGFn_Call)
   1.378 +apply (erule_tac [!] V = "All ?P" in thin_rl) (* assumptions on Methd *)
   1.379 +apply (erule_tac [24] MGFn_Init)
   1.380 +prefer 19 apply (erule (1) MGFn_Loop)
   1.381 +apply (tactic "ALLGOALS compl_prepare_tac")
   1.382 +
   1.383 +apply (rule ax_derivs.LVar [THEN conseq1], tactic "eval_Force_tac 1")
   1.384 +
   1.385 +apply (rule ax_derivs.FVar)
   1.386 +apply  (erule MGFn_InitD)
   1.387 +apply (tactic "forw_hyp_eval_Force_tac 1")
   1.388 +
   1.389 +apply (rule ax_derivs.AVar)
   1.390 +apply  (erule MGFnD [THEN ax_NormalD])
   1.391 +apply (tactic "forw_hyp_eval_Force_tac 1")
   1.392 +
   1.393 +apply (rule ax_derivs.NewC)
   1.394 +apply (erule MGFn_InitD [THEN conseq2])
   1.395 +apply (tactic "eval_Force_tac 1")
   1.396 +
   1.397 +apply (rule_tac Q = "(\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>In1r (init_comp_ty ty) \<succ>\<rightarrow> (Y',s')) \<and>. G\<turnstile>init\<le>n" in ax_derivs.NewA)
   1.398 +apply  (simp add: init_comp_ty_def split add: split_if)
   1.399 +apply   (rule conjI, clarsimp)
   1.400 +apply   (erule MGFn_InitD [THEN conseq2])
   1.401 +apply   (tactic "clarsimp_tac eval_css 1")
   1.402 +apply  clarsimp
   1.403 +apply  (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1")
   1.404 +apply (tactic "forw_hyp_eval_Force_tac 1")
   1.405 +
   1.406 +apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast],tactic"eval_Force_tac 1")
   1.407 +
   1.408 +apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst],tactic"eval_Force_tac 1")
   1.409 +apply (rule ax_derivs.Lit [THEN conseq1], tactic "eval_Force_tac 1")
   1.410 +apply (rule ax_derivs.Super [THEN conseq1], tactic "eval_Force_tac 1")
   1.411 +apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc],tactic"eval_Force_tac 1")
   1.412 +
   1.413 +apply (rule ax_derivs.Ass)
   1.414 +apply  (erule MGFnD [THEN ax_NormalD])
   1.415 +apply (tactic "forw_hyp_eval_Force_tac 1")
   1.416 +
   1.417 +apply (rule ax_derivs.Cond)
   1.418 +apply  (erule MGFnD [THEN ax_NormalD])
   1.419 +apply (rule allI)
   1.420 +apply (rule ax_Normal_cases)
   1.421 +prefer 2
   1.422 +apply  (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
   1.423 +apply  (tactic "eval_Force_tac 1")
   1.424 +apply (case_tac "b")
   1.425 +apply  (simp, tactic "forw_hyp_eval_Force_tac 1")
   1.426 +apply (simp, tactic "forw_hyp_eval_Force_tac 1")
   1.427 +
   1.428 +apply (rule_tac Q = " (\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>Init pid_field_type\<rightarrow> s') \<and>. G\<turnstile>init\<le>n" in ax_derivs.Body)
   1.429 + apply (erule MGFn_InitD [THEN conseq2])
   1.430 + apply (tactic "eval_Force_tac 1")
   1.431 +apply (tactic "forw_hyp_tac 1")
   1.432 +apply (tactic {* clarsimp_tac (eval_css delsimps2 [split_paired_all]) 1 *})
   1.433 +apply (erule (1) eval.Body)
   1.434 +
   1.435 +apply (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1")
   1.436 +
   1.437 +apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr],tactic"eval_Force_tac 1")
   1.438 +
   1.439 +apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
   1.440 +apply (tactic "clarsimp_tac eval_css 1")
   1.441 +
   1.442 +apply (rule ax_derivs.Comp)
   1.443 +apply  (erule MGFnD [THEN ax_NormalD])
   1.444 +apply (tactic "forw_hyp_eval_Force_tac 1")
   1.445 +
   1.446 +apply (rule ax_derivs.If)
   1.447 +apply  (erule MGFnD [THEN ax_NormalD])
   1.448 +apply (rule allI)
   1.449 +apply (rule ax_Normal_cases)
   1.450 +prefer 2
   1.451 +apply  (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
   1.452 +apply  (tactic "eval_Force_tac 1")
   1.453 +apply (case_tac "b")
   1.454 +apply  (simp, tactic "forw_hyp_eval_Force_tac 1")
   1.455 +apply (simp, tactic "forw_hyp_eval_Force_tac 1")
   1.456 +
   1.457 +apply (rule ax_derivs.Do [THEN conseq1])
   1.458 +apply (tactic {* force_tac (eval_css addsimps2 [thm "abupd_def2"]) 1 *})
   1.459 +
   1.460 +apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
   1.461 +apply (tactic "clarsimp_tac eval_css 1")
   1.462 +
   1.463 +apply (rule_tac Q = " (\<lambda>Y' s' s. normal s \<and> (\<exists>s''. G\<turnstile>s \<midarrow>In1r stmt1\<succ>\<rightarrow> (Y',s'') \<and> G\<turnstile>s'' \<midarrow>sxalloc\<rightarrow> s')) \<and>. G\<turnstile>init\<le>n" in ax_derivs.Try)
   1.464 +apply   (tactic "eval_Force_tac 3")
   1.465 +apply  (tactic "forw_hyp_eval_Force_tac 2")
   1.466 +apply (erule MGFnD [THEN ax_NormalD, THEN conseq2])
   1.467 +apply (tactic "clarsimp_tac eval_css 1")
   1.468 +apply (force elim: sxalloc_gext [THEN card_nyinitcls_gext])
   1.469 +
   1.470 +apply (rule_tac Q = " (\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>stmt1\<rightarrow> s') \<and>. G\<turnstile>init\<le>n" in ax_derivs.Fin)
   1.471 +apply  (tactic "forw_hyp_eval_Force_tac 1")
   1.472 +apply (rule allI)
   1.473 +apply (tactic "forw_hyp_tac 1")
   1.474 +apply (tactic {* pair_tac "sb" 1 *})
   1.475 +apply (tactic"clarsimp_tac (claset(),simpset() addsimprocs [eval_stmt_proc]) 1")
   1.476 +apply (drule (1) eval.Fin)
   1.477 +apply clarsimp
   1.478 +
   1.479 +apply (rule ax_derivs.Nil [THEN conseq1], tactic "eval_Force_tac 1")
   1.480 +
   1.481 +apply (rule ax_derivs.Cons)
   1.482 +apply  (erule MGFnD [THEN ax_NormalD])
   1.483 +apply (tactic "forw_hyp_eval_Force_tac 1")
   1.484 +done
   1.485 +
   1.486 +lemma MGF_asm: "\<forall>C sig. is_methd G C sig \<longrightarrow> G,A\<turnstile>{\<doteq>} In1l (Methd C sig)\<succ> {G\<rightarrow>} \<Longrightarrow>
   1.487 +  G,(A::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
   1.488 +apply (simp (no_asm_use) add: MGF_MGFn_iff)
   1.489 +apply (rule allI)
   1.490 +apply (rule MGFn_lemma)
   1.491 +apply (intro strip)
   1.492 +apply (rule MGFn_free_wt)
   1.493 +apply (force dest: wt_Methd_is_methd)
   1.494 +done
   1.495 +
   1.496 +declare splitI2 [intro!]
   1.497 +ML_setup {*
   1.498 +Addsimprocs [ eval_expr_proc, eval_var_proc, eval_exprs_proc, eval_stmt_proc]
   1.499 +*}
   1.500 +
   1.501 +
   1.502 +section "nested version"
   1.503 +
   1.504 +lemma nesting_lemma' [rule_format (no_asm)]: "[| !!A ts. ts <= A ==> P A ts; 
   1.505 +  !!A pn. !b:bdy pn. P (insert (mgf_call pn) A) {mgf b} ==> P A {mgf_call pn}; 
   1.506 +  !!A t. !pn:U. P A {mgf_call pn} ==> P A {mgf t};  
   1.507 +          finite U; uA = mgf_call`U |] ==>  
   1.508 +  !A. A <= uA --> n <= card uA --> card A = card uA - n --> (!t. P A {mgf t})"
   1.509 +proof -
   1.510 +  assume ax_derivs_asm:    "!!A ts. ts <= A ==> P A ts"
   1.511 +  assume MGF_nested_Methd: "!!A pn. !b:bdy pn. P (insert (mgf_call pn) A) 
   1.512 +                                                  {mgf b} ==> P A {mgf_call pn}"
   1.513 +  assume MGF_asm:          "!!A t. !pn:U. P A {mgf_call pn} ==> P A {mgf t}"
   1.514 +  assume "finite U" "uA = mgf_call`U"
   1.515 +  then show ?thesis
   1.516 +    apply -
   1.517 +    apply (induct_tac "n")
   1.518 +    apply  (tactic "ALLGOALS Clarsimp_tac")
   1.519 +    apply  (tactic "dtac (permute_prems 0 1 card_seteq) 1")
   1.520 +    apply    simp
   1.521 +    apply   (erule finite_imageI)
   1.522 +    apply  (simp add: MGF_asm ax_derivs_asm)
   1.523 +    apply (rule MGF_asm)
   1.524 +    apply (rule ballI)
   1.525 +    apply (case_tac "mgf_call pn : A")
   1.526 +    apply  (fast intro: ax_derivs_asm)
   1.527 +    apply (rule MGF_nested_Methd)
   1.528 +    apply (rule ballI)
   1.529 +    apply (drule spec, erule impE, erule_tac [2] impE, erule_tac [3] impE, 
   1.530 +           erule_tac [4] spec)
   1.531 +    apply   fast
   1.532 +    apply  (erule Suc_leD)
   1.533 +    apply (drule finite_subset)
   1.534 +    apply (erule finite_imageI)
   1.535 +    apply auto
   1.536 +    apply arith
   1.537 +  done
   1.538 +qed
   1.539 +
   1.540 +lemma nesting_lemma [rule_format (no_asm)]: "[| !!A ts. ts <= A ==> P A ts; 
   1.541 +  !!A pn. !b:bdy pn. P (insert (mgf (f pn)) A) {mgf b} ==> P A {mgf (f pn)}; 
   1.542 +          !!A t. !pn:U. P A {mgf (f pn)} ==> P A {mgf t}; 
   1.543 +          finite U |] ==> P {} {mgf t}"
   1.544 +proof -
   1.545 +  assume 2: "!!A pn. !b:bdy pn. P (insert (mgf (f pn)) A) {mgf b} ==> P A {mgf (f pn)}"
   1.546 +  assume 3: "!!A t. !pn:U. P A {mgf (f pn)} ==> P A {mgf t}"
   1.547 +  assume "!!A ts. ts <= A ==> P A ts" "finite U"
   1.548 +  then show ?thesis
   1.549 +    apply -
   1.550 +    apply (rule_tac mgf = "mgf" in nesting_lemma')
   1.551 +    apply (erule_tac [2] 2)
   1.552 +    apply (rule_tac [2] 3)
   1.553 +    apply (rule_tac [6] le_refl)
   1.554 +    apply auto
   1.555 +  done
   1.556 +qed
   1.557 +
   1.558 +lemma MGF_nested_Methd: "\<lbrakk>  
   1.559 +  G,insert ({Normal \<doteq>} In1l (Methd  C sig) \<succ>{G\<rightarrow>}) A\<turnstile>  
   1.560 +            {Normal \<doteq>} In1l (body G C sig) \<succ>{G\<rightarrow>}  
   1.561 + \<rbrakk> \<Longrightarrow>  G,A\<turnstile>{Normal \<doteq>} In1l (Methd  C sig) \<succ>{G\<rightarrow>}"
   1.562 +apply (unfold MGF_def)
   1.563 +apply (rule ax_MethdN)
   1.564 +apply (erule conseq2)
   1.565 +apply clarsimp
   1.566 +apply (erule MethdI)
   1.567 +done
   1.568 +
   1.569 +lemma MGF_deriv: "ws_prog G \<Longrightarrow> G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
   1.570 +apply (rule MGFNormalI)
   1.571 +apply (rule_tac mgf = "\<lambda>t. {Normal \<doteq>} t\<succ> {G\<rightarrow>}" and 
   1.572 +                bdy = "\<lambda> (C,sig) .{In1l (body G C sig) }" and 
   1.573 +                f = "\<lambda> (C,sig) . In1l (Methd C sig) " in nesting_lemma)
   1.574 +apply    (erule ax_derivs.asm)
   1.575 +apply   (clarsimp simp add: split_tupled_all)
   1.576 +apply   (erule MGF_nested_Methd)
   1.577 +apply  (erule_tac [2] finite_is_methd)
   1.578 +apply (rule MGF_asm [THEN MGFNormalD])
   1.579 +apply clarify
   1.580 +apply (rule MGFNormalI)
   1.581 +apply force
   1.582 +done
   1.583 +
   1.584 +
   1.585 +section "simultaneous version"
   1.586 +
   1.587 +lemma MGF_simult_Methd_lemma: "finite ms \<Longrightarrow>  
   1.588 +  G,A\<union> (\<lambda>(C,sig). {Normal \<doteq>} In1l (Methd  C sig)\<succ> {G\<rightarrow>}) ` ms  
   1.589 +     |\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} In1l (body G C sig)\<succ> {G\<rightarrow>}) ` ms \<Longrightarrow>  
   1.590 +  G,A|\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} In1l (Methd  C sig)\<succ> {G\<rightarrow>}) ` ms"
   1.591 +apply (unfold MGF_def)
   1.592 +apply (rule ax_derivs.Methd [unfolded mtriples_def])
   1.593 +apply (erule ax_finite_pointwise)
   1.594 +prefer 2
   1.595 +apply  (rule ax_derivs.asm)
   1.596 +apply  fast
   1.597 +apply clarsimp
   1.598 +apply (rule conseq2)
   1.599 +apply  (erule (1) ax_methods_spec)
   1.600 +apply clarsimp
   1.601 +apply (erule eval_Methd)
   1.602 +done
   1.603 +
   1.604 +lemma MGF_simult_Methd: "ws_prog G \<Longrightarrow> 
   1.605 +   G,({}::state triple set)|\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} In1l (Methd C sig)\<succ> {G\<rightarrow>}) 
   1.606 +   ` Collect (split (is_methd G)) "
   1.607 +apply (frule finite_is_methd)
   1.608 +apply (rule MGF_simult_Methd_lemma)
   1.609 +apply  assumption
   1.610 +apply (erule ax_finite_pointwise)
   1.611 +prefer 2
   1.612 +apply  (rule ax_derivs.asm)
   1.613 +apply  blast
   1.614 +apply clarsimp
   1.615 +apply (rule MGF_asm [THEN MGFNormalD])
   1.616 +apply clarify
   1.617 +apply (rule MGFNormalI)
   1.618 +apply force
   1.619 +done
   1.620 +
   1.621 +lemma MGF_deriv: "ws_prog G \<Longrightarrow> G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
   1.622 +apply (rule MGF_asm)
   1.623 +apply (intro strip)
   1.624 +apply (rule MGFNormalI)
   1.625 +apply (rule ax_derivs.weaken)
   1.626 +apply  (erule MGF_simult_Methd)
   1.627 +apply force
   1.628 +done
   1.629 +
   1.630 +
   1.631 +section "corollaries"
   1.632 +
   1.633 +lemma MGF_complete: "G,{}\<Turnstile>{P} t\<succ> {Q} \<Longrightarrow> G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow>
   1.634 +  G,({}::state triple set)\<turnstile>{P::state assn} t\<succ> {Q}"
   1.635 +apply (rule ax_no_hazard)
   1.636 +apply (unfold MGF_def)
   1.637 +apply (erule conseq12)
   1.638 +apply (simp (no_asm_use) add: ax_valids_def triple_valid_def)
   1.639 +apply (fast dest!: eval_evaln)
   1.640 +done
   1.641 +
   1.642 +theorem ax_complete: "ws_prog G \<Longrightarrow>  
   1.643 +  G,{}\<Turnstile>{P::state assn} t\<succ> {Q} \<Longrightarrow> G,({}::state triple set)\<turnstile>{P} t\<succ> {Q}"
   1.644 +apply (erule MGF_complete)
   1.645 +apply (erule MGF_deriv)
   1.646 +done
   1.647 +
   1.648 +end
     2.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     2.2 +++ b/src/HOL/Bali/AxExample.thy	Mon Jan 28 17:00:19 2002 +0100
     2.3 @@ -0,0 +1,270 @@
     2.4 +(*  Title:      isabelle/Bali/AxExample.thy
     2.5 +    ID:         $Id$
     2.6 +    Author:     David von Oheimb
     2.7 +    Copyright   2000 Technische Universitaet Muenchen
     2.8 +*)
     2.9 +header {* Example of a proof based on the Bali axiomatic semantics *}
    2.10 +
    2.11 +theory AxExample = AxSem + Example:
    2.12 +
    2.13 +constdefs
    2.14 +  arr_inv :: "st \<Rightarrow> bool"
    2.15 + "arr_inv \<equiv> \<lambda>s. \<exists>obj a T el. globs s (Stat Base) = Some obj \<and>
    2.16 +                              values obj (Inl (arr, Base)) = Some (Addr a) \<and>
    2.17 +                              heap s a = Some \<lparr>tag=Arr T 2,values=el\<rparr>"
    2.18 +
    2.19 +lemma arr_inv_new_obj: 
    2.20 +"\<And>a. \<lbrakk>arr_inv s; new_Addr (heap s)=Some a\<rbrakk> \<Longrightarrow> arr_inv (gupd(Inl a\<mapsto>x) s)"
    2.21 +apply (unfold arr_inv_def)
    2.22 +apply (force dest!: new_AddrD2)
    2.23 +done
    2.24 +
    2.25 +lemma arr_inv_set_locals [simp]: "arr_inv (set_locals l s) = arr_inv s"
    2.26 +apply (unfold arr_inv_def)
    2.27 +apply (simp (no_asm))
    2.28 +done
    2.29 +
    2.30 +lemma arr_inv_gupd_Stat [simp]: 
    2.31 +  "Base \<noteq> C \<Longrightarrow> arr_inv (gupd(Stat C\<mapsto>obj) s) = arr_inv s"
    2.32 +apply (unfold arr_inv_def)
    2.33 +apply (simp (no_asm_simp))
    2.34 +done
    2.35 +
    2.36 +lemma ax_inv_lupd [simp]: "arr_inv (lupd(x\<mapsto>y) s) = arr_inv s"
    2.37 +apply (unfold arr_inv_def)
    2.38 +apply (simp (no_asm))
    2.39 +done
    2.40 +
    2.41 +
    2.42 +declare split_if_asm [split del]
    2.43 +declare lvar_def [simp]
    2.44 +
    2.45 +ML {*
    2.46 +fun inst1_tac s t = instantiate_tac [(s,t)];
    2.47 +val ax_tac = REPEAT o rtac allI THEN'
    2.48 +             resolve_tac(thm "ax_Skip"::thm "ax_StatRef"::thm "ax_MethdN"::
    2.49 +                         thm "ax_Alloc"::thm "ax_Alloc_Arr"::
    2.50 +                         thm "ax_SXAlloc_Normal"::
    2.51 +                         funpow 7 tl (thms "ax_derivs.intros"))
    2.52 +*}
    2.53 +
    2.54 +
    2.55 +theorem ax_test: "tprg,({}::'a triple set)\<turnstile> 
    2.56 +  {Normal (\<lambda>Y s Z::'a. heap_free four s \<and> \<not>initd Base s \<and> \<not> initd Ext s)} 
    2.57 +  .test [Class Base]. {\<lambda>Y s Z. abrupt s = Some (Xcpt (Std IndOutBound))}"
    2.58 +apply (unfold test_def arr_viewed_from_def)
    2.59 +apply (tactic "ax_tac 1" (*;;*))
    2.60 +defer
    2.61 +apply  (tactic "ax_tac 1" (* Try *))
    2.62 +defer
    2.63 +apply    (tactic {* inst1_tac "Q1" 
    2.64 +                 "\<lambda>Y s Z. arr_inv (snd s) \<and> tprg,s\<turnstile>catch SXcpt NullPointer" *})
    2.65 +prefer 2
    2.66 +apply    simp
    2.67 +apply   (rule_tac P' = "Normal (\<lambda>Y s Z. arr_inv (snd s))" in conseq1)
    2.68 +prefer 2
    2.69 +apply    clarsimp
    2.70 +apply   (rule_tac Q' = "(\<lambda>Y s Z. ?Q Y s Z)\<leftarrow>=False\<down>=\<diamondsuit>" in conseq2)
    2.71 +prefer 2
    2.72 +apply    simp
    2.73 +apply   (tactic "ax_tac 1" (* While *))
    2.74 +prefer 2
    2.75 +apply    (rule ax_impossible [THEN conseq1], clarsimp)
    2.76 +apply   (rule_tac P' = "Normal ?P" in conseq1)
    2.77 +prefer 2
    2.78 +apply    clarsimp
    2.79 +apply   (tactic "ax_tac 1")
    2.80 +apply   (tactic "ax_tac 1" (* AVar *))
    2.81 +prefer 2
    2.82 +apply    (rule ax_subst_Val_allI)
    2.83 +apply    (tactic {* inst1_tac "P'21" "\<lambda>u a. Normal (?PP a\<leftarrow>?x) u" *})
    2.84 +apply    (simp del: avar_def2 peek_and_def2)
    2.85 +apply    (tactic "ax_tac 1")
    2.86 +apply   (tactic "ax_tac 1")
    2.87 +      (* just for clarification: *)
    2.88 +apply   (rule_tac Q' = "Normal (\<lambda>Var:(v, f) u ua. fst (snd (avar tprg (Intg 2) v u)) = Some (Xcpt (Std IndOutBound)))" in conseq2)
    2.89 +prefer 2
    2.90 +apply    (clarsimp simp add: split_beta)
    2.91 +apply   (tactic "ax_tac 1" (* FVar *))
    2.92 +apply    (tactic "ax_tac 2" (* StatRef *))
    2.93 +apply   (rule ax_derivs.Done [THEN conseq1])
    2.94 +apply   (clarsimp simp add: arr_inv_def inited_def in_bounds_def)
    2.95 +defer
    2.96 +apply  (rule ax_SXAlloc_catch_SXcpt)
    2.97 +apply  (rule_tac Q' = "(\<lambda>Y (x, s) Z. x = Some (Xcpt (Std NullPointer)) \<and> arr_inv s) \<and>. heap_free two" in conseq2)
    2.98 +prefer 2
    2.99 +apply   (simp add: arr_inv_new_obj)
   2.100 +apply  (tactic "ax_tac 1") 
   2.101 +apply  (rule_tac C = "Ext" in ax_Call_known_DynT)
   2.102 +apply     (unfold DynT_prop_def)
   2.103 +apply     (simp (no_asm))
   2.104 +apply    (intro strip)
   2.105 +apply    (rule_tac P' = "Normal ?P" in conseq1)
   2.106 +apply     (tactic "ax_tac 1" (* Methd *))
   2.107 +apply     (rule ax_thin [OF _ empty_subsetI])
   2.108 +apply     (simp (no_asm) add: body_def2)
   2.109 +apply     (tactic "ax_tac 1" (* Body *))
   2.110 +(* apply       (rule_tac [2] ax_derivs.Abrupt) *)
   2.111 +defer
   2.112 +apply      (simp (no_asm))
   2.113 +apply      (tactic "ax_tac 1")
   2.114 +apply      (tactic "ax_tac 1") (* Ass *)
   2.115 +prefer 2
   2.116 +apply       (rule ax_subst_Var_allI)
   2.117 +apply       (tactic {* inst1_tac "P'27" "\<lambda>a vs l vf. ?PP a vs l vf\<leftarrow>?x \<and>. ?p" *})
   2.118 +apply       (rule allI)
   2.119 +apply       (tactic {* simp_tac (simpset() delloop "split_all_tac" delsimps [thm "peek_and_def2"]) 1 *})
   2.120 +apply       (rule ax_derivs.Abrupt)
   2.121 +apply      (simp (no_asm))
   2.122 +apply      (tactic "ax_tac 1" (* FVar *))
   2.123 +apply       (tactic "ax_tac 2", tactic "ax_tac 2", tactic "ax_tac 2")
   2.124 +apply      (tactic "ax_tac 1")
   2.125 +apply     clarsimp
   2.126 +apply     (tactic {* inst1_tac "R14" "\<lambda>a'. Normal ((\<lambda>Vals:vs (x, s) Z. arr_inv s \<and> inited Ext (globs s) \<and> a' \<noteq> Null \<and> hd vs = Null) \<and>. heap_free two)" *})
   2.127 +prefer 5
   2.128 +apply     (rule ax_derivs.Done [THEN conseq1], force)
   2.129 +apply    force
   2.130 +apply   (rule ax_subst_Val_allI)
   2.131 +apply   (tactic {* inst1_tac "P'33" "\<lambda>u a. Normal (?PP a\<leftarrow>?x) u" *})
   2.132 +apply   (simp (no_asm) del: peek_and_def2)
   2.133 +apply   (tactic "ax_tac 1")
   2.134 +prefer 2
   2.135 +apply   (rule ax_subst_Val_allI)
   2.136 +apply    (tactic {* inst1_tac "P'4" "\<lambda>aa v. Normal (?QQ aa v\<leftarrow>?y)" *})
   2.137 +apply    (simp del: peek_and_def2)
   2.138 +apply    (tactic "ax_tac 1")
   2.139 +apply   (tactic "ax_tac 1")
   2.140 +apply  (tactic "ax_tac 1")
   2.141 +apply  (tactic "ax_tac 1")
   2.142 +(* end method call *)
   2.143 +apply (simp (no_asm))
   2.144 +    (* just for clarification: *)
   2.145 +apply (rule_tac Q' = "Normal ((\<lambda>Y (x, s) Z. arr_inv s \<and> (\<exists>a. the (locals s (VName e)) = Addr a \<and> obj_class (the (globs s (Inl a))) = Ext \<and> 
   2.146 + invocation_declclass tprg IntVir s (the (locals s (VName e))) (ClassT Base)  
   2.147 +     \<lparr>name = foo, parTs = [Class Base]\<rparr> = Ext)) \<and>. initd Ext \<and>. heap_free two)"
   2.148 +  in conseq2)
   2.149 +prefer 2
   2.150 +apply  clarsimp
   2.151 +apply (tactic "ax_tac 1")
   2.152 +apply (tactic "ax_tac 1")
   2.153 +defer
   2.154 +apply  (rule ax_subst_Var_allI)
   2.155 +apply  (tactic {* inst1_tac "P'14" "\<lambda>u vf. Normal (?PP vf \<and>. ?p) u" *})
   2.156 +apply  (simp (no_asm) del: split_paired_All peek_and_def2)
   2.157 +apply  (tactic "ax_tac 1" (* NewC *))
   2.158 +apply  (tactic "ax_tac 1" (* ax_Alloc *))
   2.159 +     (* just for clarification: *)
   2.160 +apply  (rule_tac Q' = "Normal ((\<lambda>Y s Z. arr_inv (store s) \<and> vf=lvar (VName e) (store s)) \<and>. heap_free tree \<and>. initd Ext)" in conseq2)
   2.161 +prefer 2
   2.162 +apply   (simp add: invocation_declclass_def dynmethd_def)
   2.163 +apply   (unfold dynlookup_def)
   2.164 +apply   (simp add: dynmethd_Ext_foo)
   2.165 +apply   (force elim!: arr_inv_new_obj atleast_free_SucD atleast_free_weaken)
   2.166 +     (* begin init *)
   2.167 +apply  (rule ax_InitS)
   2.168 +apply     force
   2.169 +apply    (simp (no_asm))
   2.170 +apply   (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *})
   2.171 +apply   (rule ax_Init_Skip_lemma)
   2.172 +apply  (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *})
   2.173 +apply  (rule ax_InitS [THEN conseq1] (* init Base *))
   2.174 +apply      force
   2.175 +apply     (simp (no_asm))
   2.176 +apply    (unfold arr_viewed_from_def)
   2.177 +apply    (rule allI)
   2.178 +apply    (rule_tac P' = "Normal ?P" in conseq1)
   2.179 +apply     (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *})
   2.180 +apply     (tactic "ax_tac 1")
   2.181 +apply     (tactic "ax_tac 1")
   2.182 +apply     (rule_tac [2] ax_subst_Var_allI)
   2.183 +apply      (tactic {* inst1_tac "P'29" "\<lambda>vf l vfa. Normal (?P vf l vfa)" *})
   2.184 +apply     (tactic {* simp_tac (simpset() delloop "split_all_tac" delsimps [split_paired_All, thm "peek_and_def2"]) 2 *})
   2.185 +apply      (tactic "ax_tac 2" (* NewA *))
   2.186 +apply       (tactic "ax_tac 3" (* ax_Alloc_Arr *))
   2.187 +apply       (tactic "ax_tac 3")
   2.188 +apply      (tactic {* inst1_tac "P" "\<lambda>vf l vfa. Normal (?P vf l vfa\<leftarrow>\<diamondsuit>)" *})
   2.189 +apply      (tactic {* simp_tac (simpset() delloop "split_all_tac") 2 *})
   2.190 +apply      (tactic "ax_tac 2")
   2.191 +apply     (tactic "ax_tac 1" (* FVar *))
   2.192 +apply      (tactic "ax_tac 2" (* StatRef *))
   2.193 +apply     (rule ax_derivs.Done [THEN conseq1])
   2.194 +apply     (tactic {* inst1_tac "Q22" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf=lvar (VName e) (snd s)) \<and>. heap_free four \<and>. initd Base \<and>. initd Ext)" *})
   2.195 +apply     (clarsimp split del: split_if)
   2.196 +apply     (frule atleast_free_weaken [THEN atleast_free_weaken])
   2.197 +apply     (drule initedD)
   2.198 +apply     (clarsimp elim!: atleast_free_SucD simp add: arr_inv_def)
   2.199 +apply    force
   2.200 +apply   (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *})
   2.201 +apply   (rule ax_triv_Init_Object [THEN peek_and_forget2, THEN conseq1])
   2.202 +apply     (rule wf_tprg)
   2.203 +apply    clarsimp
   2.204 +apply   (tactic {* inst1_tac "P22" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf = lvar (VName e) (snd s)) \<and>. heap_free four \<and>. initd Ext)" *})
   2.205 +apply   clarsimp
   2.206 +apply  (tactic {* inst1_tac "PP" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf = lvar (VName e) (snd s)) \<and>. heap_free four \<and>. Not \<circ> initd Base)" *})
   2.207 +apply  clarsimp
   2.208 +     (* end init *)
   2.209 +apply (rule conseq1)
   2.210 +apply (tactic "ax_tac 1")
   2.211 +apply clarsimp
   2.212 +done
   2.213 +
   2.214 +(*
   2.215 +while (true) {
   2.216 +  if (i) {throw xcpt;}
   2.217 +  else i=j
   2.218 +}
   2.219 +*)
   2.220 +lemma Loop_Xcpt_benchmark: 
   2.221 + "Q = (\<lambda>Y (x,s) Z. x \<noteq> None \<longrightarrow> the_Bool (the (locals s i))) \<Longrightarrow>  
   2.222 +  G,({}::'a triple set)\<turnstile>{Normal (\<lambda>Y s Z::'a. True)}  
   2.223 +  .lab1\<bullet> While(Lit (Bool True)) (If(Acc (LVar i)) (Throw (Acc (LVar xcpt))) Else
   2.224 +        (Expr (Ass (LVar i) (Acc (LVar j))))). {Q}"
   2.225 +apply (rule_tac P' = "Q" and Q' = "Q\<leftarrow>=False\<down>=\<diamondsuit>" in conseq12)
   2.226 +apply  safe
   2.227 +apply  (tactic "ax_tac 1" (* Loop *))
   2.228 +apply   (rule ax_Normal_cases)
   2.229 +prefer 2
   2.230 +apply    (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
   2.231 +apply   (rule conseq1)
   2.232 +apply    (tactic "ax_tac 1")
   2.233 +apply   clarsimp
   2.234 +prefer 2
   2.235 +apply  clarsimp
   2.236 +apply (tactic "ax_tac 1" (* If *))
   2.237 +apply  (tactic 
   2.238 +  {* inst1_tac "P'21" "Normal (\<lambda>s.. (\<lambda>Y s Z. True)\<down>=Val (the (locals s i)))" *})
   2.239 +apply  (tactic "ax_tac 1")
   2.240 +apply  (rule conseq1)
   2.241 +apply   (tactic "ax_tac 1")
   2.242 +apply  clarsimp
   2.243 +apply (rule allI)
   2.244 +apply (rule ax_escape)
   2.245 +apply auto
   2.246 +apply  (rule conseq1)
   2.247 +apply   (tactic "ax_tac 1" (* Throw *))
   2.248 +apply   (tactic "ax_tac 1")
   2.249 +apply   (tactic "ax_tac 1")
   2.250 +apply  clarsimp
   2.251 +apply (rule_tac Q' = "Normal (\<lambda>Y s Z. True)" in conseq2)
   2.252 +prefer 2
   2.253 +apply  clarsimp
   2.254 +apply (rule conseq1)
   2.255 +apply  (tactic "ax_tac 1")
   2.256 +apply  (tactic "ax_tac 1")
   2.257 +prefer 2
   2.258 +apply   (rule ax_subst_Var_allI)
   2.259 +apply   (tactic {* inst1_tac "P'29" "\<lambda>b Y ba Z vf. \<lambda>Y (x,s) Z. x=None \<and> snd vf = snd (lvar i s)" *})
   2.260 +apply   (rule allI)
   2.261 +apply   (rule_tac P' = "Normal ?P" in conseq1)
   2.262 +prefer 2
   2.263 +apply    clarsimp
   2.264 +apply   (tactic "ax_tac 1")
   2.265 +apply   (rule conseq1)
   2.266 +apply    (tactic "ax_tac 1")
   2.267 +apply   clarsimp
   2.268 +apply  (tactic "ax_tac 1")
   2.269 +apply clarsimp
   2.270 +done
   2.271 +
   2.272 +end
   2.273 +
     3.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     3.2 +++ b/src/HOL/Bali/AxSem.thy	Mon Jan 28 17:00:19 2002 +0100
     3.3 @@ -0,0 +1,1191 @@
     3.4 +(*  Title:      isabelle/Bali/AxSem.thy
     3.5 +    ID:         $Id$
     3.6 +    Author:     David von Oheimb
     3.7 +    Copyright   1998 Technische Universitaet Muenchen
     3.8 +*)
     3.9 +
    3.10 +header {* Axiomatic semantics of Java expressions and statements 
    3.11 +          (see also Eval.thy)
    3.12 +        *}
    3.13 +
    3.14 +theory AxSem = Evaln + TypeSafe:
    3.15 +
    3.16 +text {*
    3.17 +design issues:
    3.18 +\begin{itemize}
    3.19 +\item a strong version of validity for triples with premises, namely one that 
    3.20 +      takes the recursive depth needed to complete execution, enables 
    3.21 +      correctness proof
    3.22 +\item auxiliary variables are handled first-class (-> Thomas Kleymann)
    3.23 +\item expressions not flattened to elementary assignments (as usual for 
    3.24 +      axiomatic semantics) but treated first-class => explicit result value 
    3.25 +      handling
    3.26 +\item intermediate values not on triple, but on assertion level 
    3.27 +      (with result entry)
    3.28 +\item multiple results with semantical substitution mechnism not requiring a 
    3.29 +      stack 
    3.30 +\item because of dynamic method binding, terms need to be dependent on state.
    3.31 +  this is also useful for conditional expressions and statements
    3.32 +\item result values in triples exactly as in eval relation (also for xcpt 
    3.33 +      states)
    3.34 +\item validity: additional assumption of state conformance and well-typedness,
    3.35 +  which is required for soundness and thus rule hazard required of completeness
    3.36 +\end{itemize}
    3.37 +
    3.38 +restrictions:
    3.39 +\begin{itemize}
    3.40 +\item all triples in a derivation are of the same type (due to weak 
    3.41 +      polymorphism)
    3.42 +\end{itemize}
    3.43 +*}
    3.44 +
    3.45 +
    3.46 +
    3.47 +types  res = vals (* result entry *)
    3.48 +syntax
    3.49 +  Val  :: "val      \<Rightarrow> res"
    3.50 +  Var  :: "var      \<Rightarrow> res"
    3.51 +  Vals :: "val list \<Rightarrow> res"
    3.52 +translations
    3.53 +  "Val  x"     => "(In1 x)"
    3.54 +  "Var  x"     => "(In2 x)"
    3.55 +  "Vals x"     => "(In3 x)"
    3.56 +
    3.57 +syntax
    3.58 +  "Val_"    :: "[pttrn] => pttrn"     ("Val:_"  [951] 950)
    3.59 +  "Var_"    :: "[pttrn] => pttrn"     ("Var:_"  [951] 950)
    3.60 +  "Vals_"   :: "[pttrn] => pttrn"     ("Vals:_" [951] 950)
    3.61 +
    3.62 +translations
    3.63 +  "\<lambda>Val:v . b"  == "(\<lambda>v. b) \<circ> the_In1"
    3.64 +  "\<lambda>Var:v . b"  == "(\<lambda>v. b) \<circ> the_In2"
    3.65 +  "\<lambda>Vals:v. b"  == "(\<lambda>v. b) \<circ> the_In3"
    3.66 +
    3.67 +  (* relation on result values, state and auxiliary variables *)
    3.68 +types 'a assn   =        "res \<Rightarrow> state \<Rightarrow> 'a \<Rightarrow> bool"
    3.69 +translations
    3.70 +      "res"    <= (type) "AxSem.res"
    3.71 +      "a assn" <= (type) "vals \<Rightarrow> state \<Rightarrow> a \<Rightarrow> bool"
    3.72 +
    3.73 +constdefs
    3.74 +  assn_imp   :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> bool"             (infixr "\<Rightarrow>" 25)
    3.75 + "P \<Rightarrow> Q \<equiv> \<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z"
    3.76 +  
    3.77 +lemma assn_imp_def2 [iff]: "(P \<Rightarrow> Q) = (\<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z)"
    3.78 +apply (unfold assn_imp_def)
    3.79 +apply (rule HOL.refl)
    3.80 +done
    3.81 +
    3.82 +
    3.83 +section "assertion transformers"
    3.84 +
    3.85 +subsection "peek-and"
    3.86 +
    3.87 +constdefs
    3.88 +  peek_and   :: "'a assn \<Rightarrow> (state \<Rightarrow>  bool) \<Rightarrow> 'a assn" (infixl "\<and>." 13)
    3.89 + "P \<and>. p \<equiv> \<lambda>Y s Z. P Y s Z \<and> p s"
    3.90 +
    3.91 +lemma peek_and_def2 [simp]: "peek_and P p Y s = (\<lambda>Z. (P Y s Z \<and> p s))"
    3.92 +apply (unfold peek_and_def)
    3.93 +apply (simp (no_asm))
    3.94 +done
    3.95 +
    3.96 +lemma peek_and_Not [simp]: "(P \<and>. (\<lambda>s. \<not> f s)) = (P \<and>. Not \<circ> f)"
    3.97 +apply (rule ext)
    3.98 +apply (rule ext)
    3.99 +apply (simp (no_asm))
   3.100 +done
   3.101 +
   3.102 +lemma peek_and_and [simp]: "peek_and (peek_and P p) p = peek_and P p"
   3.103 +apply (unfold peek_and_def)
   3.104 +apply (simp (no_asm))
   3.105 +done
   3.106 +
   3.107 +lemma peek_and_commut: "(P \<and>. p \<and>. q) = (P \<and>. q \<and>. p)"
   3.108 +apply (rule ext)
   3.109 +apply (rule ext)
   3.110 +apply (rule ext)
   3.111 +apply auto
   3.112 +done
   3.113 +
   3.114 +syntax
   3.115 +  Normal     :: "'a assn \<Rightarrow> 'a assn"
   3.116 +translations
   3.117 +  "Normal P" == "P \<and>. normal"
   3.118 +
   3.119 +lemma peek_and_Normal [simp]: "peek_and (Normal P) p = Normal (peek_and P p)"
   3.120 +apply (rule ext)
   3.121 +apply (rule ext)
   3.122 +apply (rule ext)
   3.123 +apply auto
   3.124 +done
   3.125 +
   3.126 +subsection "assn-supd"
   3.127 +
   3.128 +constdefs
   3.129 +  assn_supd  :: "'a assn \<Rightarrow> (state \<Rightarrow> state) \<Rightarrow> 'a assn" (infixl ";." 13)
   3.130 + "P ;. f \<equiv> \<lambda>Y s' Z. \<exists>s. P Y s Z \<and> s' = f s"
   3.131 +
   3.132 +lemma assn_supd_def2 [simp]: "assn_supd P f Y s' Z = (\<exists>s. P Y s Z \<and> s' = f s)"
   3.133 +apply (unfold assn_supd_def)
   3.134 +apply (simp (no_asm))
   3.135 +done
   3.136 +
   3.137 +subsection "supd-assn"
   3.138 +
   3.139 +constdefs
   3.140 +  supd_assn  :: "(state \<Rightarrow> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr ".;" 13)
   3.141 + "f .; P \<equiv> \<lambda>Y s. P Y (f s)"
   3.142 +
   3.143 +
   3.144 +lemma supd_assn_def2 [simp]: "(f .; P) Y s = P Y (f s)"
   3.145 +apply (unfold supd_assn_def)
   3.146 +apply (simp (no_asm))
   3.147 +done
   3.148 +
   3.149 +lemma supd_assn_supdD [elim]: "((f .; Q) ;. f) Y s Z \<Longrightarrow> Q Y s Z"
   3.150 +apply auto
   3.151 +done
   3.152 +
   3.153 +lemma supd_assn_supdI [elim]: "Q Y s Z \<Longrightarrow> (f .; (Q ;. f)) Y s Z"
   3.154 +apply (auto simp del: split_paired_Ex)
   3.155 +done
   3.156 +
   3.157 +subsection "subst-res"
   3.158 +
   3.159 +constdefs
   3.160 +  subst_res   :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn"              ("_\<leftarrow>_"  [60,61] 60)
   3.161 + "P\<leftarrow>w \<equiv> \<lambda>Y. P w"
   3.162 +
   3.163 +lemma subst_res_def2 [simp]: "(P\<leftarrow>w) Y = P w"
   3.164 +apply (unfold subst_res_def)
   3.165 +apply (simp (no_asm))
   3.166 +done
   3.167 +
   3.168 +lemma subst_subst_res [simp]: "P\<leftarrow>w\<leftarrow>v = P\<leftarrow>w"
   3.169 +apply (rule ext)
   3.170 +apply (simp (no_asm))
   3.171 +done
   3.172 +
   3.173 +lemma peek_and_subst_res [simp]: "(P \<and>. p)\<leftarrow>w = (P\<leftarrow>w \<and>. p)"
   3.174 +apply (rule ext)
   3.175 +apply (rule ext)
   3.176 +apply (simp (no_asm))
   3.177 +done
   3.178 +
   3.179 +(*###Do not work for some strange (unification?) reason
   3.180 +lemma subst_res_Val_beta [simp]: "(\<lambda>Y. P (the_In1 Y))\<leftarrow>Val v = (\<lambda>Y. P v)"
   3.181 +apply (rule ext)
   3.182 +by simp
   3.183 +
   3.184 +lemma subst_res_Var_beta [simp]: "(\<lambda>Y. P (the_In2 Y))\<leftarrow>Var vf = (\<lambda>Y. P vf)";
   3.185 +apply (rule ext)
   3.186 +by simp
   3.187 +
   3.188 +lemma subst_res_Vals_beta [simp]: "(\<lambda>Y. P (the_In3 Y))\<leftarrow>Vals vs = (\<lambda>Y. P vs)";
   3.189 +apply (rule ext)
   3.190 +by simp
   3.191 +*)
   3.192 +
   3.193 +subsection "subst-Bool"
   3.194 +
   3.195 +constdefs
   3.196 +  subst_Bool  :: "'a assn \<Rightarrow> bool \<Rightarrow> 'a assn"             ("_\<leftarrow>=_" [60,61] 60)
   3.197 + "P\<leftarrow>=b \<equiv> \<lambda>Y s Z. \<exists>v. P (Val v) s Z \<and> (normal s \<longrightarrow> the_Bool v=b)"
   3.198 +
   3.199 +lemma subst_Bool_def2 [simp]: 
   3.200 +"(P\<leftarrow>=b) Y s Z = (\<exists>v. P (Val v) s Z \<and> (normal s \<longrightarrow> the_Bool v=b))"
   3.201 +apply (unfold subst_Bool_def)
   3.202 +apply (simp (no_asm))
   3.203 +done
   3.204 +
   3.205 +lemma subst_Bool_the_BoolI: "P (Val b) s Z \<Longrightarrow> (P\<leftarrow>=the_Bool b) Y s Z"
   3.206 +apply auto
   3.207 +done
   3.208 +
   3.209 +subsection "peek-res"
   3.210 +
   3.211 +constdefs
   3.212 +  peek_res    :: "(res \<Rightarrow> 'a assn) \<Rightarrow> 'a assn"
   3.213 + "peek_res Pf \<equiv> \<lambda>Y. Pf Y Y"
   3.214 +
   3.215 +syntax
   3.216 +"@peek_res"  :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"            ("\<lambda>_:. _" [0,3] 3)
   3.217 +translations
   3.218 +  "\<lambda>w:. P"   == "peek_res (\<lambda>w. P)"
   3.219 +
   3.220 +lemma peek_res_def2 [simp]: "peek_res P Y = P Y Y"
   3.221 +apply (unfold peek_res_def)
   3.222 +apply (simp (no_asm))
   3.223 +done
   3.224 +
   3.225 +lemma peek_res_subst_res [simp]: "peek_res P\<leftarrow>w = P w\<leftarrow>w"
   3.226 +apply (rule ext)
   3.227 +apply (simp (no_asm))
   3.228 +done
   3.229 +
   3.230 +(* unused *)
   3.231 +lemma peek_subst_res_allI: 
   3.232 + "(\<And>a. T a (P (f a)\<leftarrow>f a)) \<Longrightarrow> \<forall>a. T a (peek_res P\<leftarrow>f a)"
   3.233 +apply (rule allI)
   3.234 +apply (simp (no_asm))
   3.235 +apply fast
   3.236 +done
   3.237 +
   3.238 +subsection "ign-res"
   3.239 +
   3.240 +constdefs
   3.241 +  ign_res    ::  "        'a assn \<Rightarrow> 'a assn"            ("_\<down>" [1000] 1000)
   3.242 +  "P\<down>        \<equiv> \<lambda>Y s Z. \<exists>Y. P Y s Z"
   3.243 +
   3.244 +lemma ign_res_def2 [simp]: "P\<down> Y s Z = (\<exists>Y. P Y s Z)"
   3.245 +apply (unfold ign_res_def)
   3.246 +apply (simp (no_asm))
   3.247 +done
   3.248 +
   3.249 +lemma ign_ign_res [simp]: "P\<down>\<down> = P\<down>"
   3.250 +apply (rule ext)
   3.251 +apply (rule ext)
   3.252 +apply (rule ext)
   3.253 +apply (simp (no_asm))
   3.254 +done
   3.255 +
   3.256 +lemma ign_subst_res [simp]: "P\<down>\<leftarrow>w = P\<down>"
   3.257 +apply (rule ext)
   3.258 +apply (rule ext)
   3.259 +apply (rule ext)
   3.260 +apply (simp (no_asm))
   3.261 +done
   3.262 +
   3.263 +lemma peek_and_ign_res [simp]: "(P \<and>. p)\<down> = (P\<down> \<and>. p)"
   3.264 +apply (rule ext)
   3.265 +apply (rule ext)
   3.266 +apply (rule ext)
   3.267 +apply (simp (no_asm))
   3.268 +done
   3.269 +
   3.270 +subsection "peek-st"
   3.271 +
   3.272 +constdefs
   3.273 +  peek_st    :: "(st \<Rightarrow> 'a assn) \<Rightarrow> 'a assn"
   3.274 + "peek_st P \<equiv> \<lambda>Y s. P (store s) Y s"
   3.275 +
   3.276 +syntax
   3.277 +"@peek_st"   :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"            ("\<lambda>_.. _" [0,3] 3)
   3.278 +translations
   3.279 +  "\<lambda>s.. P"   == "peek_st (\<lambda>s. P)"
   3.280 +
   3.281 +lemma peek_st_def2 [simp]: "(\<lambda>s.. Pf s) Y s = Pf (store s) Y s"
   3.282 +apply (unfold peek_st_def)
   3.283 +apply (simp (no_asm))
   3.284 +done
   3.285 +
   3.286 +lemma peek_st_triv [simp]: "(\<lambda>s.. P) = P"
   3.287 +apply (rule ext)
   3.288 +apply (rule ext)
   3.289 +apply (simp (no_asm))
   3.290 +done
   3.291 +
   3.292 +lemma peek_st_st [simp]: "(\<lambda>s.. \<lambda>s'.. P s s') = (\<lambda>s.. P s s)"
   3.293 +apply (rule ext)
   3.294 +apply (rule ext)
   3.295 +apply (simp (no_asm))
   3.296 +done
   3.297 +
   3.298 +lemma peek_st_split [simp]: "(\<lambda>s.. \<lambda>Y s'. P s Y s') = (\<lambda>Y s. P (store s) Y s)"
   3.299 +apply (rule ext)
   3.300 +apply (rule ext)
   3.301 +apply (simp (no_asm))
   3.302 +done
   3.303 +
   3.304 +lemma peek_st_subst_res [simp]: "(\<lambda>s.. P s)\<leftarrow>w = (\<lambda>s.. P s\<leftarrow>w)"
   3.305 +apply (rule ext)
   3.306 +apply (simp (no_asm))
   3.307 +done
   3.308 +
   3.309 +lemma peek_st_Normal [simp]: "(\<lambda>s..(Normal (P s))) = Normal (\<lambda>s.. P s)"
   3.310 +apply (rule ext)
   3.311 +apply (rule ext)
   3.312 +apply (simp (no_asm))
   3.313 +done
   3.314 +
   3.315 +subsection "ign-res-eq"
   3.316 +
   3.317 +constdefs
   3.318 +  ign_res_eq :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn"               ("_\<down>=_"  [60,61] 60)
   3.319 + "P\<down>=w       \<equiv> \<lambda>Y:. P\<down> \<and>. (\<lambda>s. Y=w)"
   3.320 +
   3.321 +lemma ign_res_eq_def2 [simp]: "(P\<down>=w) Y s Z = ((\<exists>Y. P Y s Z) \<and> Y=w)"
   3.322 +apply (unfold ign_res_eq_def)
   3.323 +apply auto
   3.324 +done
   3.325 +
   3.326 +lemma ign_ign_res_eq [simp]: "(P\<down>=w)\<down> = P\<down>"
   3.327 +apply (rule ext)
   3.328 +apply (rule ext)
   3.329 +apply (rule ext)
   3.330 +apply (simp (no_asm))
   3.331 +done
   3.332 +
   3.333 +(* unused *)
   3.334 +lemma ign_res_eq_subst_res: "P\<down>=w\<leftarrow>w = P\<down>"
   3.335 +apply (rule ext)
   3.336 +apply (rule ext)
   3.337 +apply (rule ext)
   3.338 +apply (simp (no_asm))
   3.339 +done
   3.340 +
   3.341 +(* unused *)
   3.342 +lemma subst_Bool_ign_res_eq: "((P\<leftarrow>=b)\<down>=x) Y s Z = ((P\<leftarrow>=b) Y s Z  \<and> Y=x)"
   3.343 +apply (simp (no_asm))
   3.344 +done
   3.345 +
   3.346 +subsection "RefVar"
   3.347 +
   3.348 +constdefs
   3.349 +  RefVar    :: "(state \<Rightarrow> vvar \<times> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn"(infixr "..;" 13)
   3.350 + "vf ..; P \<equiv> \<lambda>Y s. let (v,s') = vf s in P (Var v) s'"
   3.351 + 
   3.352 +lemma RefVar_def2 [simp]: "(vf ..; P) Y s =  
   3.353 +  P (Var (fst (vf s))) (snd (vf s))"
   3.354 +apply (unfold RefVar_def Let_def)
   3.355 +apply (simp (no_asm) add: split_beta)
   3.356 +done
   3.357 +
   3.358 +subsection "allocation"
   3.359 +
   3.360 +constdefs
   3.361 +  Alloc      :: "prog \<Rightarrow> obj_tag \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
   3.362 + "Alloc G otag P \<equiv> \<lambda>Y s Z.
   3.363 +                   \<forall>s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a\<rightarrow> s'\<longrightarrow> P (Val (Addr a)) s' Z"
   3.364 +
   3.365 +  SXAlloc     :: "prog \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
   3.366 + "SXAlloc G P \<equiv> \<lambda>Y s Z. \<forall>s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<longrightarrow> P Y s' Z"
   3.367 +
   3.368 +
   3.369 +lemma Alloc_def2 [simp]: "Alloc G otag P Y s Z =  
   3.370 +       (\<forall>s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a\<rightarrow> s'\<longrightarrow> P (Val (Addr a)) s' Z)"
   3.371 +apply (unfold Alloc_def)
   3.372 +apply (simp (no_asm))
   3.373 +done
   3.374 +
   3.375 +lemma SXAlloc_def2 [simp]: 
   3.376 +  "SXAlloc G P Y s Z = (\<forall>s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<longrightarrow> P Y s' Z)"
   3.377 +apply (unfold SXAlloc_def)
   3.378 +apply (simp (no_asm))
   3.379 +done
   3.380 +
   3.381 +section "validity"
   3.382 +
   3.383 +constdefs
   3.384 +  type_ok  :: "prog \<Rightarrow> term \<Rightarrow> state \<Rightarrow> bool"
   3.385 + "type_ok G t s \<equiv> \<exists>L T C. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) \<and> s\<Colon>\<preceq>(G,L)"
   3.386 +
   3.387 +datatype    'a triple = triple "('a assn)" "term" "('a assn)" (** should be
   3.388 +something like triple = \<forall>'a. triple ('a assn) term ('a assn)   **)
   3.389 +                                        ("{(1_)}/ _>/ {(1_)}"      [3,65,3]75)
   3.390 +types    'a triples = "'a triple set"
   3.391 +
   3.392 +syntax
   3.393 +
   3.394 +  var_triple   :: "['a assn, var         ,'a assn] \<Rightarrow> 'a triple"
   3.395 +                                         ("{(1_)}/ _=>/ {(1_)}"    [3,80,3] 75)
   3.396 +  expr_triple  :: "['a assn, expr        ,'a assn] \<Rightarrow> 'a triple"
   3.397 +                                         ("{(1_)}/ _->/ {(1_)}"    [3,80,3] 75)
   3.398 +  exprs_triple :: "['a assn, expr list   ,'a assn] \<Rightarrow> 'a triple"
   3.399 +                                         ("{(1_)}/ _#>/ {(1_)}"    [3,65,3] 75)
   3.400 +  stmt_triple  :: "['a assn, stmt,        'a assn] \<Rightarrow> 'a triple"
   3.401 +                                         ("{(1_)}/ ._./ {(1_)}"     [3,65,3] 75)
   3.402 +
   3.403 +syntax (xsymbols)
   3.404 +
   3.405 +  triple       :: "['a assn, term        ,'a assn] \<Rightarrow> 'a triple"
   3.406 +                                         ("{(1_)}/ _\<succ>/ {(1_)}"     [3,65,3] 75)
   3.407 +  var_triple   :: "['a assn, var         ,'a assn] \<Rightarrow> 'a triple"
   3.408 +                                         ("{(1_)}/ _=\<succ>/ {(1_)}"    [3,80,3] 75)
   3.409 +  expr_triple  :: "['a assn, expr        ,'a assn] \<Rightarrow> 'a triple"
   3.410 +                                         ("{(1_)}/ _-\<succ>/ {(1_)}"    [3,80,3] 75)
   3.411 +  exprs_triple :: "['a assn, expr list   ,'a assn] \<Rightarrow> 'a triple"
   3.412 +                                         ("{(1_)}/ _\<doteq>\<succ>/ {(1_)}"    [3,65,3] 75)
   3.413 +
   3.414 +translations
   3.415 +  "{P} e-\<succ> {Q}" == "{P} In1l e\<succ> {Q}"
   3.416 +  "{P} e=\<succ> {Q}" == "{P} In2  e\<succ> {Q}"
   3.417 +  "{P} e\<doteq>\<succ> {Q}" == "{P} In3  e\<succ> {Q}"
   3.418 +  "{P} .c. {Q}" == "{P} In1r c\<succ> {Q}"
   3.419 +
   3.420 +lemma inj_triple: "inj (\<lambda>(P,t,Q). {P} t\<succ> {Q})"
   3.421 +apply (rule injI)
   3.422 +apply auto
   3.423 +done
   3.424 +
   3.425 +lemma triple_inj_eq: "({P} t\<succ> {Q} = {P'} t'\<succ> {Q'} ) = (P=P' \<and> t=t' \<and> Q=Q')"
   3.426 +apply auto
   3.427 +done
   3.428 +
   3.429 +constdefs
   3.430 +  mtriples  :: "('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<Rightarrow> 'sig \<Rightarrow> expr) \<Rightarrow> 
   3.431 +                ('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<times>  'sig) set \<Rightarrow> 'a triples"
   3.432 +                                     ("{{(1_)}/ _-\<succ>/ {(1_)} | _}"[3,65,3,65]75)
   3.433 + "{{P} tf-\<succ> {Q} | ms} \<equiv> (\<lambda>(C,sig). {Normal(P C sig)} tf C sig-\<succ> {Q C sig})`ms"
   3.434 +  
   3.435 +consts
   3.436 +
   3.437 + triple_valid :: "prog \<Rightarrow> nat \<Rightarrow>        'a triple  \<Rightarrow> bool"
   3.438 +                                                (   "_\<Turnstile>_:_" [61,0, 58] 57)
   3.439 +    ax_valids :: "prog \<Rightarrow> 'b triples \<Rightarrow> 'a triples \<Rightarrow> bool"
   3.440 +                                                ("_,_|\<Turnstile>_"   [61,58,58] 57)
   3.441 +    ax_derivs :: "prog \<Rightarrow> ('b triples \<times> 'a triples) set"
   3.442 +
   3.443 +syntax
   3.444 +
   3.445 + triples_valid:: "prog \<Rightarrow> nat \<Rightarrow>         'a triples \<Rightarrow> bool"
   3.446 +                                                (  "_||=_:_" [61,0, 58] 57)
   3.447 +     ax_valid :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
   3.448 +                                                ( "_,_|=_"   [61,58,58] 57)
   3.449 +     ax_Derivs:: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triples \<Rightarrow> bool"
   3.450 +                                                ("_,_||-_"   [61,58,58] 57)
   3.451 +     ax_Deriv :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
   3.452 +                                                ( "_,_|-_"   [61,58,58] 57)
   3.453 +
   3.454 +syntax (xsymbols)
   3.455 +
   3.456 + triples_valid:: "prog \<Rightarrow> nat \<Rightarrow>         'a triples \<Rightarrow> bool"
   3.457 +                                                (  "_|\<Turnstile>_:_" [61,0, 58] 57)
   3.458 +     ax_valid :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
   3.459 +                                                ( "_,_\<Turnstile>_"   [61,58,58] 57)
   3.460 +     ax_Derivs:: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triples \<Rightarrow> bool"
   3.461 +                                                ("_,_|\<turnstile>_"   [61,58,58] 57)
   3.462 +     ax_Deriv :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
   3.463 +                                                ( "_,_\<turnstile>_"   [61,58,58] 57)
   3.464 +
   3.465 +defs  triple_valid_def:  "G\<Turnstile>n:t  \<equiv> case t of {P} t\<succ> {Q} \<Rightarrow>
   3.466 +                          \<forall>Y s Z. P Y s Z \<longrightarrow> type_ok G t s \<longrightarrow>
   3.467 +                          (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> Q Y' s' Z)"
   3.468 +translations         "G|\<Turnstile>n:ts" == "Ball ts (triple_valid G n)"
   3.469 +defs   ax_valids_def:"G,A|\<Turnstile>ts  \<equiv>  \<forall>n. G|\<Turnstile>n:A \<longrightarrow> G|\<Turnstile>n:ts"
   3.470 +translations         "G,A \<Turnstile>t"  == "G,A|\<Turnstile>{t}"
   3.471 +                     "G,A|\<turnstile>ts" == "(A,ts) \<in> ax_derivs G"
   3.472 +                     "G,A \<turnstile>t"  == "G,A|\<turnstile>{t}"
   3.473 +
   3.474 +lemma triple_valid_def2: "G\<Turnstile>n:{P} t\<succ> {Q} =  
   3.475 + (\<forall>Y s Z. P Y s Z 
   3.476 +  \<longrightarrow> (\<exists>L. (normal s \<longrightarrow> (\<exists>T C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T)) \<and> s\<Colon>\<preceq>(G,L)) \<longrightarrow> 
   3.477 +  (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s')\<longrightarrow> Q Y' s' Z))"
   3.478 +apply (unfold triple_valid_def type_ok_def)
   3.479 +apply (simp (no_asm))
   3.480 +done
   3.481 +
   3.482 +
   3.483 +declare split_paired_All [simp del] split_paired_Ex [simp del] 
   3.484 +declare split_if     [split del] split_if_asm     [split del] 
   3.485 +        option.split [split del] option.split_asm [split del]
   3.486 +ML_setup {*
   3.487 +simpset_ref() := simpset() delloop "split_all_tac";
   3.488 +claset_ref () := claset () delSWrapper "split_all_tac"
   3.489 +*}
   3.490 +
   3.491 +
   3.492 +inductive "ax_derivs G" intros
   3.493 +
   3.494 +  empty: " G,A|\<turnstile>{}"
   3.495 +  insert:"\<lbrakk>G,A\<turnstile>t; G,A|\<turnstile>ts\<rbrakk> \<Longrightarrow>
   3.496 +          G,A|\<turnstile>insert t ts"
   3.497 +
   3.498 +  asm:   "ts\<subseteq>A \<Longrightarrow> G,A|\<turnstile>ts"
   3.499 +
   3.500 +(* could be added for convenience and efficiency, but is not necessary
   3.501 +  cut:   "\<lbrakk>G,A'|\<turnstile>ts; G,A|\<turnstile>A'\<rbrakk> \<Longrightarrow>
   3.502 +           G,A |\<turnstile>ts"
   3.503 +*)
   3.504 +  weaken:"\<lbrakk>G,A|\<turnstile>ts'; ts \<subseteq> ts'\<rbrakk> \<Longrightarrow> G,A|\<turnstile>ts"
   3.505 +
   3.506 +  conseq:"\<forall>Y s Z . P  Y s Z  \<longrightarrow> (\<exists>P' Q'. G,A\<turnstile>{P'} t\<succ> {Q'} \<and> (\<forall>Y' s'. 
   3.507 +         (\<forall>Y   Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>
   3.508 +                                 Q  Y' s' Z ))
   3.509 +                                         \<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q }"
   3.510 +
   3.511 +  hazard:"G,A\<turnstile>{P \<and>. Not \<circ> type_ok G t} t\<succ> {Q}"
   3.512 +
   3.513 +  Abrupt:  "G,A\<turnstile>{P\<leftarrow>(arbitrary3 t) \<and>. Not \<circ> normal} t\<succ> {P}"
   3.514 +
   3.515 +  (* variables *)
   3.516 +  LVar:  " G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Var (lvar vn s))} LVar vn=\<succ> {P}"
   3.517 +
   3.518 +  FVar: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Q};
   3.519 +          G,A\<turnstile>{Q} e-\<succ> {\<lambda>Val:a:. fvar C stat fn a ..; R}\<rbrakk> \<Longrightarrow>
   3.520 +                                 G,A\<turnstile>{Normal P} {C,stat}e..fn=\<succ> {R}"
   3.521 +
   3.522 +  AVar:  "\<lbrakk>G,A\<turnstile>{Normal P} e1-\<succ> {Q};
   3.523 +          \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} e2-\<succ> {\<lambda>Val:i:. avar G i a ..; R}\<rbrakk> \<Longrightarrow>
   3.524 +                                 G,A\<turnstile>{Normal P} e1.[e2]=\<succ> {R}"
   3.525 +  (* expressions *)
   3.526 +
   3.527 +  NewC: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Alloc G (CInst C) Q}\<rbrakk> \<Longrightarrow>
   3.528 +                                 G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}"
   3.529 +
   3.530 +  NewA: "\<lbrakk>G,A\<turnstile>{Normal P} .init_comp_ty T. {Q};  G,A\<turnstile>{Q} e-\<succ>
   3.531 +	  {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\<rbrakk> \<Longrightarrow>
   3.532 +                                 G,A\<turnstile>{Normal P} New T[e]-\<succ> {R}"
   3.533 +
   3.534 +  Cast: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
   3.535 +          abupd (raise_if (\<not>G,s\<turnstile>v fits T) ClassCast) .; Q\<leftarrow>Val v}\<rbrakk> \<Longrightarrow>
   3.536 +                                 G,A\<turnstile>{Normal P} Cast T e-\<succ> {Q}"
   3.537 +
   3.538 +  Inst: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
   3.539 +                  Q\<leftarrow>Val (Bool (v\<noteq>Null \<and> G,s\<turnstile>v fits RefT T))}\<rbrakk> \<Longrightarrow>
   3.540 +                                 G,A\<turnstile>{Normal P} e InstOf T-\<succ> {Q}"
   3.541 +
   3.542 +  Lit:                          "G,A\<turnstile>{Normal (P\<leftarrow>Val v)} Lit v-\<succ> {P}"
   3.543 +
   3.544 +  Super:" G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Val (val_this s))} Super-\<succ> {P}"
   3.545 +
   3.546 +  Acc:  "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {\<lambda>Var:(v,f):. Q\<leftarrow>Val v}\<rbrakk> \<Longrightarrow>
   3.547 +                                 G,A\<turnstile>{Normal P} Acc va-\<succ> {Q}"
   3.548 +
   3.549 +  Ass:  "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {Q};
   3.550 +     \<forall>vf. G,A\<turnstile>{Q\<leftarrow>Var vf} e-\<succ> {\<lambda>Val:v:. assign (snd vf) v .; R}\<rbrakk> \<Longrightarrow>
   3.551 +                                 G,A\<turnstile>{Normal P} va:=e-\<succ> {R}"
   3.552 +
   3.553 +  Cond: "\<lbrakk>G,A \<turnstile>{Normal P} e0-\<succ> {P'};
   3.554 +          \<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} (if b then e1 else e2)-\<succ> {Q}\<rbrakk> \<Longrightarrow>
   3.555 +                                 G,A\<turnstile>{Normal P} e0 ? e1 : e2-\<succ> {Q}"
   3.556 +
   3.557 +  Call: 
   3.558 +"\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q}; \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {R a};
   3.559 +  \<forall>a vs invC declC l. G,A\<turnstile>{(R a\<leftarrow>Vals vs \<and>.
   3.560 + (\<lambda>s. declC=invocation_declclass G mode (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> \<and>
   3.561 +      invC = invocation_class mode (store s) a statT \<and>
   3.562 +         l = locals (store s)) ;.
   3.563 +      init_lvars G declC \<lparr>name=mn,parTs=pTs\<rparr> mode a vs) \<and>.
   3.564 +      (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)}
   3.565 + Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}\<rbrakk> \<Longrightarrow>
   3.566 +         G,A\<turnstile>{Normal P} {statT,mode}e\<cdot>mn({pTs}args)-\<succ> {S}"
   3.567 +
   3.568 +  Methd:"\<lbrakk>G,A\<union> {{P} Methd-\<succ> {Q} | ms} |\<turnstile> {{P} body G-\<succ> {Q} | ms}\<rbrakk> \<Longrightarrow>
   3.569 +                                 G,A|\<turnstile>{{P} Methd-\<succ>  {Q} | ms}"
   3.570 +
   3.571 +  Body: "\<lbrakk>G,A\<turnstile>{Normal P} .Init D. {Q}; 
   3.572 +  G,A\<turnstile>{Q} .c. {\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>(In1 (the (locals s Result)))}\<rbrakk> 
   3.573 +    \<Longrightarrow>
   3.574 +                                 G,A\<turnstile>{Normal P} Body D c-\<succ> {R}"
   3.575 +  
   3.576 +  (* expression lists *)
   3.577 +
   3.578 +  Nil:                          "G,A\<turnstile>{Normal (P\<leftarrow>Vals [])} []\<doteq>\<succ> {P}"
   3.579 +
   3.580 +  Cons: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q};
   3.581 +          \<forall>v. G,A\<turnstile>{Q\<leftarrow>Val v} es\<doteq>\<succ> {\<lambda>Vals:vs:. R\<leftarrow>Vals (v#vs)}\<rbrakk> \<Longrightarrow>
   3.582 +                                 G,A\<turnstile>{Normal P} e#es\<doteq>\<succ> {R}"
   3.583 +
   3.584 +  (* statements *)
   3.585 +
   3.586 +  Skip:                         "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit>)} .Skip. {P}"
   3.587 +
   3.588 +  Expr: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>
   3.589 +                                 G,A\<turnstile>{Normal P} .Expr e. {Q}"
   3.590 +
   3.591 +  Lab: "\<lbrakk>G,A\<turnstile>{Normal P} .c. {abupd (absorb (Break l)) .; Q}\<rbrakk> \<Longrightarrow>
   3.592 +                           G,A\<turnstile>{Normal P} .l\<bullet> c. {Q}"
   3.593 +
   3.594 +  Comp: "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q};
   3.595 +          G,A\<turnstile>{Q} .c2. {R}\<rbrakk> \<Longrightarrow>
   3.596 +                                 G,A\<turnstile>{Normal P} .c1;;c2. {R}"
   3.597 +
   3.598 +  If:   "\<lbrakk>G,A \<turnstile>{Normal P} e-\<succ> {P'};
   3.599 +          \<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c1 else c2). {Q}\<rbrakk> \<Longrightarrow>
   3.600 +                                 G,A\<turnstile>{Normal P} .If(e) c1 Else c2. {Q}"
   3.601 +(* unfolding variant of Loop, not needed here
   3.602 +  LoopU:"\<lbrakk>G,A \<turnstile>{Normal P} e-\<succ> {P'};
   3.603 +          \<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c;;While(e) c else Skip).{Q}\<rbrakk>
   3.604 +         \<Longrightarrow>              G,A\<turnstile>{Normal P} .While(e) c. {Q}"
   3.605 +*)
   3.606 +  Loop: "\<lbrakk>G,A\<turnstile>{P} e-\<succ> {P'}; 
   3.607 +          G,A\<turnstile>{Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}\<rbrakk> \<Longrightarrow>
   3.608 +                            G,A\<turnstile>{P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}"
   3.609 +(** Beware of polymorphic_Loop below: should be identical terms **)
   3.610 +  
   3.611 +  Do: "G,A\<turnstile>{Normal (abupd (\<lambda>a. (Some (Jump j))) .; P\<leftarrow>\<diamondsuit>)} .Do j. {P}"
   3.612 +
   3.613 +  Throw:"\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>
   3.614 +                                 G,A\<turnstile>{Normal P} .Throw e. {Q}"
   3.615 +
   3.616 +  Try:  "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {SXAlloc G Q};
   3.617 +          G,A\<turnstile>{Q \<and>. (\<lambda>s.  G,s\<turnstile>catch C) ;. new_xcpt_var vn} .c2. {R};
   3.618 +              (Q \<and>. (\<lambda>s. \<not>G,s\<turnstile>catch C)) \<Rightarrow> R\<rbrakk> \<Longrightarrow>
   3.619 +                                 G,A\<turnstile>{Normal P} .Try c1 Catch(C vn) c2. {R}"
   3.620 +
   3.621 +  Fin:  "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q};
   3.622 +      \<forall>x. G,A\<turnstile>{Q \<and>. (\<lambda>s. x = fst s) ;. abupd (\<lambda>x. None)}
   3.623 +              .c2. {abupd (abrupt_if (x\<noteq>None) x) .; R}\<rbrakk> \<Longrightarrow>
   3.624 +                                 G,A\<turnstile>{Normal P} .c1 Finally c2. {R}"
   3.625 +
   3.626 +  Done:                       "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit> \<and>. initd C)} .Init C. {P}"
   3.627 +
   3.628 +  Init: "\<lbrakk>the (class G C) = c;
   3.629 +          G,A\<turnstile>{Normal ((P \<and>. Not \<circ> initd C) ;. supd (init_class_obj G C))}
   3.630 +              .(if C = Object then Skip else Init (super c)). {Q};
   3.631 +      \<forall>l. G,A\<turnstile>{Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars empty}
   3.632 +              .init c. {set_lvars l .; R}\<rbrakk> \<Longrightarrow>
   3.633 +                               G,A\<turnstile>{Normal (P \<and>. Not \<circ> initd C)} .Init C. {R}"
   3.634 +
   3.635 +axioms (** these terms are the same as above, but with generalized typing **)
   3.636 +  polymorphic_conseq:
   3.637 +        "\<forall>Y s Z . P  Y s Z  \<longrightarrow> (\<exists>P' Q'. G,A\<turnstile>{P'} t\<succ> {Q'} \<and> (\<forall>Y' s'. 
   3.638 +        (\<forall>Y   Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>
   3.639 +                                Q  Y' s' Z ))
   3.640 +                                         \<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q }"
   3.641 +
   3.642 +  polymorphic_Loop:
   3.643 +        "\<lbrakk>G,A\<turnstile>{P} e-\<succ> {P'}; 
   3.644 +          G,A\<turnstile>{Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}\<rbrakk> \<Longrightarrow>
   3.645 +                            G,A\<turnstile>{P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}"
   3.646 +
   3.647 +constdefs
   3.648 + adapt_pre :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
   3.649 +"adapt_pre P Q Q'\<equiv>\<lambda>Y s Z. \<forall>Y' s'. \<exists>Z'. P Y s Z' \<and> (Q Y' s' Z' \<longrightarrow> Q' Y' s' Z)"
   3.650 +
   3.651 +
   3.652 +section "rules derived by induction"
   3.653 +
   3.654 +lemma cut_valid: "\<lbrakk>G,A'|\<Turnstile>ts; G,A|\<Turnstile>A'\<rbrakk> \<Longrightarrow> G,A|\<Turnstile>ts"
   3.655 +apply (unfold ax_valids_def)
   3.656 +apply fast
   3.657 +done
   3.658 +
   3.659 +(*if cut is available
   3.660 +Goal "\<lbrakk>G,A'|\<turnstile>ts; A' \<subseteq> A; \<forall>P Q t. {P} t\<succ> {Q} \<in> A' \<longrightarrow> (\<exists>T. (G,L)\<turnstile>t\<Colon>T) \<rbrakk> \<Longrightarrow>  
   3.661 +       G,A|\<turnstile>ts"
   3.662 +b y etac ax_derivs.cut 1;
   3.663 +b y eatac ax_derivs.asm 1 1;
   3.664 +qed "ax_thin";
   3.665 +*)
   3.666 +lemma ax_thin [rule_format (no_asm)]: 
   3.667 +  "G,(A'::'a triple set)|\<turnstile>(ts::'a triple set) \<Longrightarrow> \<forall>A. A' \<subseteq> A \<longrightarrow> G,A|\<turnstile>ts"
   3.668 +apply (erule ax_derivs.induct)
   3.669 +apply                (tactic "ALLGOALS(EVERY'[Clarify_tac,REPEAT o smp_tac 1])")
   3.670 +apply                (rule ax_derivs.empty)
   3.671 +apply               (erule (1) ax_derivs.insert)
   3.672 +apply              (fast intro: ax_derivs.asm)
   3.673 +(*apply           (fast intro: ax_derivs.cut) *)
   3.674 +apply            (fast intro: ax_derivs.weaken)
   3.675 +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)
   3.676 +(* 31 subgoals *)
   3.677 +prefer 16 (* Methd *)
   3.678 +apply (rule ax_derivs.Methd, drule spec, erule mp, fast)
   3.679 +apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) (thms "ax_derivs.intros")) 
   3.680 +                     THEN_ALL_NEW Blast_tac) *})
   3.681 +apply (erule ax_derivs.Call)
   3.682 +apply   clarify 
   3.683 +apply   blast
   3.684 +
   3.685 +apply   (rule allI)+ 
   3.686 +apply   (drule spec)+
   3.687 +apply   blast
   3.688 +done
   3.689 +
   3.690 +lemma ax_thin_insert: "G,(A::'a triple set)\<turnstile>(t::'a triple) \<Longrightarrow> G,insert x A\<turnstile>t"
   3.691 +apply (erule ax_thin)
   3.692 +apply fast
   3.693 +done
   3.694 +
   3.695 +lemma subset_mtriples_iff: 
   3.696 +  "ts \<subseteq> {{P} mb-\<succ> {Q} | ms} = (\<exists>ms'. ms'\<subseteq>ms \<and>  ts = {{P} mb-\<succ> {Q} | ms'})"
   3.697 +apply (unfold mtriples_def)
   3.698 +apply (rule subset_image_iff)
   3.699 +done
   3.700 +
   3.701 +lemma weaken: 
   3.702 + "G,(A::'a triple set)|\<turnstile>(ts'::'a triple set) \<Longrightarrow> !ts. ts \<subseteq> ts' \<longrightarrow> G,A|\<turnstile>ts"
   3.703 +apply (erule ax_derivs.induct)
   3.704 +(*36 subgoals*)
   3.705 +apply       (tactic "ALLGOALS strip_tac")
   3.706 +apply       (tactic {* ALLGOALS(REPEAT o (EVERY'[dtac (thm "subset_singletonD"),
   3.707 +         etac disjE, fast_tac (claset() addSIs [thm "ax_derivs.empty"])]))*})
   3.708 +apply       (tactic "TRYALL hyp_subst_tac")
   3.709 +apply       (simp, rule ax_derivs.empty)
   3.710 +apply      (drule subset_insertD)
   3.711 +apply      (blast intro: ax_derivs.insert)
   3.712 +apply     (fast intro: ax_derivs.asm)
   3.713 +(*apply  (blast intro: ax_derivs.cut) *)
   3.714 +apply   (fast intro: ax_derivs.weaken)
   3.715 +apply  (rule ax_derivs.conseq, clarify, tactic "smp_tac 3 1", blast(* unused *))
   3.716 +(*31 subgoals*)
   3.717 +apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) (thms "ax_derivs.intros")) 
   3.718 +                   THEN_ALL_NEW Fast_tac) *})
   3.719 +(*1 subgoal*)
   3.720 +apply (clarsimp simp add: subset_mtriples_iff)
   3.721 +apply (rule ax_derivs.Methd)
   3.722 +apply (drule spec)
   3.723 +apply (erule impE)
   3.724 +apply  (rule exI)
   3.725 +apply  (erule conjI)
   3.726 +apply  (rule HOL.refl)
   3.727 +oops (* dead end, Methd is to blame *)
   3.728 +
   3.729 +
   3.730 +section "rules derived from conseq"
   3.731 +
   3.732 +lemma conseq12: "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q'};  
   3.733 + \<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>  
   3.734 +  Q Y' s' Z)\<rbrakk>  
   3.735 +  \<Longrightarrow>  G,A\<turnstile>{P ::'a assn} t\<succ> {Q }"
   3.736 +apply (rule polymorphic_conseq)
   3.737 +apply clarsimp
   3.738 +apply blast
   3.739 +done
   3.740 +
   3.741 +(*unused, but nice variant*)
   3.742 +lemma conseq12': "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q'}; \<forall>s Y' s'.  
   3.743 +       (\<forall>Y Z. P' Y s Z \<longrightarrow> Q' Y' s' Z) \<longrightarrow>  
   3.744 +       (\<forall>Y Z. P  Y s Z \<longrightarrow> Q  Y' s' Z)\<rbrakk>  
   3.745 +  \<Longrightarrow>  G,A\<turnstile>{P } t\<succ> {Q }"
   3.746 +apply (erule conseq12)
   3.747 +apply fast
   3.748 +done
   3.749 +
   3.750 +lemma conseq12_from_conseq12': "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q'};  
   3.751 + \<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>  
   3.752 +  Q Y' s' Z)\<rbrakk>  
   3.753 +  \<Longrightarrow>  G,A\<turnstile>{P } t\<succ> {Q }"
   3.754 +apply (erule conseq12')
   3.755 +apply blast
   3.756 +done
   3.757 +
   3.758 +lemma conseq1: "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q}; P \<Rightarrow> P'\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q}"
   3.759 +apply (erule conseq12)
   3.760 +apply blast
   3.761 +done
   3.762 +
   3.763 +lemma conseq2: "\<lbrakk>G,A\<turnstile>{P} t\<succ> {Q'}; Q' \<Rightarrow> Q\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
   3.764 +apply (erule conseq12)
   3.765 +apply blast
   3.766 +done
   3.767 +
   3.768 +lemma ax_escape: "\<lbrakk>\<forall>Y s Z. P Y s Z \<longrightarrow> G,A\<turnstile>{\<lambda>Y' s' Z'. (Y',s') = (Y,s)} t\<succ> {\<lambda>Y s Z'. Q Y s Z}\<rbrakk> \<Longrightarrow>  
   3.769 +  G,A\<turnstile>{P} t\<succ> {Q}"
   3.770 +apply (rule polymorphic_conseq)
   3.771 +apply force
   3.772 +done
   3.773 +
   3.774 +(* unused *)
   3.775 +lemma ax_constant: "\<lbrakk> C \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}\<rbrakk> \<Longrightarrow> G,A\<turnstile>{\<lambda>Y s Z. C \<and> P Y s Z} t\<succ> {Q}"
   3.776 +apply (rule ax_escape (* unused *))
   3.777 +apply clarify
   3.778 +apply (rule conseq12)
   3.779 +apply  fast
   3.780 +apply auto
   3.781 +done
   3.782 +(*alternative (more direct) proof:
   3.783 +apply (rule ax_derivs.conseq) *)(* unused *)(*
   3.784 +apply (fast)
   3.785 +*)
   3.786 +
   3.787 +
   3.788 +lemma ax_impossible [intro]: "G,A\<turnstile>{\<lambda>Y s Z. False} t\<succ> {Q}"
   3.789 +apply (rule ax_escape)
   3.790 +apply clarify
   3.791 +done
   3.792 +
   3.793 +(* unused *)
   3.794 +lemma ax_nochange_lemma: "\<lbrakk>P Y s; All (op = w)\<rbrakk> \<Longrightarrow> P w s"
   3.795 +apply auto
   3.796 +done
   3.797 +lemma ax_nochange:"G,A\<turnstile>{\<lambda>Y s Z. (Y,s)=Z} t\<succ> {\<lambda>Y s Z. (Y,s)=Z} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {P}"
   3.798 +apply (erule conseq12)
   3.799 +apply auto
   3.800 +apply (erule (1) ax_nochange_lemma)
   3.801 +done
   3.802 +
   3.803 +(* unused *)
   3.804 +lemma ax_trivial: "G,A\<turnstile>{P}  t\<succ> {\<lambda>Y s Z. True}"
   3.805 +apply (rule polymorphic_conseq(* unused *))
   3.806 +apply auto
   3.807 +done
   3.808 +
   3.809 +(* unused *)
   3.810 +lemma ax_disj: "\<lbrakk>G,A\<turnstile>{P1} t\<succ> {Q1}; G,A\<turnstile>{P2} t\<succ> {Q2}\<rbrakk> \<Longrightarrow>  
   3.811 +  G,A\<turnstile>{\<lambda>Y s Z. P1 Y s Z \<or> P2 Y s Z} t\<succ> {\<lambda>Y s Z. Q1 Y s Z \<or> Q2 Y s Z}"
   3.812 +apply (rule ax_escape (* unused *))
   3.813 +apply safe
   3.814 +apply  (erule conseq12, fast)+
   3.815 +done
   3.816 +
   3.817 +(* unused *)
   3.818 +lemma ax_supd_shuffle: "(\<exists>Q. G,A\<turnstile>{P} .c1. {Q} \<and> G,A\<turnstile>{Q ;. f} .c2. {R}) =  
   3.819 +       (\<exists>Q'. G,A\<turnstile>{P} .c1. {f .; Q'} \<and> G,A\<turnstile>{Q'} .c2. {R})"
   3.820 +apply (best elim!: conseq1 conseq2)
   3.821 +done
   3.822 +
   3.823 +lemma ax_cases: "\<lbrakk>G,A\<turnstile>{P \<and>.       C} t\<succ> {Q};  
   3.824 +                       G,A\<turnstile>{P \<and>. Not \<circ> C} t\<succ> {Q}\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
   3.825 +apply (unfold peek_and_def)
   3.826 +apply (rule ax_escape)
   3.827 +apply clarify
   3.828 +apply (case_tac "C s")
   3.829 +apply  (erule conseq12, force)+
   3.830 +done
   3.831 +(*alternative (more direct) proof:
   3.832 +apply (rule rtac ax_derivs.conseq) *)(* unused *)(*
   3.833 +apply clarify
   3.834 +apply (case_tac "C s")
   3.835 +apply  force+
   3.836 +*)
   3.837 +
   3.838 +lemma ax_adapt: "G,A\<turnstile>{P} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}"
   3.839 +apply (unfold adapt_pre_def)
   3.840 +apply (erule conseq12)
   3.841 +apply fast
   3.842 +done
   3.843 +
   3.844 +lemma adapt_pre_adapts: "G,A\<Turnstile>{P} t\<succ> {Q} \<longrightarrow> G,A\<Turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}"
   3.845 +apply (unfold adapt_pre_def)
   3.846 +apply (simp add: ax_valids_def triple_valid_def2)
   3.847 +apply fast
   3.848 +done
   3.849 +
   3.850 +
   3.851 +lemma adapt_pre_weakest: 
   3.852 +"\<forall>G (A::'a triple set) t. G,A\<Turnstile>{P} t\<succ> {Q} \<longrightarrow> G,A\<Turnstile>{P'} t\<succ> {Q'} \<Longrightarrow>  
   3.853 +  P' \<Rightarrow> adapt_pre P Q (Q'::'a assn)"
   3.854 +apply (unfold adapt_pre_def)
   3.855 +apply (drule spec)
   3.856 +apply (drule_tac x = "{}" in spec)
   3.857 +apply (drule_tac x = "In1r Skip" in spec)
   3.858 +apply (simp add: ax_valids_def triple_valid_def2)
   3.859 +oops
   3.860 +
   3.861 +(*
   3.862 +Goal "\<forall>(A::'a triple set) t. G,A\<Turnstile>{P} t\<succ> {Q} \<longrightarrow> G,A\<Turnstile>{P'} t\<succ> {Q'} \<Longrightarrow>  
   3.863 +  wf_prog G \<Longrightarrow> G,(A::'a triple set)\<turnstile>{P} t\<succ> {Q::'a assn} \<Longrightarrow> G,A\<turnstile>{P'} t\<succ> {Q'::'a assn}"
   3.864 +b y fatac ax_sound 1 1;
   3.865 +b y asm_full_simp_tac (simpset() addsimps [ax_valids_def,triple_valid_def2]) 1;
   3.866 +b y rtac ax_no_hazard 1; 
   3.867 +b y etac conseq12 1;
   3.868 +b y Clarify_tac 1;
   3.869 +b y case_tac "\<forall>Z. \<not>P Y s Z" 1;
   3.870 +b y smp_tac 2 1;
   3.871 +b y etac thin_rl 1;
   3.872 +b y etac thin_rl 1;
   3.873 +b y clarsimp_tac (claset(), simpset() addsimps [type_ok_def]) 1;
   3.874 +b y subgoal_tac "G|\<Turnstile>n:A" 1;
   3.875 +b y smp_tac 1 1;
   3.876 +b y smp_tac 3 1;
   3.877 +b y etac impE 1;
   3.878 + back();
   3.879 + b y Fast_tac 1;
   3.880 +b y 
   3.881 +b y rotate_tac 2 1;
   3.882 +b y etac thin_rl 1;
   3.883 +b y  etac thin_rl 2;
   3.884 +b y  etac thin_rl 2;
   3.885 +b y  Clarify_tac 2;
   3.886 +b y  dtac spec 2;
   3.887 +b y  EVERY'[dtac spec, mp_tac] 2;
   3.888 +b y  thin_tac "\<forall>n Y s Z. ?PP n Y s Z" 2;
   3.889 +b y  thin_tac "P' Y s Z" 2;
   3.890 +b y  Blast_tac 2;
   3.891 +b y smp_tac 3 1;
   3.892 +b y case_tac "\<forall>Z. \<not>P Y s Z" 1;
   3.893 +b y dres_inst_tac [("x","In1r Skip")] spec 1;
   3.894 +b y Full_simp_tac 1;
   3.895 +*)
   3.896 +
   3.897 +lemma peek_and_forget1_Normal: 
   3.898 + "G,A\<turnstile>{Normal P} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{Normal (P \<and>. p)} t\<succ> {Q}"
   3.899 +apply (erule conseq1)
   3.900 +apply (simp (no_asm))
   3.901 +done
   3.902 +
   3.903 +lemma peek_and_forget1: "G,A\<turnstile>{P} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{P \<and>. p} t\<succ> {Q}"
   3.904 +apply (erule conseq1)
   3.905 +apply (simp (no_asm))
   3.906 +done
   3.907 +
   3.908 +lemmas ax_NormalD = peek_and_forget1 [of _ _ _ _ _ normal] 
   3.909 +
   3.910 +lemma peek_and_forget2: "G,A\<turnstile>{P} t\<succ> {Q \<and>. p} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
   3.911 +apply (erule conseq2)
   3.912 +apply (simp (no_asm))
   3.913 +done
   3.914 +
   3.915 +lemma ax_subst_Val_allI: "\<forall>v. G,A\<turnstile>{(P'               v )\<leftarrow>Val v} t\<succ> {Q v} \<Longrightarrow>  
   3.916 +      \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In1 w))\<leftarrow>Val v} t\<succ> {Q v}"
   3.917 +apply (force elim!: conseq1)
   3.918 +done
   3.919 +
   3.920 +lemma ax_subst_Var_allI: "\<forall>v. G,A\<turnstile>{(P'               v )\<leftarrow>Var v} t\<succ> {Q v} \<Longrightarrow>  
   3.921 +      \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In2 w))\<leftarrow>Var v} t\<succ> {Q v}"
   3.922 +apply (force elim!: conseq1)
   3.923 +done
   3.924 +
   3.925 +lemma ax_subst_Vals_allI: "(\<forall>v. G,A\<turnstile>{(     P'          v )\<leftarrow>Vals v} t\<succ> {Q v}) \<Longrightarrow>  
   3.926 +       \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In3 w))\<leftarrow>Vals v} t\<succ> {Q v}"
   3.927 +apply (force elim!: conseq1)
   3.928 +done
   3.929 +
   3.930 +
   3.931 +section "alternative axioms"
   3.932 +
   3.933 +lemma ax_Lit2: 
   3.934 +  "G,(A::'a triple set)\<turnstile>{Normal P::'a assn} Lit v-\<succ> {Normal (P\<down>=Val v)}"
   3.935 +apply (rule ax_derivs.Lit [THEN conseq1])
   3.936 +apply force
   3.937 +done
   3.938 +lemma ax_Lit2_test_complete: 
   3.939 +  "G,(A::'a triple set)\<turnstile>{Normal (P\<leftarrow>Val v)::'a assn} Lit v-\<succ> {P}"
   3.940 +apply (rule ax_Lit2 [THEN conseq2])
   3.941 +apply force
   3.942 +done
   3.943 +
   3.944 +lemma ax_LVar2: "G,(A::'a triple set)\<turnstile>{Normal P::'a assn} LVar vn=\<succ> {Normal (\<lambda>s.. P\<down>=Var (lvar vn s))}"
   3.945 +apply (rule ax_derivs.LVar [THEN conseq1])
   3.946 +apply force
   3.947 +done
   3.948 +
   3.949 +lemma ax_Super2: "G,(A::'a triple set)\<turnstile>
   3.950 +  {Normal P::'a assn} Super-\<succ> {Normal (\<lambda>s.. P\<down>=Val (val_this s))}"
   3.951 +apply (rule ax_derivs.Super [THEN conseq1])
   3.952 +apply force
   3.953 +done
   3.954 +
   3.955 +lemma ax_Nil2: 
   3.956 +  "G,(A::'a triple set)\<turnstile>{Normal P::'a assn} []\<doteq>\<succ> {Normal (P\<down>=Vals [])}"
   3.957 +apply (rule ax_derivs.Nil [THEN conseq1])
   3.958 +apply force
   3.959 +done
   3.960 +
   3.961 +
   3.962 +section "misc derived structural rules"
   3.963 +
   3.964 +(* unused *)
   3.965 +lemma ax_finite_mtriples_lemma: "\<lbrakk>F \<subseteq> ms; finite ms; \<forall>(C,sig)\<in>ms. 
   3.966 +    G,(A::'a triple set)\<turnstile>{Normal (P C sig)::'a assn} mb C sig-\<succ> {Q C sig}\<rbrakk> \<Longrightarrow> 
   3.967 +       G,A|\<turnstile>{{P} mb-\<succ> {Q} | F}"
   3.968 +apply (frule (1) finite_subset)
   3.969 +apply (erule make_imp)
   3.970 +apply (erule thin_rl)
   3.971 +apply (erule finite_induct)
   3.972 +apply  (unfold mtriples_def)
   3.973 +apply  (clarsimp intro!: ax_derivs.empty ax_derivs.insert)+
   3.974 +apply force
   3.975 +done
   3.976 +lemmas ax_finite_mtriples = ax_finite_mtriples_lemma [OF subset_refl]
   3.977 +
   3.978 +lemma ax_derivs_insertD: 
   3.979 + "G,(A::'a triple set)|\<turnstile>insert (t::'a triple) ts \<Longrightarrow> G,A\<turnstile>t \<and> G,A|\<turnstile>ts"
   3.980 +apply (fast intro: ax_derivs.weaken)
   3.981 +done
   3.982 +
   3.983 +lemma ax_methods_spec: 
   3.984 +"\<lbrakk>G,(A::'a triple set)|\<turnstile>split f ` ms; (C,sig) \<in> ms\<rbrakk>\<Longrightarrow> G,A\<turnstile>((f C sig)::'a triple)"
   3.985 +apply (erule ax_derivs.weaken)
   3.986 +apply (force del: image_eqI intro: rev_image_eqI)
   3.987 +done
   3.988 +
   3.989 +(* this version is used to avoid using the cut rule *)
   3.990 +lemma ax_finite_pointwise_lemma [rule_format]: "\<lbrakk>F \<subseteq> ms; finite ms\<rbrakk> \<Longrightarrow>  
   3.991 +  ((\<forall>(C,sig)\<in>F. G,(A::'a triple set)\<turnstile>(f C sig::'a triple)) \<longrightarrow> (\<forall>(C,sig)\<in>ms. G,A\<turnstile>(g C sig::'a triple))) \<longrightarrow>  
   3.992 +      G,A|\<turnstile>split f ` F \<longrightarrow> G,A|\<turnstile>split g ` F"
   3.993 +apply (frule (1) finite_subset)
   3.994 +apply (erule make_imp)
   3.995 +apply (erule thin_rl)
   3.996 +apply (erule finite_induct)
   3.997 +apply  clarsimp+
   3.998 +apply (drule ax_derivs_insertD)
   3.999 +apply (rule ax_derivs.insert)
  3.1000 +apply  (simp (no_asm_simp) only: split_tupled_all)
  3.1001 +apply  (auto elim: ax_methods_spec)
  3.1002 +done
  3.1003 +lemmas ax_finite_pointwise = ax_finite_pointwise_lemma [OF subset_refl]
  3.1004 + 
  3.1005 +lemma ax_no_hazard: 
  3.1006 +  "G,(A::'a triple set)\<turnstile>{P \<and>. type_ok G t} t\<succ> {Q::'a assn} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
  3.1007 +apply (erule ax_cases)
  3.1008 +apply (rule ax_derivs.hazard [THEN conseq1])
  3.1009 +apply force
  3.1010 +done
  3.1011 +
  3.1012 +lemma ax_free_wt: 
  3.1013 + "(\<exists>T L C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) 
  3.1014 +  \<longrightarrow> G,(A::'a triple set)\<turnstile>{Normal P} t\<succ> {Q::'a assn} \<Longrightarrow> 
  3.1015 +  G,A\<turnstile>{Normal P} t\<succ> {Q}"
  3.1016 +apply (rule ax_no_hazard)
  3.1017 +apply (rule ax_escape)
  3.1018 +apply clarify
  3.1019 +apply (erule mp [THEN conseq12])
  3.1020 +apply  (auto simp add: type_ok_def)
  3.1021 +done
  3.1022 +
  3.1023 +ML {*
  3.1024 +bind_thms ("ax_Abrupts", sum3_instantiate (thm "ax_derivs.Abrupt"))
  3.1025 +*}
  3.1026 +declare ax_Abrupts [intro!]
  3.1027 +
  3.1028 +lemmas ax_Normal_cases = ax_cases [of _ _ normal]
  3.1029 +
  3.1030 +lemma ax_Skip [intro!]: "G,(A::'a triple set)\<turnstile>{P\<leftarrow>\<diamondsuit>} .Skip. {P::'a assn}"
  3.1031 +apply (rule ax_Normal_cases)
  3.1032 +apply  (rule ax_derivs.Skip)
  3.1033 +apply fast
  3.1034 +done
  3.1035 +lemmas ax_SkipI = ax_Skip [THEN conseq1, standard]
  3.1036 +
  3.1037 +
  3.1038 +section "derived rules for methd call"
  3.1039 +
  3.1040 +lemma ax_Call_known_DynT: 
  3.1041 +"\<lbrakk>G\<turnstile>IntVir\<rightarrow>C\<preceq>statT; 
  3.1042 +  \<forall>a vs l. G,A\<turnstile>{(R a\<leftarrow>Vals vs \<and>. (\<lambda>s. l = locals (store s)) ;.
  3.1043 +  init_lvars G C \<lparr>name=mn,parTs=pTs\<rparr> IntVir a vs)} 
  3.1044 +    Methd C \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}; 
  3.1045 +  \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ>  
  3.1046 +       {R a \<and>. (\<lambda>s. C = obj_class (the (heap (store s) (the_Addr a))) \<and>
  3.1047 +                     C = invocation_declclass 
  3.1048 +                            G IntVir (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> )};  
  3.1049 +       G,(A::'a triple set)\<turnstile>{Normal P} e-\<succ> {Q::'a assn}\<rbrakk>  
  3.1050 +   \<Longrightarrow> G,A\<turnstile>{Normal P} {statT,IntVir}e\<cdot>mn({pTs}args)-\<succ> {S}"
  3.1051 +apply (erule ax_derivs.Call)
  3.1052 +apply  safe
  3.1053 +apply  (erule spec)
  3.1054 +apply (rule ax_escape, clarsimp)
  3.1055 +apply (drule spec, drule spec, drule spec,erule conseq12)
  3.1056 +apply force
  3.1057 +done
  3.1058 +
  3.1059 +
  3.1060 +lemma ax_Call_Static: 
  3.1061 + "\<lbrakk>\<forall>a vs l. G,A\<turnstile>{R a\<leftarrow>Vals vs \<and>. (\<lambda>s. l = locals (store s)) ;.  
  3.1062 +               init_lvars G C \<lparr>name=mn,parTs=pTs\<rparr> Static any_Addr vs}  
  3.1063 +              Methd C \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}; 
  3.1064 +  G,A\<turnstile>{Normal P} e-\<succ> {Q};
  3.1065 +  \<forall> a. G,(A::'a triple set)\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {(R::val \<Rightarrow> 'a assn)  a 
  3.1066 +  \<and>. (\<lambda> s. C=invocation_declclass 
  3.1067 +                G Static (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr>)}
  3.1068 +\<rbrakk>  \<Longrightarrow>  G,A\<turnstile>{Normal P} {statT,Static}e\<cdot>mn({pTs}args)-\<succ> {S}"
  3.1069 +apply (erule ax_derivs.Call)
  3.1070 +apply  safe
  3.1071 +apply  (erule spec)
  3.1072 +apply (rule ax_escape, clarsimp)
  3.1073 +apply (erule_tac V = "?P \<longrightarrow> ?Q" in thin_rl)
  3.1074 +apply (drule spec,drule spec,drule spec, erule conseq12)
  3.1075 +apply (force simp add: init_lvars_def)
  3.1076 +done
  3.1077 +
  3.1078 +lemma ax_Methd1: 
  3.1079 + "\<lbrakk>G,A\<union>{{P} Methd-\<succ> {Q} | ms}|\<turnstile> {{P} body G-\<succ> {Q} | ms}; (C,sig)\<in> ms\<rbrakk> \<Longrightarrow> 
  3.1080 +       G,A\<turnstile>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
  3.1081 +apply (drule ax_derivs.Methd)
  3.1082 +apply (unfold mtriples_def)
  3.1083 +apply (erule (1) ax_methods_spec)
  3.1084 +done
  3.1085 +
  3.1086 +lemma ax_MethdN: 
  3.1087 +"G,insert({Normal P} Methd  C sig-\<succ> {Q}) A\<turnstile> 
  3.1088 +          {Normal P} body G C sig-\<succ> {Q} \<Longrightarrow>  
  3.1089 +      G,A\<turnstile>{Normal P} Methd   C sig-\<succ> {Q}"
  3.1090 +apply (rule ax_Methd1)
  3.1091 +apply  (rule_tac [2] singletonI)
  3.1092 +apply (unfold mtriples_def)
  3.1093 +apply clarsimp
  3.1094 +done
  3.1095 +
  3.1096 +lemma ax_StatRef: 
  3.1097 +  "G,(A::'a triple set)\<turnstile>{Normal (P\<leftarrow>Val Null)} StatRef rt-\<succ> {P::'a assn}"
  3.1098 +apply (rule ax_derivs.Cast)
  3.1099 +apply (rule ax_Lit2 [THEN conseq2])
  3.1100 +apply clarsimp
  3.1101 +done
  3.1102 +
  3.1103 +section "rules derived from Init and Done"
  3.1104 +
  3.1105 +  lemma ax_InitS: "\<lbrakk>the (class G C) = c; C \<noteq> Object;  
  3.1106 +     \<forall>l. G,A\<turnstile>{Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars empty}  
  3.1107 +            .init c. {set_lvars l .; R};   
  3.1108 +         G,A\<turnstile>{Normal ((P \<and>. Not \<circ> initd C) ;. supd (init_class_obj G C))}  
  3.1109 +  .Init (super c). {Q}\<rbrakk> \<Longrightarrow>  
  3.1110 +  G,(A::'a triple set)\<turnstile>{Normal (P \<and>. Not \<circ> initd C)} .Init C. {R::'a assn}"
  3.1111 +apply (erule ax_derivs.Init)
  3.1112 +apply  (simp (no_asm_simp))
  3.1113 +apply assumption
  3.1114 +done
  3.1115 +
  3.1116 +lemma ax_Init_Skip_lemma: 
  3.1117 +"\<forall>l. G,(A::'a triple set)\<turnstile>{P\<leftarrow>\<diamondsuit> \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars l'}
  3.1118 +  .Skip. {(set_lvars l .; P)::'a assn}"
  3.1119 +apply (rule allI)
  3.1120 +apply (rule ax_SkipI)
  3.1121 +apply clarsimp
  3.1122 +done
  3.1123 +
  3.1124 +lemma ax_triv_InitS: "\<lbrakk>the (class G C) = c;init c = Skip; C \<noteq> Object; 
  3.1125 +       P\<leftarrow>\<diamondsuit> \<Rightarrow> (supd (init_class_obj G C) .; P);  
  3.1126 +       G,A\<turnstile>{Normal (P \<and>. initd C)} .Init (super c). {(P \<and>. initd C)\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>  
  3.1127 +       G,(A::'a triple set)\<turnstile>{Normal P\<leftarrow>\<diamondsuit>} .Init C. {(P \<and>. initd C)::'a assn}"
  3.1128 +apply (rule_tac C = "initd C" in ax_cases)
  3.1129 +apply  (rule conseq1, rule ax_derivs.Done, clarsimp)
  3.1130 +apply (simp (no_asm))
  3.1131 +apply (erule (1) ax_InitS)
  3.1132 +apply  simp
  3.1133 +apply  (rule ax_Init_Skip_lemma)
  3.1134 +apply (erule conseq1)
  3.1135 +apply force
  3.1136 +done
  3.1137 +
  3.1138 +lemma ax_Init_Object: "wf_prog G \<Longrightarrow> G,(A::'a triple set)\<turnstile>
  3.1139 +  {Normal ((supd (init_class_obj G Object) .; P\<leftarrow>\<diamondsuit>) \<and>. Not \<circ> initd Object)} 
  3.1140 +       .Init Object. {(P \<and>. initd Object)::'a assn}"
  3.1141 +apply (rule ax_derivs.Init)
  3.1142 +apply   (drule class_Object, force)
  3.1143 +apply (simp_all (no_asm))
  3.1144 +apply (rule_tac [2] ax_Init_Skip_lemma)
  3.1145 +apply (rule ax_SkipI, force)
  3.1146 +done
  3.1147 +
  3.1148 +lemma ax_triv_Init_Object: "\<lbrakk>wf_prog G;  
  3.1149 +       (P::'a assn) \<Rightarrow> (supd (init_class_obj G Object) .; P)\<rbrakk> \<Longrightarrow>  
  3.1150 +  G,(A::'a triple set)\<turnstile>{Normal P\<leftarrow>\<diamondsuit>} .Init Object. {P \<and>. initd Object}"
  3.1151 +apply (rule_tac C = "initd Object" in ax_cases)
  3.1152 +apply  (rule conseq1, rule ax_derivs.Done, clarsimp)
  3.1153 +apply (erule ax_Init_Object [THEN conseq1])
  3.1154 +apply force
  3.1155 +done
  3.1156 +
  3.1157 +
  3.1158 +section "introduction rules for Alloc and SXAlloc"
  3.1159 +
  3.1160 +lemma ax_SXAlloc_Normal: "G,A\<turnstile>{P} .c. {Normal Q} \<Longrightarrow> G,A\<turnstile>{P} .c. {SXAlloc G Q}"
  3.1161 +apply (erule conseq2)
  3.1162 +apply (clarsimp elim!: sxalloc_elim_cases simp add: split_tupled_all)
  3.1163 +done
  3.1164 +
  3.1165 +lemma ax_Alloc: 
  3.1166 +  "G,A\<turnstile>{P} t\<succ> {Normal (\<lambda>Y (x,s) Z. (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
  3.1167 + Q (Val (Addr a)) (Norm(init_obj G (CInst C) (Heap a) s)) Z)) \<and>. 
  3.1168 +    heap_free (Suc (Suc 0))}
  3.1169 +   \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Alloc G (CInst C) Q}"
  3.1170 +apply (erule conseq2)
  3.1171 +apply (auto elim!: halloc_elim_cases)
  3.1172 +done
  3.1173 +
  3.1174 +lemma ax_Alloc_Arr: 
  3.1175 + "G,A\<turnstile>{P} t\<succ> {\<lambda>Val:i:. Normal (\<lambda>Y (x,s) Z. \<not>the_Intg i<0 \<and>  
  3.1176 +  (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
  3.1177 +  Q (Val (Addr a)) (Norm (init_obj G (Arr T (the_Intg i)) (Heap a) s)) Z)) \<and>. 
  3.1178 +   heap_free (Suc (Suc 0))} \<Longrightarrow>  
  3.1179 + G,A\<turnstile>{P} t\<succ> {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T(the_Intg i)) Q}"
  3.1180 +apply (erule conseq2)
  3.1181 +apply (auto elim!: halloc_elim_cases)
  3.1182 +done
  3.1183 +
  3.1184 +lemma ax_SXAlloc_catch_SXcpt: 
  3.1185 + "\<lbrakk>G,A\<turnstile>{P} t\<succ> {(\<lambda>Y (x,s) Z. x=Some (Xcpt (Std xn)) \<and>  
  3.1186 +  (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
  3.1187 +  Q Y (Some (Xcpt (Loc a)),init_obj G (CInst (SXcpt xn)) (Heap a) s) Z))  
  3.1188 +  \<and>. heap_free (Suc (Suc 0))}\<rbrakk> \<Longrightarrow>  
  3.1189 +  G,A\<turnstile>{P} t\<succ> {SXAlloc G (\<lambda>Y s Z. Q Y s Z \<and> G,s\<turnstile>catch SXcpt xn)}"
  3.1190 +apply (erule conseq2)
  3.1191 +apply (auto elim!: sxalloc_elim_cases halloc_elim_cases)
  3.1192 +done
  3.1193 +
  3.1194 +end
     4.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     4.2 +++ b/src/HOL/Bali/AxSound.thy	Mon Jan 28 17:00:19 2002 +0100
     4.3 @@ -0,0 +1,434 @@
     4.4 +(*  Title:      isabelle/Bali/AxSound.thy
     4.5 +    ID:         $Id$
     4.6 +    Author:     David von Oheimb
     4.7 +    Copyright   1999 Technische Universitaet Muenchen
     4.8 +*)
     4.9 +header {* Soundness proof for Axiomatic semantics of Java expressions and 
    4.10 +          statements
    4.11 +       *}
    4.12 +
    4.13 +
    4.14 +
    4.15 +theory AxSound = AxSem:
    4.16 +
    4.17 +section "validity"
    4.18 +
    4.19 +consts
    4.20 +
    4.21 + triple_valid2:: "prog \<Rightarrow> nat \<Rightarrow>        'a triple  \<Rightarrow> bool"
    4.22 +                                                (   "_\<Turnstile>_\<Colon>_"[61,0, 58] 57)
    4.23 +    ax_valids2:: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool"
    4.24 +                                                ("_,_|\<Turnstile>\<Colon>_" [61,58,58] 57)
    4.25 +
    4.26 +defs  triple_valid2_def: "G\<Turnstile>n\<Colon>t \<equiv> case t of {P} t\<succ> {Q} \<Rightarrow>
    4.27 + \<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>L. s\<Colon>\<preceq>(G,L) 
    4.28 + \<longrightarrow> (\<forall>T C. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) \<longrightarrow>
    4.29 + (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> Q Y' s' Z \<and> s'\<Colon>\<preceq>(G,L))))"
    4.30 +
    4.31 +defs  ax_valids2_def:    "G,A|\<Turnstile>\<Colon>ts \<equiv>  \<forall>n. (\<forall>t\<in>A. G\<Turnstile>n\<Colon>t) \<longrightarrow> (\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t)"
    4.32 +
    4.33 +lemma triple_valid2_def2: "G\<Turnstile>n\<Colon>{P} t\<succ> {Q} =  
    4.34 + (\<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s')\<longrightarrow>  
    4.35 +  (\<forall>L. s\<Colon>\<preceq>(G,L) \<longrightarrow> (\<forall>T C. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) \<longrightarrow>  
    4.36 +  Q Y' s' Z \<and> s'\<Colon>\<preceq>(G,L)))))"
    4.37 +apply (unfold triple_valid2_def)
    4.38 +apply (simp (no_asm) add: split_paired_All)
    4.39 +apply blast
    4.40 +done
    4.41 +
    4.42 +lemma triple_valid2_eq [rule_format (no_asm)]: 
    4.43 +  "wf_prog G ==> triple_valid2 G = triple_valid G"
    4.44 +apply (rule ext)
    4.45 +apply (rule ext)
    4.46 +apply (rule triple.induct)
    4.47 +apply (simp (no_asm) add: triple_valid_def2 triple_valid2_def2)
    4.48 +apply (rule iffI)
    4.49 +apply  fast
    4.50 +apply clarify
    4.51 +apply (tactic "smp_tac 3 1")
    4.52 +apply (case_tac "normal s")
    4.53 +apply  clarsimp
    4.54 +apply  (blast dest: evaln_eval eval_type_sound [THEN conjunct1])
    4.55 +apply clarsimp
    4.56 +done
    4.57 +
    4.58 +lemma ax_valids2_eq: "wf_prog G \<Longrightarrow> G,A|\<Turnstile>\<Colon>ts = G,A|\<Turnstile>ts"
    4.59 +apply (unfold ax_valids_def ax_valids2_def)
    4.60 +apply (force simp add: triple_valid2_eq)
    4.61 +done
    4.62 +
    4.63 +lemma triple_valid2_Suc [rule_format (no_asm)]: "G\<Turnstile>Suc n\<Colon>t \<longrightarrow> G\<Turnstile>n\<Colon>t"
    4.64 +apply (induct_tac "t")
    4.65 +apply (subst triple_valid2_def2)
    4.66 +apply (subst triple_valid2_def2)
    4.67 +apply (fast intro: evaln_nonstrict_Suc)
    4.68 +done
    4.69 +
    4.70 +lemma Methd_triple_valid2_0: "G\<Turnstile>0\<Colon>{Normal P} Methd C sig-\<succ> {Q}"
    4.71 +apply (clarsimp elim!: evaln_elim_cases simp add: triple_valid2_def2)
    4.72 +done
    4.73 +
    4.74 +lemma Methd_triple_valid2_SucI: 
    4.75 +"\<lbrakk>G\<Turnstile>n\<Colon>{Normal P} body G C sig-\<succ>{Q}\<rbrakk> 
    4.76 +  \<Longrightarrow> G\<Turnstile>Suc n\<Colon>{Normal P} Methd C sig-\<succ> {Q}"
    4.77 +apply (simp (no_asm_use) add: triple_valid2_def2)
    4.78 +apply (intro strip, tactic "smp_tac 3 1", clarify)
    4.79 +apply (erule wt_elim_cases, erule evaln_elim_cases)
    4.80 +apply (unfold body_def Let_def)
    4.81 +apply clarsimp
    4.82 +apply blast
    4.83 +done
    4.84 +
    4.85 +lemma triples_valid2_Suc: 
    4.86 + "Ball ts (triple_valid2 G (Suc n)) \<Longrightarrow> Ball ts (triple_valid2 G n)"
    4.87 +apply (fast intro: triple_valid2_Suc)
    4.88 +done
    4.89 +
    4.90 +lemma "G|\<Turnstile>n:insert t A = (G\<Turnstile>n:t \<and> G|\<Turnstile>n:A)";
    4.91 +oops
    4.92 +
    4.93 +
    4.94 +section "soundness"
    4.95 +
    4.96 +lemma Methd_sound: 
    4.97 +"\<lbrakk>G,A\<union>  {{P} Methd-\<succ> {Q} | ms}|\<Turnstile>\<Colon>{{P} body G-\<succ> {Q} | ms}\<rbrakk> \<Longrightarrow> 
    4.98 +  G,A|\<Turnstile>\<Colon>{{P} Methd-\<succ> {Q} | ms}"
    4.99 +apply (unfold ax_valids2_def mtriples_def)
   4.100 +apply (rule allI)
   4.101 +apply (induct_tac "n")
   4.102 +apply  (clarify, tactic {* pair_tac "x" 1 *}, simp (no_asm))
   4.103 +apply  (fast intro: Methd_triple_valid2_0)
   4.104 +apply (clarify, tactic {* pair_tac "xa" 1 *}, simp (no_asm))
   4.105 +apply (drule triples_valid2_Suc)
   4.106 +apply (erule (1) notE impE)
   4.107 +apply (drule_tac x = na in spec)
   4.108 +apply (tactic {* auto_tac (claset() addSIs [thm "Methd_triple_valid2_SucI"],
   4.109 +   simpset() addsimps [ball_Un] addloop ("split_all_tac", split_all_tac)) *})
   4.110 +done
   4.111 +
   4.112 +
   4.113 +lemma valids2_inductI: "\<forall>s t n Y' s'. G\<turnstile>s\<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> t = c \<longrightarrow>    
   4.114 +  Ball A (triple_valid2 G n) \<longrightarrow> (\<forall>Y Z. P Y s Z \<longrightarrow>  
   4.115 +  (\<forall>L. s\<Colon>\<preceq>(G,L) \<longrightarrow> (\<forall>T C. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) \<longrightarrow>  
   4.116 +  Q Y' s' Z \<and> s'\<Colon>\<preceq>(G, L)))) \<Longrightarrow>  
   4.117 +  G,A|\<Turnstile>\<Colon>{ {P} c\<succ> {Q}}"
   4.118 +apply (simp (no_asm) add: ax_valids2_def triple_valid2_def2)
   4.119 +apply clarsimp
   4.120 +done
   4.121 +
   4.122 +ML_setup {*
   4.123 +Delsimprocs [evaln_expr_proc,evaln_var_proc,evaln_exprs_proc,evaln_stmt_proc]
   4.124 +*}
   4.125 +
   4.126 +lemma Loop_sound: "\<lbrakk>G,A|\<Turnstile>\<Colon>{ {P} e-\<succ> {P'}};  
   4.127 +       G,A|\<Turnstile>\<Colon>{ {Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}}\<rbrakk> \<Longrightarrow>  
   4.128 +       G,A|\<Turnstile>\<Colon>{ {P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}}"
   4.129 +apply (rule valids2_inductI)
   4.130 +apply ((rule allI)+, rule impI, tactic {* pair_tac "s" 1*}, tactic {* pair_tac "s'" 1*})
   4.131 +apply (erule evaln.induct)
   4.132 +apply  simp_all (* takes half a minute *)
   4.133 +apply  clarify
   4.134 +apply  (erule_tac V = "G,A|\<Turnstile>\<Colon>{ {?P'} .c. {?P}}" in thin_rl)
   4.135 +apply  (simp_all (no_asm_use) add: ax_valids2_def triple_valid2_def2)
   4.136 +apply  (tactic "smp_tac 1 1", tactic "smp_tac 3 1", force)
   4.137 +apply clarify
   4.138 +apply (rule wt_elim_cases, assumption)
   4.139 +apply (tactic "smp_tac 1 1", tactic "smp_tac 1 1", tactic "smp_tac 3 1", 
   4.140 +       tactic "smp_tac 2 1", tactic "smp_tac 1 1")
   4.141 +apply (erule impE,simp (no_asm),blast)
   4.142 +apply (simp add: imp_conjL split_tupled_all split_paired_All)
   4.143 +apply (case_tac "the_Bool b")
   4.144 +apply  clarsimp
   4.145 +apply  (case_tac "a")
   4.146 +apply (simp_all)
   4.147 +apply clarsimp
   4.148 +apply  (erule_tac V = "c = l\<bullet> While(e) c \<longrightarrow> ?P" in thin_rl)
   4.149 +apply (blast intro: conforms_absorb)
   4.150 +apply blast+
   4.151 +done
   4.152 +
   4.153 +declare subst_Bool_def2 [simp del]
   4.154 +lemma all_empty: "(!x. P) = P"
   4.155 +by simp
   4.156 +lemma sound_valid2_lemma: 
   4.157 +"\<lbrakk>\<forall>v n. Ball A (triple_valid2 G n) \<longrightarrow> P v n; Ball A (triple_valid2 G n)\<rbrakk>
   4.158 + \<Longrightarrow>P v n"
   4.159 +by blast
   4.160 +ML {*
   4.161 +val fullsimptac = full_simp_tac(simpset() delsimps [thm "all_empty"]);
   4.162 +val sound_prepare_tac = EVERY'[REPEAT o thin_tac "?x \<in> ax_derivs G",
   4.163 + full_simp_tac (simpset()addsimps[thm "ax_valids2_def",thm "triple_valid2_def2",
   4.164 +                                  thm "imp_conjL"] delsimps[thm "all_empty"]),
   4.165 + Clarify_tac];
   4.166 +val sound_elim_tac = EVERY'[eresolve_tac (thms "evaln_elim_cases"), 
   4.167 +        TRY o eresolve_tac (thms "wt_elim_cases"), fullsimptac, Clarify_tac];
   4.168 +val sound_valid2_tac = REPEAT o FIRST'[smp_tac 1, 
   4.169 +                  datac (thm "sound_valid2_lemma") 1];
   4.170 +val sound_forw_hyp_tac = 
   4.171 + EVERY'[smp_tac 3 
   4.172 +          ORELSE' EVERY'[dtac spec,dtac spec, dtac spec,etac impE, Fast_tac] 
   4.173 +          ORELSE' EVERY'[dtac spec,dtac spec,etac impE, Fast_tac],
   4.174 +        fullsimptac, 
   4.175 +        smp_tac 2,TRY o smp_tac 1,
   4.176 +        TRY o EVERY'[etac impE, TRY o rtac impI, 
   4.177 +        atac ORELSE' (EVERY' [REPEAT o rtac exI,Blast_tac]),
   4.178 +        fullsimptac, Clarify_tac, TRY o smp_tac 1]]
   4.179 +*}
   4.180 +(* ### rtac conjI,rtac HOL.refl *)
   4.181 +lemma Call_sound: 
   4.182 +"\<lbrakk>wf_prog G; G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q}}; \<forall>a. G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>Val a} ps\<doteq>\<succ> {R a}};
   4.183 +  \<forall>a vs invC declC l. G,A|\<Turnstile>\<Colon>{ {(R a\<leftarrow>Vals vs \<and>.  
   4.184 +   (\<lambda>s. declC = invocation_declclass 
   4.185 +                    G mode (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> \<and>
   4.186 +         invC = invocation_class mode (store s) a statT \<and>
   4.187 +            l = locals (store s)) ;.  
   4.188 +   init_lvars G declC \<lparr>name=mn,parTs=pTs\<rparr> mode a vs) \<and>.  
   4.189 +   (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)}  
   4.190 +   Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}}\<rbrakk> \<Longrightarrow>  
   4.191 +  G,A|\<Turnstile>\<Colon>{ {Normal P} {statT,mode}e\<cdot>mn({pTs}ps)-\<succ> {S}}"
   4.192 +apply (tactic "EVERY'[sound_prepare_tac, sound_elim_tac, sound_valid2_tac] 1")
   4.193 +apply (rename_tac x1 s1 x2 s2 ab bb v vs m pTsa statDeclC)
   4.194 +apply (tactic "smp_tac 6 1")
   4.195 +apply (tactic "sound_forw_hyp_tac 1")
   4.196 +apply (tactic "sound_forw_hyp_tac 1")
   4.197 +apply (drule max_spec2mheads)
   4.198 +apply (drule evaln_eval, drule (3) eval_ts)
   4.199 +apply (drule evaln_eval, frule (3) evals_ts)
   4.200 +apply (drule spec,erule impE,rule exI, blast)
   4.201 +(* apply (drule spec,drule spec,drule spec,erule impE,rule exI,blast) *)
   4.202 +apply (case_tac "if static m then x2 else (np a') x2")
   4.203 +defer 1
   4.204 +apply  (rename_tac x, subgoal_tac "(Some x, s2)\<Colon>\<preceq>(G, L)" (* used two times *))
   4.205 +prefer 2 
   4.206 +apply   (force split add: split_if_asm)
   4.207 +apply  (simp del: if_raise_if_None)
   4.208 +apply  (tactic "smp_tac 2 1")
   4.209 +apply (simp only: init_lvars_def2 invmode_Static_eq)
   4.210 +apply (clarsimp simp del: resTy_mthd)
   4.211 +apply  (drule spec,erule swap,erule conforms_set_locals [OF _ lconf_empty])
   4.212 +apply clarsimp
   4.213 +apply (drule Null_staticD)
   4.214 +apply (drule eval_gext', drule (1) conf_gext, frule (3) DynT_propI)
   4.215 +apply (erule impE) apply blast
   4.216 +apply (subgoal_tac 
   4.217 + "G\<turnstile>invmode (mhd (statDeclC,m)) e
   4.218 +     \<rightarrow>invocation_class (invmode m e) s2 a' statT\<preceq>statT")
   4.219 +defer   apply simp
   4.220 +apply (drule (3) DynT_mheadsD,simp,simp)
   4.221 +apply (clarify, drule wf_mdeclD1, clarify)
   4.222 +apply (frule ty_expr_is_type) apply simp
   4.223 +apply (subgoal_tac "invmode (mhd (statDeclC,m)) e = IntVir \<longrightarrow> a' \<noteq> Null")
   4.224 +defer   apply simp
   4.225 +apply (frule (2) wt_MethdI)
   4.226 +apply clarify
   4.227 +apply (drule (2) conforms_init_lvars)
   4.228 +apply   (simp) 
   4.229 +apply   (assumption)+
   4.230 +apply   simp
   4.231 +apply   (assumption)+
   4.232 +apply   (rule impI) apply simp
   4.233 +apply   simp
   4.234 +apply   simp
   4.235 +apply   (rule Ball_weaken)
   4.236 +apply     assumption
   4.237 +apply     (force simp add: is_acc_type_def)
   4.238 +apply (tactic "smp_tac 2 1")
   4.239 +apply simp
   4.240 +apply (tactic "smp_tac 1 1")
   4.241 +apply (erule_tac V = "?P \<longrightarrow> ?Q" in thin_rl) 
   4.242 +apply (erule impE)
   4.243 +apply   (rule exI)+ 
   4.244 +apply   (subgoal_tac "is_static dm = (static m)") 
   4.245 +prefer 2  apply (simp add: member_is_static_simp)
   4.246 +apply   (simp only: )
   4.247 +apply   (simp only: sig.simps)
   4.248 +apply (force dest!: evaln_eval eval_gext' elim: conforms_return 
   4.249 +             del: impCE simp add: init_lvars_def2)
   4.250 +done
   4.251 +
   4.252 +lemma Init_sound: "\<lbrakk>wf_prog G; the (class G C) = c;  
   4.253 +      G,A|\<Turnstile>\<Colon>{ {Normal ((P \<and>. Not \<circ> initd C) ;. supd (init_class_obj G C))}  
   4.254 +             .(if C = Object then Skip else Init (super c)). {Q}};  
   4.255 +  \<forall>l. G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars empty}  
   4.256 +            .init c. {set_lvars l .; R}}\<rbrakk> \<Longrightarrow>  
   4.257 +      G,A|\<Turnstile>\<Colon>{ {Normal (P \<and>. Not \<circ> initd C)} .Init C. {R}}"
   4.258 +apply (tactic "EVERY'[sound_prepare_tac, sound_elim_tac,sound_valid2_tac] 1")
   4.259 +apply (tactic {* instantiate_tac [("l24","\<lambda> n Y Z sa Y' s' L y a b aa ba ab bb. locals b")]*})
   4.260 +apply (clarsimp simp add: split_paired_Ex)
   4.261 +apply (drule spec, drule spec, drule spec, erule impE)
   4.262 +apply  (erule_tac V = "All ?P" in thin_rl, fast)
   4.263 +apply clarsimp
   4.264 +apply (tactic "smp_tac 2 1", drule spec, erule impE, 
   4.265 +       erule (3) conforms_init_class_obj)
   4.266 +apply (drule (1) wf_prog_cdecl)
   4.267 +apply (erule impE, rule exI,erule_tac V = "All ?P" in thin_rl,
   4.268 +       force dest: wf_cdecl_supD split add: split_if simp add: is_acc_class_def)
   4.269 +apply clarify
   4.270 +apply (drule spec)
   4.271 +apply (drule spec)
   4.272 +apply (drule spec)
   4.273 +apply  (erule impE)
   4.274 +apply ( fast)
   4.275 +apply (simp (no_asm_use) del: empty_def2)
   4.276 +apply (tactic "smp_tac 2 1")
   4.277 +apply (drule spec, erule impE, erule conforms_set_locals, rule lconf_empty)
   4.278 +apply (erule impE,rule impI,rule exI, erule wf_cdecl_wt_init)
   4.279 +apply clarsimp
   4.280 +apply (erule (1) conforms_return, force dest: evaln_eval eval_gext')
   4.281 +done
   4.282 +
   4.283 +lemma all_conjunct2: "\<forall>l. P' l \<and> P l \<Longrightarrow> \<forall>l. P l"
   4.284 +by fast
   4.285 +
   4.286 +lemma all4_conjunct2: 
   4.287 +  "\<forall>a vs D l. (P' a vs D l \<and> P a vs D l) \<Longrightarrow> \<forall>a vs D l. P a vs D l"
   4.288 +by fast
   4.289 +
   4.290 +lemmas sound_lemmas = Init_sound Loop_sound Methd_sound
   4.291 +
   4.292 +lemma ax_sound2: "wf_prog G \<Longrightarrow> G,A|\<turnstile>ts \<Longrightarrow> G,A|\<Turnstile>\<Colon>ts"
   4.293 +apply (erule ax_derivs.induct)
   4.294 +prefer 20 (* Call *)
   4.295 +apply (erule (1) Call_sound) apply simp apply force apply force 
   4.296 +
   4.297 +apply (tactic {* TRYALL (eresolve_tac (thms "sound_lemmas") THEN_ALL_NEW 
   4.298 +    eresolve_tac [asm_rl, thm "all_conjunct2", thm "all4_conjunct2"]) *})
   4.299 +
   4.300 +apply(tactic "COND (has_fewer_prems(30+3)) (ALLGOALS sound_prepare_tac) no_tac")
   4.301 +
   4.302 +               (*empty*)
   4.303 +apply        fast (* insert *)
   4.304 +apply       fast (* asm *)
   4.305 +(*apply    fast *) (* cut *)
   4.306 +apply     fast (* weaken *)
   4.307 +apply    (tactic "smp_tac 3 1", clarify, tactic "smp_tac 1 1", frule evaln_eval,
   4.308 +(* conseq *)case_tac"fst s",clarsimp simp add: eval_type_sound [THEN conjunct1],
   4.309 +clarsimp)
   4.310 +apply   (simp (no_asm_use) add: type_ok_def, drule evaln_eval,fast) (* hazard *)
   4.311 +apply  force (* Abrupt *)
   4.312 +
   4.313 +(* 25 subgoals *)
   4.314 +apply (tactic {* ALLGOALS sound_elim_tac *})(* LVar, Lit, Super, Nil, Skip,Do *)
   4.315 +apply (tactic {* ALLGOALS (asm_simp_tac (noAll_simpset() 
   4.316 +                          delsimps [thm "all_empty"])) *})    (* Done *)
   4.317 +(* for FVar *)
   4.318 +apply (frule wf_ws_prog) 
   4.319 +apply (frule ty_expr_is_type [THEN type_is_class, 
   4.320 +                              THEN accfield_declC_is_class])
   4.321 +apply (simp,simp,simp) 
   4.322 +apply (frule_tac [4] wt_init_comp_ty) (* for NewA*)
   4.323 +apply (tactic "ALLGOALS sound_valid2_tac")
   4.324 +apply (tactic "TRYALL sound_forw_hyp_tac") (* Cast, Inst, Acc, Expr *)
   4.325 +apply (tactic {* TRYALL (EVERY'[dtac spec, TRY o EVERY'[rotate_tac ~1, 
   4.326 +  dtac spec], dtac conjunct2, smp_tac 1, 
   4.327 +  TRY o dres_inst_tac [("P","P'")] (thm "subst_Bool_the_BoolI")]) *})
   4.328 +apply (frule_tac [14] x = x1 in conforms_NormI)  (* for Fin *)
   4.329 +
   4.330 +
   4.331 +(* 15 subgoals *)
   4.332 +(* FVar *)
   4.333 +apply (tactic "sound_forw_hyp_tac 1")
   4.334 +apply (clarsimp simp add: fvar_def2 Let_def split add: split_if_asm)
   4.335 +
   4.336 +(* AVar *)
   4.337 +(*
   4.338 +apply (drule spec, drule spec, erule impE, fast)
   4.339 +apply (simp)
   4.340 +apply (tactic "smp_tac 2 1")
   4.341 +apply (tactic "smp_tac 1 1")
   4.342 +apply (erule impE)
   4.343 +apply (rule impI)
   4.344 +apply (rule exI)+
   4.345 +apply blast
   4.346 +apply (clarsimp simp add: avar_def2)
   4.347 +*)
   4.348 +apply (tactic "sound_forw_hyp_tac 1")
   4.349 +apply (clarsimp simp add: avar_def2)
   4.350 +
   4.351 +(* NewC *)
   4.352 +apply (clarsimp simp add: is_acc_class_def)
   4.353 +apply (erule (1) halloc_conforms, simp, simp)
   4.354 +
   4.355 +(* NewA *)
   4.356 +apply (tactic "sound_forw_hyp_tac 1")
   4.357 +apply (rule conjI,blast)
   4.358 +apply (erule (1) halloc_conforms, simp, simp, simp add: is_acc_type_def)
   4.359 +
   4.360 +(* Ass *)
   4.361 +apply (tactic "sound_forw_hyp_tac 1")
   4.362 +apply (case_tac "aa")
   4.363 +prefer 2
   4.364 +apply  clarsimp
   4.365 +apply (drule evaln_eval)+
   4.366 +apply (frule (3) eval_ts)
   4.367 +apply clarsimp
   4.368 +apply (frule (3) evar_ts [THEN conjunct2])
   4.369 +apply (frule wf_ws_prog)
   4.370 +apply (drule (2) conf_widen)
   4.371 +apply (drule_tac "s1.0" = b in eval_gext')
   4.372 +apply (clarsimp simp add: assign_conforms_def)
   4.373 +
   4.374 +(* Cond *)
   4.375 +
   4.376 +apply (tactic "smp_tac 3 1") apply (tactic "smp_tac 2 1") 
   4.377 +apply (tactic "smp_tac 1 1") apply (erule impE) 
   4.378 +apply (rule impI,rule exI) 
   4.379 +apply (rule_tac x = "if the_Bool b then T1 else T2" in exI)
   4.380 +apply (force split add: split_if)
   4.381 +apply assumption
   4.382 +
   4.383 +(* Body *)
   4.384 +apply (tactic "sound_forw_hyp_tac 1")
   4.385 +apply (rule conforms_absorb,assumption)
   4.386 +
   4.387 +(* Lab *)
   4.388 +apply (tactic "sound_forw_hyp_tac 1")
   4.389 +apply (rule conforms_absorb,assumption)
   4.390 +
   4.391 +(* If *)
   4.392 +apply (tactic "sound_forw_hyp_tac 1")
   4.393 +apply (tactic "sound_forw_hyp_tac 1")
   4.394 +apply (force split add: split_if)
   4.395 +
   4.396 +(* Throw *)
   4.397 +apply (drule evaln_eval, drule (3) eval_ts)
   4.398 +apply clarsimp
   4.399 +apply (drule (3) Throw_lemma)
   4.400 +apply clarsimp
   4.401 +
   4.402 +(* Try *)
   4.403 +apply (frule (1) sxalloc_type_sound)
   4.404 +apply (erule sxalloc_elim_cases2)
   4.405 +apply  (tactic "smp_tac 3 1")
   4.406 +apply  (clarsimp split add: option.split_asm)
   4.407 +apply (clarsimp split add: option.split_asm)
   4.408 +apply (tactic "smp_tac 1 1")
   4.409 +apply (simp only: split add: split_if_asm)
   4.410 +prefer 2
   4.411 +apply  (tactic "smp_tac 3 1", erule_tac V = "All ?P" in thin_rl, clarsimp)
   4.412 +apply (drule spec, erule_tac V = "All ?P" in thin_rl, drule spec, drule spec, 
   4.413 +       erule impE, force)
   4.414 +apply (frule (2) Try_lemma)
   4.415 +apply clarsimp
   4.416 +apply (fast elim!: conforms_deallocL)
   4.417 +
   4.418 +(* Fin *)
   4.419 +apply (tactic "sound_forw_hyp_tac 1")
   4.420 +apply (case_tac "x1", force)
   4.421 +apply clarsimp
   4.422 +apply (drule evaln_eval, drule (4) Fin_lemma)
   4.423 +done
   4.424 +
   4.425 +
   4.426 +
   4.427 +declare subst_Bool_def2 [simp]
   4.428 +
   4.429 +theorem ax_sound: 
   4.430 + "wf_prog G \<Longrightarrow> G,(A::'a triple set)|\<turnstile>(ts::'a triple set) \<Longrightarrow> G,A|\<Turnstile>ts"
   4.431 +apply (subst ax_valids2_eq [symmetric])
   4.432 +apply  assumption
   4.433 +apply (erule (1) ax_sound2)
   4.434 +done
   4.435 +
   4.436 +
   4.437 +end
     5.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     5.2 +++ b/src/HOL/Bali/Basis.thy	Mon Jan 28 17:00:19 2002 +0100
     5.3 @@ -0,0 +1,370 @@
     5.4 +(*  Title:      isabelle/Bali/Basis.thy
     5.5 +    ID:         $Id$
     5.6 +    Author:     David von Oheimb
     5.7 +    Copyright   1997 Technische Universitaet Muenchen
     5.8 +
     5.9 +*)
    5.10 +header {* Definitions extending HOL as logical basis of Bali *}
    5.11 +
    5.12 +theory Basis = Main:
    5.13 +
    5.14 +ML_setup {*
    5.15 +Unify.search_bound := 40;
    5.16 +Unify.trace_bound  := 40;
    5.17 +
    5.18 +quick_and_dirty:=true;
    5.19 +
    5.20 +Pretty.setmargin 77;
    5.21 +goals_limit:=2;
    5.22 +*}
    5.23 +(*print_depth 100;*)
    5.24 +(*Syntax.ambiguity_level := 1;*)
    5.25 +
    5.26 +section "misc"
    5.27 +
    5.28 +declare same_fstI [intro!] (*### TO HOL/Wellfounded_Relations *)
    5.29 +
    5.30 +(* ###TO HOL/???.ML?? *)
    5.31 +ML {*
    5.32 +fun make_simproc name pat pred thm = Simplifier.mk_simproc name
    5.33 +   [Thm.read_cterm (Thm.sign_of_thm thm) (pat, HOLogic.typeT)] 
    5.34 +   (K (K (fn s => if pred s then None else Some (standard (mk_meta_eq thm)))))
    5.35 +*}
    5.36 +
    5.37 +declare split_if_asm  [split] option.split [split] option.split_asm [split]
    5.38 +ML {*
    5.39 +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
    5.40 +*}
    5.41 +declare if_weak_cong [cong del] option.weak_case_cong [cong del]
    5.42 +declare length_Suc_conv [iff];
    5.43 +
    5.44 +(*###to be phased out *)
    5.45 +ML {*
    5.46 +bind_thm ("make_imp", rearrange_prems [1,0] mp)
    5.47 +*}
    5.48 +
    5.49 +lemma Collect_split_eq: "{p. P (split f p)} = {(a,b). P (f a b)}"
    5.50 +apply auto
    5.51 +done
    5.52 +
    5.53 +lemma subset_insertD: 
    5.54 +  "A <= insert x B ==> A <= B & x ~: A | (EX B'. A = insert x B' & B' <= B)"
    5.55 +apply (case_tac "x:A")
    5.56 +apply (rule disjI2)
    5.57 +apply (rule_tac x = "A-{x}" in exI)
    5.58 +apply fast+
    5.59 +done
    5.60 +
    5.61 +syntax
    5.62 +  "3" :: nat   ("3")
    5.63 +  "4" :: nat   ("4")
    5.64 +translations
    5.65 + "3" == "Suc 2"
    5.66 + "4" == "Suc 3"
    5.67 +
    5.68 +(*unused*)
    5.69 +lemma range_bool_domain: "range f = {f True, f False}"
    5.70 +apply auto
    5.71 +apply (case_tac "xa")
    5.72 +apply auto
    5.73 +done
    5.74 +
    5.75 +(* context (theory "Transitive_Closure") *)
    5.76 +lemma irrefl_tranclI': "r^-1 Int r^+ = {} ==> !x. (x, x) ~: r^+"
    5.77 +apply (rule allI)
    5.78 +apply (erule irrefl_tranclI)
    5.79 +done
    5.80 +
    5.81 +lemma trancl_rtrancl_trancl:
    5.82 +"\<lbrakk>(x,y)\<in>r^+; (y,z)\<in>r^*\<rbrakk> \<Longrightarrow> (x,z)\<in>r^+"
    5.83 +by (auto dest: tranclD rtrancl_trans rtrancl_into_trancl2)
    5.84 +
    5.85 +lemma rtrancl_into_trancl3:
    5.86 +"\<lbrakk>(a,b)\<in>r^*; a\<noteq>b\<rbrakk> \<Longrightarrow> (a,b)\<in>r^+"
    5.87 +apply (drule rtranclD)
    5.88 +apply auto
    5.89 +done
    5.90 +
    5.91 +lemma rtrancl_into_rtrancl2: 
    5.92 +  "\<lbrakk> (a, b) \<in>  r; (b, c) \<in> r^* \<rbrakk> \<Longrightarrow> (a, c) \<in>  r^*"
    5.93 +by (auto intro: r_into_rtrancl rtrancl_trans)
    5.94 +
    5.95 +lemma triangle_lemma:
    5.96 + "\<lbrakk> \<And> a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b=c; (a,x)\<in>r\<^sup>*; (a,y)\<in>r\<^sup>*\<rbrakk> 
    5.97 + \<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*"
    5.98 +proof -
    5.99 +  note converse_rtrancl_induct = converse_rtrancl_induct [consumes 1]
   5.100 +  note converse_rtranclE = converse_rtranclE [consumes 1] 
   5.101 +  assume unique: "\<And> a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b=c"
   5.102 +  assume "(a,x)\<in>r\<^sup>*" 
   5.103 +  then show "(a,y)\<in>r\<^sup>* \<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*"
   5.104 +  proof (induct rule: converse_rtrancl_induct)
   5.105 +    assume "(x,y)\<in>r\<^sup>*"
   5.106 +    then show ?thesis 
   5.107 +      by blast
   5.108 +  next
   5.109 +    fix a v
   5.110 +    assume a_v_r: "(a, v) \<in> r" and
   5.111 +          v_x_rt: "(v, x) \<in> r\<^sup>*" and
   5.112 +          a_y_rt: "(a, y) \<in> r\<^sup>*"  and
   5.113 +             hyp: "(v, y) \<in> r\<^sup>* \<Longrightarrow> (x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"
   5.114 +    from a_y_rt 
   5.115 +    show "(x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"
   5.116 +    proof (cases rule: converse_rtranclE)
   5.117 +      assume "a=y"
   5.118 +      with a_v_r v_x_rt have "(y,x) \<in> r\<^sup>*"
   5.119 +	by (auto intro: r_into_rtrancl rtrancl_trans)
   5.120 +      then show ?thesis 
   5.121 +	by blast
   5.122 +    next
   5.123 +      fix w 
   5.124 +      assume a_w_r: "(a, w) \<in> r" and
   5.125 +            w_y_rt: "(w, y) \<in> r\<^sup>*"
   5.126 +      from a_v_r a_w_r unique 
   5.127 +      have "v=w" 
   5.128 +	by auto
   5.129 +      with w_y_rt hyp 
   5.130 +      show ?thesis
   5.131 +	by blast
   5.132 +    qed
   5.133 +  qed
   5.134 +qed
   5.135 +
   5.136 +
   5.137 +lemma rtrancl_cases [consumes 1, case_names Refl Trancl]:
   5.138 + "\<lbrakk>(a,b)\<in>r\<^sup>*;  a = b \<Longrightarrow> P; (a,b)\<in>r\<^sup>+ \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
   5.139 +apply (erule rtranclE)
   5.140 +apply (auto dest: rtrancl_into_trancl1)
   5.141 +done
   5.142 +
   5.143 +(* ### To Transitive_Closure *)
   5.144 +theorems converse_rtrancl_induct 
   5.145 + = converse_rtrancl_induct [consumes 1,case_names Id Step]
   5.146 +
   5.147 +theorems converse_trancl_induct 
   5.148 +         = converse_trancl_induct [consumes 1,case_names Single Step]
   5.149 +
   5.150 +(* context (theory "Set") *)
   5.151 +lemma Ball_weaken:"\<lbrakk>Ball s P;\<And> x. P x\<longrightarrow>Q x\<rbrakk>\<Longrightarrow>Ball s Q"
   5.152 +by auto
   5.153 +
   5.154 +(* context (theory "Finite") *)
   5.155 +lemma finite_SetCompr2: "[| finite (Collect P); !y. P y --> finite (range (f y)) |] ==>  
   5.156 +  finite {f y x |x y. P y}"
   5.157 +apply (subgoal_tac "{f y x |x y. P y} = UNION (Collect P) (%y. range (f y))")
   5.158 +prefer 2 apply  fast
   5.159 +apply (erule ssubst)
   5.160 +apply (erule finite_UN_I)
   5.161 +apply fast
   5.162 +done
   5.163 +
   5.164 +
   5.165 +(* ### TO theory "List" *)
   5.166 +lemma list_all2_trans: "\<forall> a b c. P1 a b \<longrightarrow> P2 b c \<longrightarrow> P3 a c \<Longrightarrow>
   5.167 + \<forall>xs2 xs3. list_all2 P1 xs1 xs2 \<longrightarrow> list_all2 P2 xs2 xs3 \<longrightarrow> list_all2 P3 xs1 xs3"
   5.168 +apply (induct_tac "xs1")
   5.169 +apply simp
   5.170 +apply (rule allI)
   5.171 +apply (induct_tac "xs2")
   5.172 +apply simp
   5.173 +apply (rule allI)
   5.174 +apply (induct_tac "xs3")
   5.175 +apply auto
   5.176 +done
   5.177 +
   5.178 +
   5.179 +section "pairs"
   5.180 +
   5.181 +lemma surjective_pairing5: "p = (fst p, fst (snd p), fst (snd (snd p)), fst (snd (snd (snd p))), 
   5.182 +  snd (snd (snd (snd p))))"
   5.183 +apply auto
   5.184 +done
   5.185 +
   5.186 +lemma fst_splitE [elim!]: 
   5.187 +"[| fst s' = x';  !!x s. [| s' = (x,s);  x = x' |] ==> Q |] ==> Q"
   5.188 +apply (cut_tac p = "s'" in surjective_pairing)
   5.189 +apply auto
   5.190 +done
   5.191 +
   5.192 +lemma fst_in_set_lemma [rule_format (no_asm)]: "(x, y) : set l --> x : fst ` set l"
   5.193 +apply (induct_tac "l")
   5.194 +apply  auto
   5.195 +done
   5.196 +
   5.197 +
   5.198 +section "quantifiers"
   5.199 +
   5.200 +(*###to be phased out *)
   5.201 +ML {* 
   5.202 +fun noAll_simpset () = simpset() setmksimps 
   5.203 +	mksimps (filter (fn (x,_) => x<>"All") mksimps_pairs)
   5.204 +*}
   5.205 +
   5.206 +lemma All_Ex_refl_eq2 [simp]: 
   5.207 + "(!x. (? b. x = f b & Q b) \<longrightarrow> P x) = (!b. Q b --> P (f b))"
   5.208 +apply auto
   5.209 +done
   5.210 +
   5.211 +lemma ex_ex_miniscope1 [simp]:
   5.212 +  "(EX w v. P w v & Q v) = (EX v. (EX w. P w v) & Q v)"
   5.213 +apply auto
   5.214 +done
   5.215 +
   5.216 +lemma ex_miniscope2 [simp]:
   5.217 +  "(EX v. P v & Q & R v) = (Q & (EX v. P v & R v))" 
   5.218 +apply auto
   5.219 +done
   5.220 +
   5.221 +lemma ex_reorder31: "(\<exists>z x y. P x y z) = (\<exists>x y z. P x y z)"
   5.222 +apply auto
   5.223 +done
   5.224 +
   5.225 +lemma All_Ex_refl_eq1 [simp]: "(!x. (? b. x = f b) --> P x) = (!b. P (f b))"
   5.226 +apply auto
   5.227 +done
   5.228 +
   5.229 +
   5.230 +section "sums"
   5.231 +
   5.232 +hide const In0 In1
   5.233 +
   5.234 +syntax
   5.235 +  fun_sum :: "('a => 'c) => ('b => 'c) => (('a+'b) => 'c)" (infixr "'(+')"80)
   5.236 +translations
   5.237 + "fun_sum" == "sum_case"
   5.238 +
   5.239 +consts    the_Inl  :: "'a + 'b \<Rightarrow> 'a"
   5.240 +          the_Inr  :: "'a + 'b \<Rightarrow> 'b"
   5.241 +primrec  "the_Inl (Inl a) = a"
   5.242 +primrec  "the_Inr (Inr b) = b"
   5.243 +
   5.244 +datatype ('a, 'b, 'c) sum3 = In1 'a | In2 'b | In3 'c
   5.245 +
   5.246 +consts    the_In1  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'a"
   5.247 +          the_In2  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'b"
   5.248 +          the_In3  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'c"
   5.249 +primrec  "the_In1 (In1 a) = a"
   5.250 +primrec  "the_In2 (In2 b) = b"
   5.251 +primrec  "the_In3 (In3 c) = c"
   5.252 +
   5.253 +syntax
   5.254 +	 In1l	:: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
   5.255 +	 In1r	:: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
   5.256 +translations
   5.257 +	"In1l e" == "In1 (Inl e)"
   5.258 +	"In1r c" == "In1 (Inr c)"
   5.259 +
   5.260 +ML {*
   5.261 +fun sum3_instantiate thm = map (fn s => simplify(simpset()delsimps[not_None_eq])
   5.262 + (read_instantiate [("t","In"^s^" ?x")] thm)) ["1l","2","3","1r"]
   5.263 +*}
   5.264 +(* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *)
   5.265 +
   5.266 +translations
   5.267 +  "option"<= (type) "Option.option"
   5.268 +  "list"  <= (type) "List.list"
   5.269 +  "sum3"  <= (type) "Basis.sum3"
   5.270 +
   5.271 +
   5.272 +section "quantifiers for option type"
   5.273 +
   5.274 +syntax
   5.275 +  Oall :: "[pttrn, 'a option, bool] => bool"   ("(3! _:_:/ _)" [0,0,10] 10)
   5.276 +  Oex  :: "[pttrn, 'a option, bool] => bool"   ("(3? _:_:/ _)" [0,0,10] 10)
   5.277 +
   5.278 +syntax (symbols)
   5.279 +  Oall :: "[pttrn, 'a option, bool] => bool"   ("(3\<forall>_\<in>_:/ _)"  [0,0,10] 10)
   5.280 +  Oex  :: "[pttrn, 'a option, bool] => bool"   ("(3\<exists>_\<in>_:/ _)"  [0,0,10] 10)
   5.281 +
   5.282 +translations
   5.283 +  "! x:A: P"    == "! x:o2s A. P"
   5.284 +  "? x:A: P"    == "? x:o2s A. P"
   5.285 +
   5.286 +
   5.287 +section "unique association lists"
   5.288 +
   5.289 +constdefs
   5.290 +  unique   :: "('a \<times> 'b) list \<Rightarrow> bool"
   5.291 + "unique \<equiv> nodups \<circ> map fst"
   5.292 +
   5.293 +lemma uniqueD [rule_format (no_asm)]: 
   5.294 +"unique l--> (!x y. (x,y):set l --> (!x' y'. (x',y'):set l --> x=x'-->  y=y'))"
   5.295 +apply (unfold unique_def o_def)
   5.296 +apply (induct_tac "l")
   5.297 +apply  (auto dest: fst_in_set_lemma)
   5.298 +done
   5.299 +
   5.300 +lemma unique_Nil [simp]: "unique []"
   5.301 +apply (unfold unique_def)
   5.302 +apply (simp (no_asm))
   5.303 +done
   5.304 +
   5.305 +lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))"
   5.306 +apply (unfold unique_def)
   5.307 +apply  (auto dest: fst_in_set_lemma)
   5.308 +done
   5.309 +
   5.310 +lemmas unique_ConsI = conjI [THEN unique_Cons [THEN iffD2], standard]
   5.311 +
   5.312 +lemma unique_single [simp]: "!!p. unique [p]"
   5.313 +apply auto
   5.314 +done
   5.315 +
   5.316 +lemma unique_ConsD: "unique (x#xs) ==> unique xs"
   5.317 +apply (simp add: unique_def)
   5.318 +done
   5.319 +
   5.320 +lemma unique_append [rule_format (no_asm)]: "unique l' ==> unique l -->  
   5.321 +  (!(x,y):set l. !(x',y'):set l'. x' ~= x) --> unique (l @ l')"
   5.322 +apply (induct_tac "l")
   5.323 +apply  (auto dest: fst_in_set_lemma)
   5.324 +done
   5.325 +
   5.326 +lemma unique_map_inj [rule_format (no_asm)]: "unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)"
   5.327 +apply (induct_tac "l")
   5.328 +apply  (auto dest: fst_in_set_lemma simp add: inj_eq)
   5.329 +done
   5.330 +
   5.331 +lemma map_of_SomeI [rule_format (no_asm)]: "unique l --> (k, x) : set l --> map_of l k = Some x"
   5.332 +apply (induct_tac "l")
   5.333 +apply auto
   5.334 +done
   5.335 +
   5.336 +
   5.337 +section "list patterns"
   5.338 +
   5.339 +consts
   5.340 +  lsplit         :: "[['a, 'a list] => 'b, 'a list] => 'b"
   5.341 +defs
   5.342 +  lsplit_def:    "lsplit == %f l. f (hd l) (tl l)"
   5.343 +(*  list patterns -- extends pre-defined type "pttrn" used in abstractions *)
   5.344 +syntax
   5.345 +  "_lpttrn"    :: "[pttrn,pttrn] => pttrn"     ("_#/_" [901,900] 900)
   5.346 +translations
   5.347 +  "%y#x#xs. b"  == "lsplit (%y x#xs. b)"
   5.348 +  "%x#xs  . b"  == "lsplit (%x xs  . b)"
   5.349 +
   5.350 +lemma lsplit [simp]: "lsplit c (x#xs) = c x xs"
   5.351 +apply (unfold lsplit_def)
   5.352 +apply (simp (no_asm))
   5.353 +done
   5.354 +
   5.355 +lemma lsplit2 [simp]: "lsplit P (x#xs) y z = P x xs y z"
   5.356 +apply (unfold lsplit_def)
   5.357 +apply simp
   5.358 +done 
   5.359 +
   5.360 +
   5.361 +section "dummy pattern for quantifiers, let, etc."
   5.362 +
   5.363 +syntax
   5.364 +  "@dummy_pat"   :: pttrn    ("'_")
   5.365 +
   5.366 +parse_translation {*
   5.367 +let fun dummy_pat_tr [] = Free ("_",dummyT)
   5.368 +  | dummy_pat_tr ts = raise TERM ("dummy_pat_tr", ts);
   5.369 +in [("@dummy_pat", dummy_pat_tr)] 
   5.370 +end
   5.371 +*}
   5.372 +
   5.373 +end
     6.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     6.2 +++ b/src/HOL/Bali/Conform.thy	Mon Jan 28 17:00:19 2002 +0100
     6.3 @@ -0,0 +1,447 @@
     6.4 +(*  Title:      isabelle/Bali/Conform.thy
     6.5 +    ID:         $Id$
     6.6 +    Author:     David von Oheimb
     6.7 +    Copyright   1997 Technische Universitaet Muenchen
     6.8 +*)
     6.9 +
    6.10 +header {* Conformance notions for the type soundness proof for Java *}
    6.11 +
    6.12 +theory Conform = State:
    6.13 +
    6.14 +text {*
    6.15 +design issues:
    6.16 +\begin{itemize}
    6.17 +\item lconf allows for (arbitrary) inaccessible values
    6.18 +\item ''conforms'' does not directly imply that the dynamic types of all 
    6.19 +      objects on the heap are indeed existing classes. Yet this can be 
    6.20 +      inferred for all referenced objs.
    6.21 +\end{itemize}
    6.22 +*}
    6.23 +
    6.24 +types	env_ = "prog \<times> (lname, ty) table" (* same as env of WellType.thy *)
    6.25 +
    6.26 +
    6.27 +section "extension of global store"
    6.28 +
    6.29 +constdefs
    6.30 +
    6.31 +  gext    :: "st \<Rightarrow> st \<Rightarrow> bool"                ("_\<le>|_"       [71,71]   70)
    6.32 +   "s\<le>|s' \<equiv> \<forall>r. \<forall>obj\<in>globs s r: \<exists>obj'\<in>globs s' r: tag obj'= tag obj"
    6.33 +
    6.34 +lemma gext_objD: 
    6.35 +"\<lbrakk>s\<le>|s'; globs s r = Some obj\<rbrakk> 
    6.36 +\<Longrightarrow> \<exists>obj'. globs s' r = Some obj' \<and> tag obj' = tag obj"
    6.37 +apply (simp only: gext_def)
    6.38 +by force
    6.39 +
    6.40 +lemma rev_gext_objD: 
    6.41 +"\<lbrakk>globs s r = Some obj; s\<le>|s'\<rbrakk> 
    6.42 + \<Longrightarrow> \<exists>obj'. globs s' r = Some obj' \<and> tag obj' = tag obj"
    6.43 +by (auto elim: gext_objD)
    6.44 +
    6.45 +lemma init_class_obj_inited: 
    6.46 +   "init_class_obj G C s1\<le>|s2 \<Longrightarrow> inited C (globs s2)"
    6.47 +apply (unfold inited_def init_obj_def)
    6.48 +apply (auto dest!: gext_objD)
    6.49 +done
    6.50 +
    6.51 +lemma gext_refl [intro!, simp]: "s\<le>|s"
    6.52 +apply (unfold gext_def)
    6.53 +apply (fast del: fst_splitE)
    6.54 +done
    6.55 +
    6.56 +lemma gext_gupd [simp, elim!]: "\<And>s. globs s r = None \<Longrightarrow> s\<le>|gupd(r\<mapsto>x)s"
    6.57 +by (auto simp: gext_def)
    6.58 +
    6.59 +lemma gext_new [simp, elim!]: "\<And>s. globs s r = None \<Longrightarrow> s\<le>|init_obj G oi r s"
    6.60 +apply (simp only: init_obj_def)
    6.61 +apply (erule_tac gext_gupd)
    6.62 +done
    6.63 +
    6.64 +lemma gext_trans [elim]: "\<And>X. \<lbrakk>s\<le>|s'; s'\<le>|s''\<rbrakk> \<Longrightarrow> s\<le>|s''" 
    6.65 +by (force simp: gext_def)
    6.66 +
    6.67 +lemma gext_upd_gobj [intro!]: "s\<le>|upd_gobj r n v s"
    6.68 +apply (simp only: gext_def)
    6.69 +apply auto
    6.70 +apply (case_tac "ra = r")
    6.71 +apply auto
    6.72 +apply (case_tac "globs s r = None")
    6.73 +apply auto
    6.74 +done
    6.75 +
    6.76 +lemma gext_cong1 [simp]: "set_locals l s1\<le>|s2 = s1\<le>|s2"
    6.77 +by (auto simp: gext_def)
    6.78 +
    6.79 +lemma gext_cong2 [simp]: "s1\<le>|set_locals l s2 = s1\<le>|s2"
    6.80 +by (auto simp: gext_def)
    6.81 +
    6.82 +
    6.83 +lemma gext_lupd1 [simp]: "lupd(vn\<mapsto>v)s1\<le>|s2 = s1\<le>|s2"
    6.84 +by (auto simp: gext_def)
    6.85 +
    6.86 +lemma gext_lupd2 [simp]: "s1\<le>|lupd(vn\<mapsto>v)s2 = s1\<le>|s2"
    6.87 +by (auto simp: gext_def)
    6.88 +
    6.89 +
    6.90 +lemma inited_gext: "\<lbrakk>inited C (globs s); s\<le>|s'\<rbrakk> \<Longrightarrow> inited C (globs s')"
    6.91 +apply (unfold inited_def)
    6.92 +apply (auto dest: gext_objD)
    6.93 +done
    6.94 +
    6.95 +
    6.96 +section "value conformance"
    6.97 +
    6.98 +constdefs
    6.99 +
   6.100 +  conf  :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool"    ("_,_\<turnstile>_\<Colon>\<preceq>_"   [71,71,71,71] 70)
   6.101 +	   "G,s\<turnstile>v\<Colon>\<preceq>T \<equiv> \<exists>T'\<in>typeof (\<lambda>a. option_map obj_ty (heap s a)) v:G\<turnstile>T'\<preceq>T"
   6.102 +
   6.103 +lemma conf_cong [simp]: "G,set_locals l s\<turnstile>v\<Colon>\<preceq>T = G,s\<turnstile>v\<Colon>\<preceq>T"
   6.104 +by (auto simp: conf_def)
   6.105 +
   6.106 +lemma conf_lupd [simp]: "G,lupd(vn\<mapsto>va)s\<turnstile>v\<Colon>\<preceq>T = G,s\<turnstile>v\<Colon>\<preceq>T"
   6.107 +by (auto simp: conf_def)
   6.108 +
   6.109 +lemma conf_PrimT [simp]: "\<forall>dt. typeof dt v = Some (PrimT t) \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>PrimT t"
   6.110 +apply (simp add: conf_def)
   6.111 +done
   6.112 +
   6.113 +lemma conf_litval [rule_format (no_asm)]: 
   6.114 +  "typeof (\<lambda>a. None) v = Some T \<longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T"
   6.115 +apply (unfold conf_def)
   6.116 +apply (rule val.induct)
   6.117 +apply auto
   6.118 +done
   6.119 +
   6.120 +lemma conf_Null [simp]: "G,s\<turnstile>Null\<Colon>\<preceq>T = G\<turnstile>NT\<preceq>T"
   6.121 +by (simp add: conf_def)
   6.122 +
   6.123 +lemma conf_Addr: 
   6.124 +  "G,s\<turnstile>Addr a\<Colon>\<preceq>T = (\<exists>obj. heap s a = Some obj \<and> G\<turnstile>obj_ty obj\<preceq>T)"
   6.125 +by (auto simp: conf_def)
   6.126 +
   6.127 +lemma conf_AddrI:"\<lbrakk>heap s a = Some obj; G\<turnstile>obj_ty obj\<preceq>T\<rbrakk> \<Longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq>T"
   6.128 +apply (rule conf_Addr [THEN iffD2])
   6.129 +by fast
   6.130 +
   6.131 +lemma defval_conf [rule_format (no_asm), elim]: 
   6.132 +  "is_type G T \<longrightarrow> G,s\<turnstile>default_val T\<Colon>\<preceq>T"
   6.133 +apply (unfold conf_def)
   6.134 +apply (induct "T")
   6.135 +apply (auto intro: prim_ty.induct)
   6.136 +done
   6.137 +
   6.138 +lemma conf_widen [rule_format (no_asm), elim]: 
   6.139 +  "G\<turnstile>T\<preceq>T' \<Longrightarrow> G,s\<turnstile>x\<Colon>\<preceq>T \<longrightarrow> ws_prog G \<longrightarrow> G,s\<turnstile>x\<Colon>\<preceq>T'"
   6.140 +apply (unfold conf_def)
   6.141 +apply (rule val.induct)
   6.142 +apply (auto elim: ws_widen_trans)
   6.143 +done
   6.144 +
   6.145 +lemma conf_gext [rule_format (no_asm), elim]: 
   6.146 +  "G,s\<turnstile>v\<Colon>\<preceq>T \<longrightarrow> s\<le>|s' \<longrightarrow> G,s'\<turnstile>v\<Colon>\<preceq>T"
   6.147 +apply (unfold gext_def conf_def)
   6.148 +apply (rule val.induct)
   6.149 +apply force+
   6.150 +done
   6.151 +
   6.152 +
   6.153 +lemma conf_list_widen [rule_format (no_asm)]: 
   6.154 +"ws_prog G \<Longrightarrow>  
   6.155 +  \<forall>Ts Ts'. list_all2 (conf G s) vs Ts 
   6.156 +           \<longrightarrow>   G\<turnstile>Ts[\<preceq>] Ts' \<longrightarrow> list_all2 (conf G s) vs Ts'"
   6.157 +apply (unfold widens_def)
   6.158 +apply (rule list_all2_trans)
   6.159 +apply auto
   6.160 +done
   6.161 +
   6.162 +lemma conf_RefTD [rule_format (no_asm)]: 
   6.163 + "G,s\<turnstile>a'\<Colon>\<preceq>RefT T 
   6.164 +  \<longrightarrow> a' = Null \<or> (\<exists>a obj T'. a' = Addr a \<and> heap s a = Some obj \<and>  
   6.165 +                    obj_ty obj = T' \<and> G\<turnstile>T'\<preceq>RefT T)"
   6.166 +apply (unfold conf_def)
   6.167 +apply (induct_tac "a'")
   6.168 +apply (auto dest: widen_PrimT)
   6.169 +done
   6.170 +
   6.171 +
   6.172 +section "value list conformance"
   6.173 +
   6.174 +constdefs
   6.175 +
   6.176 +  lconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool"
   6.177 +                                                ("_,_\<turnstile>_[\<Colon>\<preceq>]_" [71,71,71,71] 70)
   6.178 +           "G,s\<turnstile>vs[\<Colon>\<preceq>]Ts \<equiv> \<forall>n. \<forall>T\<in>Ts n: \<exists>v\<in>vs n: G,s\<turnstile>v\<Colon>\<preceq>T"
   6.179 +
   6.180 +lemma lconfD: "\<lbrakk>G,s\<turnstile>vs[\<Colon>\<preceq>]Ts; Ts n = Some T\<rbrakk> \<Longrightarrow> G,s\<turnstile>(the (vs n))\<Colon>\<preceq>T"
   6.181 +by (force simp: lconf_def)
   6.182 +
   6.183 +
   6.184 +lemma lconf_cong [simp]: "\<And>s. G,set_locals x s\<turnstile>l[\<Colon>\<preceq>]L = G,s\<turnstile>l[\<Colon>\<preceq>]L"
   6.185 +by (auto simp: lconf_def)
   6.186 +
   6.187 +lemma lconf_lupd [simp]: "G,lupd(vn\<mapsto>v)s\<turnstile>l[\<Colon>\<preceq>]L = G,s\<turnstile>l[\<Colon>\<preceq>]L"
   6.188 +by (auto simp: lconf_def)
   6.189 +
   6.190 +(* unused *)
   6.191 +lemma lconf_new: "\<lbrakk>L vn = None; G,s\<turnstile>l[\<Colon>\<preceq>]L\<rbrakk> \<Longrightarrow> G,s\<turnstile>l(vn\<mapsto>v)[\<Colon>\<preceq>]L"
   6.192 +by (auto simp: lconf_def)
   6.193 +
   6.194 +lemma lconf_upd: "\<lbrakk>G,s\<turnstile>l[\<Colon>\<preceq>]L; G,s\<turnstile>v\<Colon>\<preceq>T; L vn = Some T\<rbrakk> \<Longrightarrow>  
   6.195 +  G,s\<turnstile>l(vn\<mapsto>v)[\<Colon>\<preceq>]L"
   6.196 +by (auto simp: lconf_def)
   6.197 +
   6.198 +lemma lconf_ext: "\<lbrakk>G,s\<turnstile>l[\<Colon>\<preceq>]L; G,s\<turnstile>v\<Colon>\<preceq>T\<rbrakk> \<Longrightarrow> G,s\<turnstile>l(vn\<mapsto>v)[\<Colon>\<preceq>]L(vn\<mapsto>T)"
   6.199 +by (auto simp: lconf_def)
   6.200 +
   6.201 +lemma lconf_map_sum [simp]: 
   6.202 + "G,s\<turnstile>l1 (+) l2[\<Colon>\<preceq>]L1 (+) L2 = (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>l2[\<Colon>\<preceq>]L2)"
   6.203 +apply (unfold lconf_def)
   6.204 +apply safe
   6.205 +apply (case_tac [3] "n")
   6.206 +apply (force split add: sum.split)+
   6.207 +done
   6.208 +
   6.209 +lemma lconf_ext_list [rule_format (no_asm)]: "
   6.210 + \<And>X. \<lbrakk>G,s\<turnstile>l[\<Colon>\<preceq>]L\<rbrakk> \<Longrightarrow> 
   6.211 +      \<forall>vs Ts. nodups vns \<longrightarrow> length Ts = length vns 
   6.212 +      \<longrightarrow> list_all2 (conf G s) vs Ts \<longrightarrow> G,s\<turnstile>l(vns[\<mapsto>]vs)[\<Colon>\<preceq>]L(vns[\<mapsto>]Ts)"
   6.213 +apply (unfold lconf_def)
   6.214 +apply (induct_tac "vns")
   6.215 +apply  clarsimp
   6.216 +apply clarsimp
   6.217 +apply (frule list_all2_lengthD)
   6.218 +apply clarsimp
   6.219 +done
   6.220 +
   6.221 +
   6.222 +lemma lconf_deallocL: "\<lbrakk>G,s\<turnstile>l[\<Colon>\<preceq>]L(vn\<mapsto>T); L vn = None\<rbrakk> \<Longrightarrow> G,s\<turnstile>l[\<Colon>\<preceq>]L"
   6.223 +apply (simp only: lconf_def)
   6.224 +apply safe
   6.225 +apply (drule spec)
   6.226 +apply (drule ospec)
   6.227 +apply auto
   6.228 +done 
   6.229 +
   6.230 +
   6.231 +lemma lconf_gext [elim]: "\<lbrakk>G,s\<turnstile>l[\<Colon>\<preceq>]L; s\<le>|s'\<rbrakk> \<Longrightarrow> G,s'\<turnstile>l[\<Colon>\<preceq>]L"
   6.232 +apply (simp only: lconf_def)
   6.233 +apply fast
   6.234 +done
   6.235 +
   6.236 +lemma lconf_empty [simp, intro!]: "G,s\<turnstile>vs[\<Colon>\<preceq>]empty"
   6.237 +apply (unfold lconf_def)
   6.238 +apply force
   6.239 +done
   6.240 +
   6.241 +lemma lconf_init_vals [intro!]: 
   6.242 +	" \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<Colon>\<preceq>]fs"
   6.243 +apply (unfold lconf_def)
   6.244 +apply force
   6.245 +done
   6.246 +
   6.247 +
   6.248 +section "object conformance"
   6.249 +
   6.250 +constdefs
   6.251 +
   6.252 +  oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool"  ("_,_\<turnstile>_\<Colon>\<preceq>\<surd>_"  [71,71,71,71] 70)
   6.253 +	   "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r \<equiv> G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and> 
   6.254 +                           (case r of 
   6.255 +		              Heap a \<Rightarrow> is_type G (obj_ty obj) 
   6.256 +                            | Stat C \<Rightarrow> True)"
   6.257 +(*
   6.258 +lemma oconf_def2:  "G,s\<turnstile>\<lparr>tag=oi,values=fs\<rparr>\<Colon>\<preceq>\<surd>r =  
   6.259 +  (G,s\<turnstile>fs[\<Colon>\<preceq>]var_tys G oi r \<and> 
   6.260 +  (case r of Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=fs\<rparr>) | Stat C \<Rightarrow> True))"
   6.261 +by (simp add: oconf_def Let_def)
   6.262 +*)
   6.263 +(*
   6.264 +lemma oconf_def2:  "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r =  
   6.265 +  (G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and> 
   6.266 +  (case r of Heap a \<Rightarrow> is_type G (obj_ty obj) | Stat C \<Rightarrow> True))"
   6.267 +by (simp add: oconf_def Let_def)
   6.268 +*)
   6.269 +lemma oconf_is_type: "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>Heap a \<Longrightarrow> is_type G (obj_ty obj)"
   6.270 +by (auto simp: oconf_def Let_def)
   6.271 +
   6.272 +lemma oconf_lconf: "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r \<Longrightarrow> G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r"
   6.273 +by (simp add: oconf_def) 
   6.274 +
   6.275 +lemma oconf_cong [simp]: "G,set_locals l s\<turnstile>obj\<Colon>\<preceq>\<surd>r = G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r"
   6.276 +by (auto simp: oconf_def Let_def)
   6.277 +
   6.278 +lemma oconf_init_obj_lemma: 
   6.279 +"\<lbrakk>\<And>C c. class G C = Some c \<Longrightarrow> unique (DeclConcepts.fields G C);  
   6.280 +  \<And>C c f fld. \<lbrakk>class G C = Some c; 
   6.281 +                table_of (DeclConcepts.fields G C) f = Some fld \<rbrakk> 
   6.282 +            \<Longrightarrow> is_type G (type fld);  
   6.283 +  (case r of 
   6.284 +     Heap a \<Rightarrow> is_type G (obj_ty obj) 
   6.285 +  | Stat C \<Rightarrow> is_class G C)
   6.286 +\<rbrakk> \<Longrightarrow>  G,s\<turnstile>obj \<lparr>values:=init_vals (var_tys G (tag obj) r)\<rparr>\<Colon>\<preceq>\<surd>r"
   6.287 +apply (auto simp add: oconf_def)
   6.288 +apply (drule_tac var_tys_Some_eq [THEN iffD1]) 
   6.289 +defer
   6.290 +apply (subst obj_ty_cong)
   6.291 +apply(auto dest!: fields_table_SomeD obj_ty_CInst1 obj_ty_Arr1
   6.292 +           split add: sum.split_asm obj_tag.split_asm)
   6.293 +done
   6.294 +
   6.295 +(*
   6.296 +lemma oconf_init_obj_lemma: 
   6.297 +"\<lbrakk>\<And>C c. class G C = Some c \<Longrightarrow> unique (fields G C);  
   6.298 +  \<And>C c f fld. \<lbrakk>class G C = Some c; table_of (fields G C) f = Some fld \<rbrakk> 
   6.299 +            \<Longrightarrow> is_type G (type fld);  
   6.300 +  (case r of 
   6.301 +     Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=fs\<rparr>) 
   6.302 +  | Stat C \<Rightarrow> is_class G C)
   6.303 +\<rbrakk> \<Longrightarrow>  G,s\<turnstile>\<lparr>tag=oi, values=init_vals (var_tys G oi r)\<rparr>\<Colon>\<preceq>\<surd>r"
   6.304 +apply (auto simp add: oconf_def)
   6.305 +apply (drule_tac var_tys_Some_eq [THEN iffD1]) 
   6.306 +defer
   6.307 +apply (subst obj_ty_eq)
   6.308 +apply(auto dest!: fields_table_SomeD split add: sum.split_asm obj_tag.split_asm)
   6.309 +done
   6.310 +*)
   6.311 +
   6.312 +
   6.313 +section "state conformance"
   6.314 +
   6.315 +constdefs
   6.316 +
   6.317 +  conforms :: "state \<Rightarrow> env_ \<Rightarrow> bool"          (     "_\<Colon>\<preceq>_"   [71,71]      70)
   6.318 +   "xs\<Colon>\<preceq>E \<equiv> let (G, L) = E; s = snd xs; l = locals s in
   6.319 +      (\<forall>r. \<forall>obj\<in>globs s r:           G,s\<turnstile>obj   \<Colon>\<preceq>\<surd>r) \<and>
   6.320 +                  \<spacespace>                   G,s\<turnstile>l    [\<Colon>\<preceq>]L\<spacespace> \<and>
   6.321 +      (\<forall>a. fst xs=Some(Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable))"
   6.322 +
   6.323 +section "conforms"
   6.324 +
   6.325 +lemma conforms_globsD: 
   6.326 +"\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); globs s r = Some obj\<rbrakk> \<Longrightarrow> G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r"
   6.327 +by (auto simp: conforms_def Let_def)
   6.328 +
   6.329 +lemma conforms_localD: "(x, s)\<Colon>\<preceq>(G, L) \<Longrightarrow> G,s\<turnstile>locals s[\<Colon>\<preceq>]L"
   6.330 +by (auto simp: conforms_def Let_def)
   6.331 +
   6.332 +lemma conforms_XcptLocD: "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); x = Some (Xcpt (Loc a))\<rbrakk> \<Longrightarrow>  
   6.333 +	  G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)"
   6.334 +by (auto simp: conforms_def Let_def)
   6.335 +
   6.336 +lemma conforms_RefTD: 
   6.337 + "\<lbrakk>G,s\<turnstile>a'\<Colon>\<preceq>RefT t; a' \<noteq> Null; (x,s) \<Colon>\<preceq>(G, L)\<rbrakk> \<Longrightarrow>  
   6.338 +   \<exists>a obj. a' = Addr a \<and> globs s (Inl a) = Some obj \<and>  
   6.339 +   G\<turnstile>obj_ty obj\<preceq>RefT t \<and> is_type G (obj_ty obj)"
   6.340 +apply (drule_tac conf_RefTD)
   6.341 +apply clarsimp
   6.342 +apply (rule conforms_globsD [THEN oconf_is_type])
   6.343 +apply auto
   6.344 +done
   6.345 +
   6.346 +lemma conforms_Jump [iff]:
   6.347 +  "((Some (Jump j), s)\<Colon>\<preceq>(G, L)) = (Norm s\<Colon>\<preceq>(G, L))"
   6.348 +by (auto simp: conforms_def)
   6.349 +
   6.350 +lemma conforms_StdXcpt [iff]: 
   6.351 +  "((Some (Xcpt (Std xn)), s)\<Colon>\<preceq>(G, L)) = (Norm s\<Colon>\<preceq>(G, L))"
   6.352 +by (auto simp: conforms_def)
   6.353 +
   6.354 +lemma conforms_raise_if [iff]: 
   6.355 +  "((raise_if c xn x, s)\<Colon>\<preceq>(G, L)) = ((x, s)\<Colon>\<preceq>(G, L))"
   6.356 +by (auto simp: abrupt_if_def)
   6.357 +
   6.358 +
   6.359 +lemma conforms_NormI: "(x, s)\<Colon>\<preceq>(G, L) \<Longrightarrow> Norm s\<Colon>\<preceq>(G, L)"
   6.360 +by (auto simp: conforms_def Let_def)
   6.361 +
   6.362 +
   6.363 +lemma conforms_absorb [rule_format]:
   6.364 +  "(a, b)\<Colon>\<preceq>(G, L) \<longrightarrow> (absorb j a, b)\<Colon>\<preceq>(G, L)"
   6.365 +apply (rule impI)
   6.366 +apply ( case_tac a)
   6.367 +apply (case_tac "absorb j a")
   6.368 +apply auto
   6.369 +apply (case_tac "absorb j (Some a)",auto)
   6.370 +apply (erule conforms_NormI)
   6.371 +done
   6.372 +
   6.373 +lemma conformsI: "\<lbrakk>\<forall>r. \<forall>obj\<in>globs s r: G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r;  
   6.374 +     G,s\<turnstile>locals s[\<Colon>\<preceq>]L;  
   6.375 +     \<forall>a. x = Some (Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)\<rbrakk> \<Longrightarrow> 
   6.376 +  (x, s)\<Colon>\<preceq>(G, L)"
   6.377 +by (auto simp: conforms_def Let_def)
   6.378 +
   6.379 +lemma conforms_xconf: "\<lbrakk>(x, s)\<Colon>\<preceq>(G,L);   
   6.380 + \<forall>a. x' = Some (Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)\<rbrakk> \<Longrightarrow> 
   6.381 + (x',s)\<Colon>\<preceq>(G,L)"
   6.382 +by (fast intro: conformsI elim: conforms_globsD conforms_localD)
   6.383 +
   6.384 +lemma conforms_lupd: 
   6.385 + "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); L vn = Some T; G,s\<turnstile>v\<Colon>\<preceq>T\<rbrakk> \<Longrightarrow> (x, lupd(vn\<mapsto>v)s)\<Colon>\<preceq>(G, L)"
   6.386 +by (force intro: conformsI lconf_upd dest: conforms_globsD conforms_localD 
   6.387 +                                           conforms_XcptLocD simp: oconf_def)
   6.388 +
   6.389 +
   6.390 +lemmas conforms_allocL_aux = conforms_localD [THEN lconf_ext]
   6.391 +
   6.392 +lemma conforms_allocL: 
   6.393 +  "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); G,s\<turnstile>v\<Colon>\<preceq>T\<rbrakk> \<Longrightarrow> (x, lupd(vn\<mapsto>v)s)\<Colon>\<preceq>(G, L(vn\<mapsto>T))"
   6.394 +by (force intro: conformsI dest: conforms_globsD 
   6.395 +          elim: conforms_XcptLocD conforms_allocL_aux simp: oconf_def)
   6.396 +
   6.397 +lemmas conforms_deallocL_aux = conforms_localD [THEN lconf_deallocL]
   6.398 +
   6.399 +lemma conforms_deallocL: "\<And>s.\<lbrakk>s\<Colon>\<preceq>(G, L(vn\<mapsto>T)); L vn = None\<rbrakk> \<Longrightarrow> s\<Colon>\<preceq>(G,L)"
   6.400 +by (fast intro: conformsI dest: conforms_globsD 
   6.401 +         elim: conforms_XcptLocD conforms_deallocL_aux)
   6.402 +
   6.403 +lemma conforms_gext: "\<lbrakk>(x, s)\<Colon>\<preceq>(G,L); s\<le>|s';  
   6.404 +  \<forall>r. \<forall>obj\<in>globs s' r: G,s'\<turnstile>obj\<Colon>\<preceq>\<surd>r;  
   6.405 +   locals s'=locals s\<rbrakk> \<Longrightarrow> (x,s')\<Colon>\<preceq>(G,L)"
   6.406 +by (force intro!: conformsI dest: conforms_localD conforms_XcptLocD)
   6.407 +
   6.408 +
   6.409 +lemma conforms_xgext: 
   6.410 +  "\<lbrakk>(x ,s)\<Colon>\<preceq>(G,L); (x', s')\<Colon>\<preceq>(G, L); s'\<le>|s\<rbrakk> \<Longrightarrow> (x',s)\<Colon>\<preceq>(G,L)"
   6.411 +apply (erule_tac conforms_xconf)
   6.412 +apply (fast dest: conforms_XcptLocD)
   6.413 +done
   6.414 +
   6.415 +lemma conforms_gupd: "\<And>obj. \<lbrakk>(x, s)\<Colon>\<preceq>(G, L); G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r; s\<le>|gupd(r\<mapsto>obj)s\<rbrakk> 
   6.416 +\<Longrightarrow>  (x, gupd(r\<mapsto>obj)s)\<Colon>\<preceq>(G, L)"
   6.417 +apply (rule conforms_gext)
   6.418 +apply    auto
   6.419 +apply (force dest: conforms_globsD simp add: oconf_def)+
   6.420 +done
   6.421 +
   6.422 +lemma conforms_upd_gobj: "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L); globs s r = Some obj; 
   6.423 +  var_tys G (tag obj) r n = Some T; G,s\<turnstile>v\<Colon>\<preceq>T\<rbrakk> \<Longrightarrow> (x,upd_gobj r n v s)\<Colon>\<preceq>(G,L)"
   6.424 +apply (rule conforms_gext)
   6.425 +apply auto
   6.426 +apply (drule (1) conforms_globsD)
   6.427 +apply (simp add: oconf_def)
   6.428 +apply safe
   6.429 +apply (rule lconf_upd)
   6.430 +apply auto
   6.431 +apply (simp only: obj_ty_cong) 
   6.432 +apply (force dest: conforms_globsD intro!: lconf_upd 
   6.433 +       simp add: oconf_def cong del: sum.weak_case_cong)
   6.434 +done
   6.435 +
   6.436 +lemma conforms_set_locals: 
   6.437 +  "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L'); G,s\<turnstile>l[\<Colon>\<preceq>]L\<rbrakk> \<Longrightarrow> (x,set_locals l s)\<Colon>\<preceq>(G,L)"
   6.438 +apply (auto intro!: conformsI dest: conforms_globsD 
   6.439 +            elim!: conforms_XcptLocD simp add: oconf_def)
   6.440 +done
   6.441 +
   6.442 +lemma conforms_return: "\<And>s'. \<lbrakk>(x,s)\<Colon>\<preceq>(G, L); (x',s')\<Colon>\<preceq>(G, L'); s\<le>|s'\<rbrakk> \<Longrightarrow>  
   6.443 +  (x',set_locals (locals s) s')\<Colon>\<preceq>(G, L)"
   6.444 +apply (rule conforms_xconf)
   6.445 +prefer 2 apply (force dest: conforms_XcptLocD)
   6.446 +apply (erule conforms_gext)
   6.447 +apply (force dest: conforms_globsD)+
   6.448 +done
   6.449 +
   6.450 +end
     7.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     7.2 +++ b/src/HOL/Bali/Decl.thy	Mon Jan 28 17:00:19 2002 +0100
     7.3 @@ -0,0 +1,867 @@
     7.4 +(*  Title:      isabelle/Bali/Decl.thy
     7.5 +    ID:         $Id$
     7.6 +    Author:     David von Oheimb
     7.7 +    Copyright   1997 Technische Universitaet Muenchen
     7.8 +*)
     7.9 +header {* Field, method, interface, and class declarations, whole Java programs
    7.10 +*}
    7.11 +
    7.12 +(** order is significant, because of clash for "var" **)
    7.13 +theory Decl = Term + Table:
    7.14 +
    7.15 +text {*
    7.16 +improvements:
    7.17 +\begin{itemize}
    7.18 +\item clarification and correction of some aspects of the package/access concept
    7.19 +  (Also submitted as bug report to the Java Bug Database:
    7.20 +   Bug Id: 4485402 and Bug Id: 4493343 
    7.21 +   http://developer.java.sun.com/developer/bugParade/index.jshtml
    7.22 +  )
    7.23 +\end{itemize}
    7.24 +simplifications:
    7.25 +\begin{itemize}
    7.26 +\item the only field and method modifiers are static and the access modifiers
    7.27 +\item no constructors, which may be simulated by new + suitable methods
    7.28 +\item there is just one global initializer per class, which can simulate all 
    7.29 +      others
    7.30 +
    7.31 +\item no throws clause
    7.32 +\item a void method is replaced by one that returns Unit (of dummy type Void)
    7.33 +
    7.34 +\item no interface fields
    7.35 +
    7.36 +\item every class has an explicit superclass (unused for Object)
    7.37 +\item the (standard) methods of Object and of standard exceptions are not 
    7.38 +      specified
    7.39 +
    7.40 +\item no main method
    7.41 +\end{itemize}
    7.42 +*}
    7.43 +
    7.44 +subsection {* Modifier*}
    7.45 +
    7.46 +subsubsection {* Access modifier *}
    7.47 +
    7.48 +datatype acc_modi (* access modifier *)
    7.49 +         = Private | Package | Protected | Public 
    7.50 +
    7.51 +text {* 
    7.52 +We can define a linear order for the access modifiers. With Private yielding the
    7.53 +most restrictive access and public the most liberal access policy:
    7.54 +  Private < Package < Protected < Public
    7.55 +*}
    7.56 + 
    7.57 +instance acc_modi:: ord
    7.58 +by intro_classes
    7.59 +
    7.60 +defs (overloaded)
    7.61 +less_acc_def: 
    7.62 + "a < (b::acc_modi) 
    7.63 +      \<equiv> (case a of
    7.64 +             Private    \<Rightarrow> (b=Package \<or> b=Protected \<or> b=Public)
    7.65 +           | Package    \<Rightarrow> (b=Protected \<or> b=Public)
    7.66 +           | Protected  \<Rightarrow> (b=Public)
    7.67 +           | Public     \<Rightarrow> False)"
    7.68 +le_acc_def:
    7.69 + "a \<le> (b::acc_modi) \<equiv> (a = b) \<or> (a < b)"
    7.70 +
    7.71 +instance acc_modi:: order
    7.72 +proof (intro_classes)
    7.73 +  fix x y z::acc_modi
    7.74 +  {
    7.75 +  show "x \<le> x"               \<spacespace>\<spacespace>    -- reflexivity
    7.76 +    by (auto simp add: le_acc_def)
    7.77 +  next
    7.78 +  assume "x \<le> y" "y \<le> z"           -- transitivity 
    7.79 +  thus "x \<le> z"
    7.80 +    by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split)
    7.81 +  next
    7.82 +  assume "x \<le> y" "y \<le> x"           -- antisymmetry
    7.83 +  thus "x = y"
    7.84 +  proof -
    7.85 +    have "\<forall> x y. x < (y::acc_modi) \<and> y < x \<longrightarrow> False"
    7.86 +      by (auto simp add: less_acc_def split add: acc_modi.split)
    7.87 +    with prems show ?thesis
    7.88 +      by  (auto simp add: le_acc_def)
    7.89 +  qed
    7.90 +  next
    7.91 +  show "(x < y) = (x \<le> y \<and> x \<noteq> y)"
    7.92 +    by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split) 
    7.93 +  }
    7.94 +qed
    7.95 +  
    7.96 +instance acc_modi:: linorder
    7.97 +proof intro_classes
    7.98 +  fix x y:: acc_modi
    7.99 +  show  "x \<le> y \<or> y \<le> x"   
   7.100 +  by (auto simp add: less_acc_def le_acc_def split add: acc_modi.split)
   7.101 +qed
   7.102 +
   7.103 +lemma acc_modi_top [simp]: "Public \<le> a \<Longrightarrow> a = Public"
   7.104 +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
   7.105 +
   7.106 +lemma acc_modi_top1 [simp, intro!]: "a \<le> Public"
   7.107 +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
   7.108 +
   7.109 +lemma acc_modi_le_Public: 
   7.110 +"a \<le> Public \<Longrightarrow> a=Private \<or> a = Package \<or> a=Protected \<or> a=Public"
   7.111 +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
   7.112 +
   7.113 +lemma acc_modi_bottom: "a \<le> Private \<Longrightarrow> a = Private"
   7.114 +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
   7.115 +
   7.116 +lemma acc_modi_Private_le: 
   7.117 +"Private \<le> a \<Longrightarrow> a=Private \<or> a = Package \<or> a=Protected \<or> a=Public"
   7.118 +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
   7.119 +
   7.120 +lemma acc_modi_Package_le: 
   7.121 +  "Package \<le> a \<Longrightarrow> a = Package \<or> a=Protected \<or> a=Public"
   7.122 +by (auto simp add: le_acc_def less_acc_def split: acc_modi.split)
   7.123 +
   7.124 +lemma acc_modi_le_Package: 
   7.125 +  "a \<le> Package \<Longrightarrow> a=Private \<or> a = Package"
   7.126 +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
   7.127 +
   7.128 +lemma acc_modi_Protected_le: 
   7.129 +  "Protected \<le> a \<Longrightarrow> a=Protected \<or> a=Public"
   7.130 +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
   7.131 +
   7.132 +lemma acc_modi_le_Protected: 
   7.133 +  "a \<le> Protected  \<Longrightarrow> a=Private \<or> a = Package \<or> a = Protected"
   7.134 +by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
   7.135 +
   7.136 +
   7.137 +lemmas acc_modi_le_Dests = acc_modi_top           acc_modi_le_Public
   7.138 +                           acc_modi_Private_le    acc_modi_bottom
   7.139 +                           acc_modi_Package_le    acc_modi_le_Package
   7.140 +                           acc_modi_Protected_le  acc_modi_le_Protected
   7.141 +
   7.142 +lemma acc_modi_Package_le_cases 
   7.143 + [consumes 1,case_names Package Protected Public]:
   7.144 + "Package \<le> m \<Longrightarrow> ( m = Package \<Longrightarrow> P m) \<Longrightarrow> (m=Protected \<Longrightarrow> P m) \<Longrightarrow> 
   7.145 +   (m=Public \<Longrightarrow> P m) \<Longrightarrow> P m"
   7.146 +by (auto dest: acc_modi_Package_le)
   7.147 +
   7.148 +
   7.149 +subsubsection {* Static Modifier *}
   7.150 +types stat_modi = bool (* modifier: static *)
   7.151 +
   7.152 +subsection {* Declaration (base "class" for member,interface and class
   7.153 + declarations *}
   7.154 +
   7.155 +record decl =
   7.156 +        access :: acc_modi
   7.157 +
   7.158 +translations
   7.159 +  "decl" <= (type) "\<lparr>access::acc_modi\<rparr>"
   7.160 +  "decl" <= (type) "\<lparr>access::acc_modi,\<dots>::'a\<rparr>"
   7.161 +
   7.162 +subsection {* Member (field or method)*}
   7.163 +record  member = decl +
   7.164 +         static :: stat_modi
   7.165 +
   7.166 +translations
   7.167 +  "member" <= (type) "\<lparr>access::acc_modi,static::bool\<rparr>"
   7.168 +  "member" <= (type) "\<lparr>access::acc_modi,static::bool,\<dots>::'a\<rparr>"
   7.169 +
   7.170 +subsection {* Field *}
   7.171 +
   7.172 +record field = member +
   7.173 +        type :: ty
   7.174 +translations
   7.175 +  "field" <= (type) "\<lparr>access::acc_modi, static::bool, type::ty\<rparr>"
   7.176 +  "field" <= (type) "\<lparr>access::acc_modi, static::bool, type::ty,\<dots>::'a\<rparr>"
   7.177 +
   7.178 +types     
   7.179 +        fdecl           (* field declaration, cf. 8.3 *)
   7.180 +	= "vname \<times> field"
   7.181 +
   7.182 +
   7.183 +translations
   7.184 +  "fdecl" <= (type) "vname \<times> field"
   7.185 +
   7.186 +subsection  {* Method *}
   7.187 +
   7.188 +record mhead = member +     (* method head (excluding signature) *)
   7.189 +        pars ::"vname list" (* parameter names *)
   7.190 +        resT ::ty           (* result type *)
   7.191 +
   7.192 +record mbody =                      (* method body *)
   7.193 +        lcls::  "(vname \<times> ty) list" (* local variables *)
   7.194 +        stmt:: stmt                 (* the body statement *)
   7.195 +
   7.196 +record methd = mhead + (* method in a class *)
   7.197 +        mbody::mbody
   7.198 +
   7.199 +types mdecl = "sig \<times> methd"  (* method declaration in a class *)
   7.200 +
   7.201 +
   7.202 +translations
   7.203 +  "mhead" <= (type) "\<lparr>access::acc_modi, static::bool, 
   7.204 +                      pars::vname list, resT::ty\<rparr>"
   7.205 +  "mhead" <= (type) "\<lparr>access::acc_modi, static::bool, 
   7.206 +                      pars::vname list, resT::ty,\<dots>::'a\<rparr>"
   7.207 +  "mbody" <= (type) "\<lparr>lcls::(vname \<times> ty) list,stmt::stmt\<rparr>"
   7.208 +  "mbody" <= (type) "\<lparr>lcls::(vname \<times> ty) list,stmt::stmt,\<dots>::'a\<rparr>"      
   7.209 +  "methd" <= (type) "\<lparr>access::acc_modi, static::bool, 
   7.210 +                      pars::vname list, resT::ty,mbody::mbody\<rparr>"
   7.211 +  "methd" <= (type) "\<lparr>access::acc_modi, static::bool, 
   7.212 +                      pars::vname list, resT::ty,mbody::mbody,\<dots>::'a\<rparr>"
   7.213 +  "mdecl" <= (type) "sig \<times> methd"
   7.214 +
   7.215 +
   7.216 +constdefs 
   7.217 +  mhead::"methd \<Rightarrow> mhead"
   7.218 +  "mhead m \<equiv> \<lparr>access=access m, static=static m, pars=pars m, resT=resT m\<rparr>"
   7.219 +
   7.220 +lemma access_mhead [simp]:"access (mhead m) = access m"
   7.221 +by (simp add: mhead_def)
   7.222 +
   7.223 +lemma static_mhead [simp]:"static (mhead m) = static m"
   7.224 +by (simp add: mhead_def)
   7.225 +
   7.226 +lemma pars_mhead [simp]:"pars (mhead m) = pars m"
   7.227 +by (simp add: mhead_def)
   7.228 +
   7.229 +lemma resT_mhead [simp]:"resT (mhead m) = resT m"
   7.230 +by (simp add: mhead_def)
   7.231 +
   7.232 +text {* To be able to talk uniformaly about field and method declarations we
   7.233 +introduce the notion of a member declaration (e.g. useful to define 
   7.234 +accessiblity ) *}
   7.235 +
   7.236 +datatype memberdecl = fdecl fdecl | mdecl mdecl
   7.237 +
   7.238 +datatype memberid = fid vname | mid sig
   7.239 +
   7.240 +axclass has_memberid < "type"
   7.241 +consts
   7.242 + memberid :: "'a::has_memberid \<Rightarrow> memberid"
   7.243 +
   7.244 +instance memberdecl::has_memberid
   7.245 +by (intro_classes)
   7.246 +
   7.247 +defs (overloaded)
   7.248 +memberdecl_memberid_def:
   7.249 +  "memberid m \<equiv> (case m of
   7.250 +                    fdecl (vn,f)  \<Rightarrow> fid vn
   7.251 +                  | mdecl (sig,m) \<Rightarrow> mid sig)"
   7.252 +
   7.253 +lemma memberid_fdecl_simp[simp]: "memberid (fdecl (vn,f)) = fid vn"
   7.254 +by (simp add: memberdecl_memberid_def)
   7.255 +
   7.256 +lemma memberid_fdecl_simp1: "memberid (fdecl f) = fid (fst f)"
   7.257 +by (cases f) (simp add: memberdecl_memberid_def)
   7.258 +
   7.259 +lemma memberid_mdecl_simp[simp]: "memberid (mdecl (sig,m)) = mid sig"
   7.260 +by (simp add: memberdecl_memberid_def)
   7.261 +
   7.262 +lemma memberid_mdecl_simp1: "memberid (mdecl m) = mid (fst m)"
   7.263 +by (cases m) (simp add: memberdecl_memberid_def)
   7.264 +
   7.265 +instance * :: ("type","has_memberid") has_memberid
   7.266 +by (intro_classes)
   7.267 +
   7.268 +defs (overloaded)
   7.269 +pair_memberid_def:
   7.270 +  "memberid p \<equiv> memberid (snd p)"
   7.271 +
   7.272 +lemma memberid_pair_simp[simp]: "memberid (c,m) = memberid m"
   7.273 +by (simp add: pair_memberid_def)
   7.274 +
   7.275 +lemma memberid_pair_simp1: "memberid p  = memberid (snd p)"
   7.276 +by (simp add: pair_memberid_def)
   7.277 +
   7.278 +constdefs is_field :: "qtname \<times> memberdecl \<Rightarrow> bool"
   7.279 +"is_field m \<equiv> \<exists> declC f. m=(declC,fdecl f)"
   7.280 +  
   7.281 +lemma is_fieldD: "is_field m \<Longrightarrow> \<exists> declC f. m=(declC,fdecl f)"
   7.282 +by (simp add: is_field_def)
   7.283 +
   7.284 +lemma is_fieldI: "is_field (C,fdecl f)"
   7.285 +by (simp add: is_field_def)
   7.286 +
   7.287 +constdefs is_method :: "qtname \<times> memberdecl \<Rightarrow> bool"
   7.288 +"is_method membr \<equiv> \<exists> declC m. membr=(declC,mdecl m)"
   7.289 +  
   7.290 +lemma is_methodD: "is_method membr \<Longrightarrow> \<exists> declC m. membr=(declC,mdecl m)"
   7.291 +by (simp add: is_method_def)
   7.292 +
   7.293 +lemma is_methodI: "is_method (C,mdecl m)"
   7.294 +by (simp add: is_method_def)
   7.295 +
   7.296 +
   7.297 +subsection {* Interface *}
   7.298 +
   7.299 +
   7.300 +record  ibody = decl +  (* interface body *)
   7.301 +          imethods :: "(sig \<times> mhead) list" (* method heads *)
   7.302 +
   7.303 +record  iface = ibody + (* interface *)
   7.304 +         isuperIfs:: "qtname list" (* superinterface list *)
   7.305 +types	
   7.306 +	idecl           (* interface declaration, cf. 9.1 *)
   7.307 +	= "qtname \<times> iface"
   7.308 +
   7.309 +translations
   7.310 +  "ibody" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list\<rparr>"
   7.311 +  "ibody" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list,\<dots>::'a\<rparr>"
   7.312 +  "iface" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list,
   7.313 +                      isuperIfs::qtname list\<rparr>"
   7.314 +  "iface" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list,
   7.315 +                      isuperIfs::qtname list,\<dots>::'a\<rparr>"
   7.316 +  "idecl" <= (type) "qtname \<times> iface"
   7.317 +
   7.318 +constdefs
   7.319 +  ibody :: "iface \<Rightarrow> ibody"
   7.320 +  "ibody i \<equiv> \<lparr>access=access i,imethods=imethods i\<rparr>"
   7.321 +
   7.322 +lemma access_ibody [simp]: "(access (ibody i)) = access i"
   7.323 +by (simp add: ibody_def)
   7.324 +
   7.325 +lemma imethods_ibody [simp]: "(imethods (ibody i)) = imethods i"
   7.326 +by (simp add: ibody_def)
   7.327 +
   7.328 +subsection  {* Class *}
   7.329 +record cbody = decl +          (* class body *)
   7.330 +         cfields:: "fdecl list" 
   7.331 +         methods:: "mdecl list"
   7.332 +         init   :: "stmt"       (* initializer *)
   7.333 +
   7.334 +record class = cbody +           (* class *)
   7.335 +        super   :: "qtname"      (* superclass *)
   7.336 +        superIfs:: "qtname list" (* implemented interfaces *)
   7.337 +types	
   7.338 +	cdecl           (* class declaration, cf. 8.1 *)
   7.339 +	= "qtname \<times> class"
   7.340 +
   7.341 +translations
   7.342 +  "cbody" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
   7.343 +                      methods::mdecl list,init::stmt\<rparr>"
   7.344 +  "cbody" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
   7.345 +                      methods::mdecl list,init::stmt,\<dots>::'a\<rparr>"
   7.346 +  "class" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
   7.347 +                      methods::mdecl list,init::stmt,
   7.348 +                      super::qtname,superIfs::qtname list\<rparr>"
   7.349 +  "class" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
   7.350 +                      methods::mdecl list,init::stmt,
   7.351 +                      super::qtname,superIfs::qtname list,\<dots>::'a\<rparr>"
   7.352 +  "cdecl" <= (type) "qtname \<times> class"
   7.353 +
   7.354 +constdefs
   7.355 +  cbody :: "class \<Rightarrow> cbody"
   7.356 +  "cbody c \<equiv> \<lparr>access=access c, cfields=cfields c,methods=methods c,init=init c\<rparr>"
   7.357 +
   7.358 +lemma access_cbody [simp]:"access (cbody c) = access c"
   7.359 +by (simp add: cbody_def)
   7.360 +
   7.361 +lemma cfields_cbody [simp]:"cfields (cbody c) = cfields c"
   7.362 +by (simp add: cbody_def)
   7.363 +
   7.364 +lemma methods_cbody [simp]:"methods (cbody c) = methods c"
   7.365 +by (simp add: cbody_def)
   7.366 +
   7.367 +lemma init_cbody [simp]:"init (cbody c) = init c"
   7.368 +by (simp add: cbody_def)
   7.369 +
   7.370 +
   7.371 +section "standard classes"
   7.372 +
   7.373 +consts
   7.374 +
   7.375 +  Object_mdecls  ::  "mdecl list" (* methods of Object *)
   7.376 +  SXcpt_mdecls   ::  "mdecl list" (* methods of SXcpts *)
   7.377 +  ObjectC ::         "cdecl"      (* declaration  of root      class   *)
   7.378 +  SXcptC  ::"xname \<Rightarrow> cdecl"      (* declarations of throwable classes *)
   7.379 +
   7.380 +defs 
   7.381 +
   7.382 +
   7.383 +ObjectC_def:"ObjectC  \<equiv> (Object,\<lparr>access=Public,cfields=[],methods=Object_mdecls,
   7.384 +                                  init=Skip,super=arbitrary,superIfs=[]\<rparr>)"
   7.385 +SXcptC_def:"SXcptC xn\<equiv> (SXcpt xn,\<lparr>access=Public,cfields=[],methods=SXcpt_mdecls,
   7.386 +                                   init=Skip,
   7.387 +                                   super=if xn = Throwable then Object 
   7.388 +                                                           else SXcpt Throwable,
   7.389 +                                   superIfs=[]\<rparr>)"
   7.390 +
   7.391 +lemma ObjectC_neq_SXcptC [simp]: "ObjectC \<noteq> SXcptC xn"
   7.392 +by (simp add: ObjectC_def SXcptC_def Object_def SXcpt_def)
   7.393 +
   7.394 +lemma SXcptC_inject [simp]: "(SXcptC xn = SXcptC xm) = (xn = xm)"
   7.395 +apply (simp add: SXcptC_def)
   7.396 +apply auto
   7.397 +done
   7.398 +
   7.399 +constdefs standard_classes :: "cdecl list"
   7.400 +         "standard_classes \<equiv> [ObjectC, SXcptC Throwable,
   7.401 +		SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast,
   7.402 +		SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]"
   7.403 +
   7.404 +
   7.405 +section "programs"
   7.406 +
   7.407 +record prog =
   7.408 +        ifaces ::"idecl list"
   7.409 +        "classes"::"cdecl list"
   7.410 +
   7.411 +translations
   7.412 +     "prog"<= (type) "\<lparr>ifaces::idecl list,classes::cdecl list\<rparr>"
   7.413 +     "prog"<= (type) "\<lparr>ifaces::idecl list,classes::cdecl list,\<dots>::'a\<rparr>"
   7.414 +
   7.415 +syntax
   7.416 +  iface     :: "prog  \<Rightarrow> (qtname, iface) table"
   7.417 +  class     :: "prog  \<Rightarrow> (qtname, class) table"
   7.418 +  is_iface  :: "prog  \<Rightarrow> qtname  \<Rightarrow> bool"
   7.419 +  is_class  :: "prog  \<Rightarrow> qtname  \<Rightarrow> bool"
   7.420 +
   7.421 +translations
   7.422 +	   "iface G I" == "table_of (ifaces G) I"
   7.423 +	   "class G C" == "table_of (classes G) C"
   7.424 +	"is_iface G I" == "iface G I \<noteq> None"
   7.425 +	"is_class G C" == "class G C \<noteq> None"
   7.426 +
   7.427 +
   7.428 +section "is type"
   7.429 +
   7.430 +consts
   7.431 +  is_type :: "prog \<Rightarrow>     ty \<Rightarrow> bool"
   7.432 +  isrtype :: "prog \<Rightarrow> ref_ty \<Rightarrow> bool"
   7.433 +
   7.434 +primrec	"is_type G (PrimT pt)  = True"
   7.435 +	"is_type G (RefT  rt)  = isrtype G rt"
   7.436 +	"isrtype G (NullT    ) = True"
   7.437 +	"isrtype G (IfaceT tn) = is_iface G tn"
   7.438 +	"isrtype G (ClassT tn) = is_class G tn"
   7.439 +	"isrtype G (ArrayT T ) = is_type  G T"
   7.440 +
   7.441 +lemma type_is_iface: "is_type G (Iface I) \<Longrightarrow> is_iface G I"
   7.442 +by auto
   7.443 +
   7.444 +lemma type_is_class: "is_type G (Class C) \<Longrightarrow>  is_class G C"
   7.445 +by auto
   7.446 +
   7.447 +
   7.448 +section "subinterface and subclass relation, in anticipation of TypeRel.thy"
   7.449 +
   7.450 +consts 
   7.451 +  subint1  :: "prog \<Rightarrow> (qtname \<times> qtname) set"
   7.452 +  subcls1  :: "prog \<Rightarrow> (qtname \<times> qtname) set"
   7.453 +
   7.454 +defs
   7.455 +  subint1_def: "subint1 G \<equiv> {(I,J). \<exists>i\<in>iface G I: J\<in>set (isuperIfs i)}"
   7.456 +  subcls1_def: "subcls1 G \<equiv> {(C,D). C\<noteq>Object \<and> (\<exists>c\<in>class G C: super c = D)}"
   7.457 +
   7.458 +syntax
   7.459 + "@subcls1" :: "prog => [qtname, qtname] => bool" ("_|-_<:C1_" [71,71,71] 70)
   7.460 + "@subclseq":: "prog => [qtname, qtname] => bool" ("_|-_<=:C _"[71,71,71] 70)
   7.461 + "@subcls"  :: "prog => [qtname, qtname] => bool" ("_|-_<:C _"[71,71,71] 70)
   7.462 +
   7.463 +syntax (xsymbols)
   7.464 +  "@subcls1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>\<^sub>C\<^sub>1_"  [71,71,71] 70)
   7.465 +  "@subclseq":: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>\<^sub>C _"  [71,71,71] 70)
   7.466 +  "@subcls"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>\<^sub>C _"  [71,71,71] 70)
   7.467 +
   7.468 +translations
   7.469 +        "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" == "(C,D) \<in> subcls1 G"
   7.470 +	"G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" (* cf. 8.1.3 *)
   7.471 +        "G\<turnstile>C \<prec>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^+"
   7.472 + 
   7.473 +
   7.474 +lemma subint1I: "\<lbrakk>iface G I = Some i; J \<in> set (isuperIfs i)\<rbrakk> 
   7.475 +                 \<Longrightarrow> (I,J) \<in> subint1 G" 
   7.476 +apply (simp add: subint1_def)
   7.477 +done
   7.478 +
   7.479 +lemma subcls1I:"\<lbrakk>class G C = Some c; C \<noteq> Object\<rbrakk> \<Longrightarrow> (C,(super c)) \<in> subcls1 G"
   7.480 +apply (simp add: subcls1_def)
   7.481 +done
   7.482 +
   7.483 +
   7.484 +lemma subint1D: "(I,J)\<in>subint1 G\<Longrightarrow> \<exists>i\<in>iface G I: J\<in>set (isuperIfs i)"
   7.485 +by (simp add: subint1_def)
   7.486 +
   7.487 +lemma subcls1D: 
   7.488 +  "(C,D)\<in>subcls1 G \<Longrightarrow> C\<noteq>Object \<and> (\<exists>c. class G C = Some c \<and> (super c = D))"
   7.489 +apply (simp add: subcls1_def)
   7.490 +apply auto
   7.491 +done
   7.492 +
   7.493 +lemma subint1_def2:  
   7.494 +   "subint1 G = (\<Sigma> I\<in>{I. is_iface G I}. set (isuperIfs (the (iface G I))))"
   7.495 +apply (unfold subint1_def)
   7.496 +apply auto
   7.497 +done
   7.498 +
   7.499 +lemma subcls1_def2: 
   7.500 + "subcls1 G = (\<Sigma>C\<in>{C. is_class G C}. {D. C\<noteq>Object \<and> super (the(class G C))=D})"
   7.501 +apply (unfold subcls1_def)
   7.502 +apply auto
   7.503 +done
   7.504 +
   7.505 +lemma subcls_is_class:
   7.506 +"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D\<rbrakk> \<Longrightarrow> \<exists> c. class G C = Some c"
   7.507 +by (auto simp add: subcls1_def dest: tranclD)
   7.508 +
   7.509 +lemma no_subcls1_Object:"G\<turnstile>Object\<prec>\<^sub>C\<^sub>1 D \<Longrightarrow> P"
   7.510 +by (auto simp add: subcls1_def)
   7.511 +
   7.512 +lemma no_subcls_Object: "G\<turnstile>Object\<prec>\<^sub>C D \<Longrightarrow> P"
   7.513 +apply (erule trancl_induct)
   7.514 +apply (auto intro: no_subcls1_Object)
   7.515 +done
   7.516 +
   7.517 +section "well-structured programs"
   7.518 +
   7.519 +constdefs
   7.520 +  ws_idecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname list \<Rightarrow> bool"
   7.521 + "ws_idecl G I si \<equiv> \<forall>J\<in>set si.  is_iface G J   \<and> (J,I)\<notin>(subint1 G)^+"
   7.522 +  
   7.523 +  ws_cdecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   7.524 + "ws_cdecl G C sc \<equiv> C\<noteq>Object \<longrightarrow> is_class G sc \<and> (sc,C)\<notin>(subcls1 G)^+"
   7.525 +  
   7.526 +  ws_prog  :: "prog \<Rightarrow> bool"
   7.527 + "ws_prog G \<equiv> (\<forall>(I,i)\<in>set (ifaces  G). ws_idecl G I (isuperIfs i)) \<and> 
   7.528 +	      (\<forall>(C,c)\<in>set (classes G). ws_cdecl G C (super c))"
   7.529 +
   7.530 +
   7.531 +lemma ws_progI: 
   7.532 +"\<lbrakk>\<forall>(I,i)\<in>set (ifaces G). \<forall>J\<in>set (isuperIfs i). is_iface G J \<and> 
   7.533 +                                                (J,I) \<notin> (subint1 G)^+; 
   7.534 +  \<forall>(C,c)\<in>set (classes G). C\<noteq>Object \<longrightarrow> is_class G (super c) \<and> 
   7.535 +                                        ((super c),C) \<notin> (subcls1 G)^+  
   7.536 + \<rbrakk> \<Longrightarrow> ws_prog G"
   7.537 +apply (unfold ws_prog_def ws_idecl_def ws_cdecl_def)
   7.538 +apply (erule_tac conjI)
   7.539 +apply blast
   7.540 +done
   7.541 +
   7.542 +lemma ws_prog_ideclD: 
   7.543 +"\<lbrakk>iface G I = Some i; J\<in>set (isuperIfs i); ws_prog G\<rbrakk> \<Longrightarrow>  
   7.544 +  is_iface G J \<and> (J,I)\<notin>(subint1 G)^+"
   7.545 +apply (unfold ws_prog_def ws_idecl_def)
   7.546 +apply clarify
   7.547 +apply (drule_tac map_of_SomeD)
   7.548 +apply auto
   7.549 +done
   7.550 +
   7.551 +lemma ws_prog_cdeclD: 
   7.552 +"\<lbrakk>class G C = Some c; C\<noteq>Object; ws_prog G\<rbrakk> \<Longrightarrow>  
   7.553 +  is_class G (super c) \<and> (super c,C)\<notin>(subcls1 G)^+"
   7.554 +apply (unfold ws_prog_def ws_cdecl_def)
   7.555 +apply clarify
   7.556 +apply (drule_tac map_of_SomeD)
   7.557 +apply auto
   7.558 +done
   7.559 +
   7.560 +
   7.561 +section "well-foundedness"
   7.562 +
   7.563 +lemma finite_is_iface: "finite {I. is_iface G I}"
   7.564 +apply (fold dom_def)
   7.565 +apply (rule_tac finite_dom_map_of)
   7.566 +done
   7.567 +
   7.568 +lemma finite_is_class: "finite {C. is_class G C}"
   7.569 +apply (fold dom_def)
   7.570 +apply (rule_tac finite_dom_map_of)
   7.571 +done
   7.572 +
   7.573 +lemma finite_subint1: "finite (subint1 G)"
   7.574 +apply (subst subint1_def2)
   7.575 +apply (rule finite_SigmaI)
   7.576 +apply (rule finite_is_iface)
   7.577 +apply (simp (no_asm))
   7.578 +done
   7.579 +
   7.580 +lemma finite_subcls1: "finite (subcls1 G)"
   7.581 +apply (subst subcls1_def2)
   7.582 +apply (rule finite_SigmaI)
   7.583 +apply (rule finite_is_class)
   7.584 +apply (rule_tac B = "{super (the (class G C))}" in finite_subset)
   7.585 +apply  auto
   7.586 +done
   7.587 +
   7.588 +lemma subint1_irrefl_lemma1: 
   7.589 +  "ws_prog G \<Longrightarrow> (subint1 G)^-1 \<inter> (subint1 G)^+ = {}"
   7.590 +apply (force dest: subint1D ws_prog_ideclD conjunct2)
   7.591 +done
   7.592 +
   7.593 +lemma subcls1_irrefl_lemma1: 
   7.594 +  "ws_prog G \<Longrightarrow> (subcls1 G)^-1 \<inter> (subcls1 G)^+ = {}"
   7.595 +apply (force dest: subcls1D ws_prog_cdeclD conjunct2)
   7.596 +done
   7.597 +
   7.598 +lemmas subint1_irrefl_lemma2 = subint1_irrefl_lemma1 [THEN irrefl_tranclI']
   7.599 +lemmas subcls1_irrefl_lemma2 = subcls1_irrefl_lemma1 [THEN irrefl_tranclI']
   7.600 +
   7.601 +lemma subint1_irrefl: "\<lbrakk>(x, y) \<in> subint1 G; ws_prog G\<rbrakk> \<Longrightarrow> x \<noteq> y"
   7.602 +apply (rule irrefl_trancl_rD)
   7.603 +apply (rule subint1_irrefl_lemma2)
   7.604 +apply auto
   7.605 +done
   7.606 +
   7.607 +lemma subcls1_irrefl: "\<lbrakk>(x, y) \<in> subcls1 G; ws_prog G\<rbrakk> \<Longrightarrow> x \<noteq> y"
   7.608 +apply (rule irrefl_trancl_rD)
   7.609 +apply (rule subcls1_irrefl_lemma2)
   7.610 +apply auto
   7.611 +done
   7.612 +
   7.613 +lemmas subint1_acyclic = subint1_irrefl_lemma2 [THEN acyclicI, standard]
   7.614 +lemmas subcls1_acyclic = subcls1_irrefl_lemma2 [THEN acyclicI, standard]
   7.615 +
   7.616 +
   7.617 +lemma wf_subint1: "ws_prog G \<Longrightarrow> wf ((subint1 G)\<inverse>)"
   7.618 +by (auto intro: finite_acyclic_wf_converse finite_subint1 subint1_acyclic)
   7.619 +
   7.620 +lemma wf_subcls1: "ws_prog G \<Longrightarrow> wf ((subcls1 G)\<inverse>)"
   7.621 +by (auto intro: finite_acyclic_wf_converse finite_subcls1 subcls1_acyclic)
   7.622 +
   7.623 +
   7.624 +lemma subint1_induct: 
   7.625 +  "\<lbrakk>ws_prog G; \<And>x. \<forall>y. (x, y) \<in> subint1 G \<longrightarrow> P y \<Longrightarrow> P x\<rbrakk> \<Longrightarrow> P a"
   7.626 +apply (frule wf_subint1)
   7.627 +apply (erule wf_induct)
   7.628 +apply (simp (no_asm_use) only: converse_iff)
   7.629 +apply blast
   7.630 +done
   7.631 +
   7.632 +lemma subcls1_induct [consumes 1]:
   7.633 +  "\<lbrakk>ws_prog G; \<And>x. \<forall>y. (x, y) \<in> subcls1 G \<longrightarrow> P y \<Longrightarrow> P x\<rbrakk> \<Longrightarrow> P a"
   7.634 +apply (frule wf_subcls1)
   7.635 +apply (erule wf_induct)
   7.636 +apply (simp (no_asm_use) only: converse_iff)
   7.637 +apply blast
   7.638 +done
   7.639 +
   7.640 +lemma ws_subint1_induct: 
   7.641 + "\<lbrakk>is_iface G I; ws_prog G; \<And>I i. \<lbrakk>iface G I = Some i \<and> 
   7.642 +   (\<forall>J \<in> set (isuperIfs i). (I,J)\<in>subint1 G \<and> P J \<and> is_iface G J)\<rbrakk> \<Longrightarrow> P I
   7.643 +  \<rbrakk> \<Longrightarrow> P I"
   7.644 +apply (erule make_imp)
   7.645 +apply (rule subint1_induct)
   7.646 +apply  assumption
   7.647 +apply safe
   7.648 +apply (fast dest: subint1I ws_prog_ideclD)
   7.649 +done
   7.650 +
   7.651 +
   7.652 +lemma ws_subcls1_induct: "\<lbrakk>is_class G C; ws_prog G;  
   7.653 +  \<And>C c. \<lbrakk>class G C = Some c;  
   7.654 + (C \<noteq> Object \<longrightarrow> (C,(super c))\<in>subcls1 G \<and> 
   7.655 +                  P (super c) \<and> is_class G (super c))\<rbrakk> \<Longrightarrow> P C
   7.656 + \<rbrakk> \<Longrightarrow> P C"
   7.657 +apply (erule make_imp)
   7.658 +apply (rule subcls1_induct)
   7.659 +apply  assumption
   7.660 +apply safe
   7.661 +apply (fast dest: subcls1I ws_prog_cdeclD)
   7.662 +done
   7.663 +
   7.664 +lemma ws_class_induct [consumes 2, case_names Object Subcls]:
   7.665 +"\<lbrakk>class G C = Some c; ws_prog G; 
   7.666 +  \<And> co. class G Object = Some co \<Longrightarrow> P Object; 
   7.667 +  \<And>  C c. \<lbrakk>class G C = Some c; C \<noteq> Object; P (super c)\<rbrakk> \<Longrightarrow> P C
   7.668 + \<rbrakk> \<Longrightarrow> P C"
   7.669 +proof -
   7.670 +  assume clsC: "class G C = Some c"
   7.671 +  and    init: "\<And> co. class G Object = Some co \<Longrightarrow> P Object"
   7.672 +  and    step: "\<And>   C c. \<lbrakk>class G C = Some c; C \<noteq> Object; P (super c)\<rbrakk> \<Longrightarrow> P C"
   7.673 +  assume ws: "ws_prog G"
   7.674 +  then have "is_class G C \<Longrightarrow> P C"  
   7.675 +  proof (induct rule: subcls1_induct)
   7.676 +    fix C
   7.677 +    assume   hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<longrightarrow> is_class G S \<longrightarrow> P S"
   7.678 +       and iscls:"is_class G C"
   7.679 +    show "P C"
   7.680 +    proof (cases "C=Object")
   7.681 +      case True with iscls init show "P C" by auto
   7.682 +    next
   7.683 +      case False with ws step hyp iscls 
   7.684 +      show "P C" by (auto dest: subcls1I ws_prog_cdeclD)
   7.685 +    qed
   7.686 +  qed
   7.687 +  with clsC show ?thesis by simp
   7.688 +qed
   7.689 +
   7.690 +lemma ws_class_induct' [consumes 2, case_names Object Subcls]:
   7.691 +"\<lbrakk>is_class G C; ws_prog G; 
   7.692 +  \<And> co. class G Object = Some co \<Longrightarrow> P Object; 
   7.693 +  \<And> C c. \<lbrakk>class G C = Some c; C \<noteq> Object; P (super c)\<rbrakk> \<Longrightarrow> P C
   7.694 + \<rbrakk> \<Longrightarrow> P C"
   7.695 +by (blast intro: ws_class_induct)
   7.696 +
   7.697 +lemma ws_class_induct'' [consumes 2, case_names Object Subcls]:
   7.698 +"\<lbrakk>class G C = Some c; ws_prog G; 
   7.699 +  \<And> co. class G Object = Some co \<Longrightarrow> P Object co; 
   7.700 +  \<And>  C c sc. \<lbrakk>class G C = Some c; class G (super c) = Some sc;
   7.701 +            C \<noteq> Object; P (super c) sc\<rbrakk> \<Longrightarrow> P C c
   7.702 + \<rbrakk> \<Longrightarrow> P C c"
   7.703 +proof -
   7.704 +  assume clsC: "class G C = Some c"
   7.705 +  and    init: "\<And> co. class G Object = Some co \<Longrightarrow> P Object co"
   7.706 +  and    step: "\<And> C c sc . \<lbrakk>class G C = Some c; class G (super c) = Some sc;
   7.707 +                             C \<noteq> Object; P (super c) sc\<rbrakk> \<Longrightarrow> P C c"
   7.708 +  assume ws: "ws_prog G"
   7.709 +  then have "\<And> c. class G C = Some c\<Longrightarrow> P C c"  
   7.710 +  proof (induct rule: subcls1_induct)
   7.711 +    fix C c
   7.712 +    assume   hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<longrightarrow> (\<forall> s. class G S = Some s \<longrightarrow> P S s)"
   7.713 +       and iscls:"class G C = Some c"
   7.714 +    show "P C c"
   7.715 +    proof (cases "C=Object")
   7.716 +      case True with iscls init show "P C c" by auto
   7.717 +    next
   7.718 +      case False
   7.719 +      with ws iscls obtain sc where
   7.720 +	sc: "class G (super c) = Some sc"
   7.721 +	by (auto dest: ws_prog_cdeclD)
   7.722 +      from iscls False have "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 (super c)" by (rule subcls1I)
   7.723 +      with False ws step hyp iscls sc
   7.724 +      show "P C c" 
   7.725 +	by (auto)  
   7.726 +    qed
   7.727 +  qed
   7.728 +  with clsC show "P C c" by auto
   7.729 +qed
   7.730 +
   7.731 +lemma ws_interface_induct [consumes 2, case_names Step]:
   7.732 + (assumes is_if_I: "is_iface G I" and 
   7.733 +               ws: "ws_prog G" and
   7.734 +          hyp_sub: "\<And>I i. \<lbrakk>iface G I = Some i; 
   7.735 +                            \<forall> J \<in> set (isuperIfs i).
   7.736 +                                 (I,J)\<in>subint1 G \<and> P J \<and> is_iface G J\<rbrakk> \<Longrightarrow> P I"
   7.737 + ) "P I"
   7.738 +proof -
   7.739 +  from is_if_I ws 
   7.740 +  show "P I"
   7.741 +  proof (rule ws_subint1_induct)
   7.742 +    fix I i
   7.743 +    assume hyp: "iface G I = Some i \<and>
   7.744 +                (\<forall>J\<in>set (isuperIfs i). (I,J) \<in>subint1 G \<and> P J \<and> is_iface G J)"
   7.745 +    then have if_I: "iface G I = Some i"
   7.746 +      by blast
   7.747 +    show "P I"
   7.748 +    proof (cases "isuperIfs i")
   7.749 +      case Nil
   7.750 +      with if_I hyp_sub 
   7.751 +      show "P I" 
   7.752 +	by auto
   7.753 +    next
   7.754 +      case (Cons hd tl)
   7.755 +      with hyp if_I hyp_sub 
   7.756 +      show "P I" 
   7.757 +	by auto
   7.758 +    qed
   7.759 +  qed
   7.760 +qed
   7.761 +
   7.762 +section "general recursion operators for the interface and class hiearchies"
   7.763 +
   7.764 +consts
   7.765 +  iface_rec  :: "prog \<times> qtname \<Rightarrow>   \<spacespace>  (qtname \<Rightarrow> iface \<Rightarrow> 'a set \<Rightarrow> 'a) \<Rightarrow> 'a"
   7.766 +  class_rec  :: "prog \<times> qtname \<Rightarrow> 'a \<Rightarrow> (qtname \<Rightarrow> class \<Rightarrow> 'a     \<Rightarrow> 'a) \<Rightarrow> 'a"
   7.767 +
   7.768 +recdef iface_rec "same_fst ws_prog (\<lambda>G. (subint1 G)^-1)" 
   7.769 +"iface_rec (G,I) = 
   7.770 +  (\<lambda>f. case iface G I of 
   7.771 +         None \<Rightarrow> arbitrary 
   7.772 +       | Some i \<Rightarrow> if ws_prog G 
   7.773 +                      then f I i 
   7.774 +                               ((\<lambda>J. iface_rec (G,J) f)`set (isuperIfs i))
   7.775 +                      else arbitrary)"
   7.776 +(hints recdef_wf: wf_subint1 intro: subint1I)
   7.777 +declare iface_rec.simps [simp del]
   7.778 +
   7.779 +lemma iface_rec: 
   7.780 +"\<lbrakk>iface G I = Some i; ws_prog G\<rbrakk> \<Longrightarrow> 
   7.781 + iface_rec (G,I) f = f I i ((\<lambda>J. iface_rec (G,J) f)`set (isuperIfs i))"
   7.782 +apply (subst iface_rec.simps)
   7.783 +apply simp
   7.784 +done
   7.785 +
   7.786 +recdef class_rec "same_fst ws_prog (\<lambda>G. (subcls1 G)^-1)"
   7.787 +"class_rec(G,C) = 
   7.788 +  (\<lambda>t f. case class G C of 
   7.789 +           None \<Rightarrow> arbitrary 
   7.790 +         | Some c \<Rightarrow> if ws_prog G 
   7.791 +                        then f C c 
   7.792 +                                 (if C = Object then t 
   7.793 +                                                else class_rec (G,super c) t f)
   7.794 +                        else arbitrary)"
   7.795 +(hints recdef_wf: wf_subcls1 intro: subcls1I)
   7.796 +declare class_rec.simps [simp del]
   7.797 +
   7.798 +lemma class_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
   7.799 + class_rec (G,C) t f = 
   7.800 +   f C c (if C = Object then t else class_rec (G,super c) t f)"
   7.801 +apply (rule class_rec.simps [THEN trans [THEN fun_cong [THEN fun_cong]]])
   7.802 +apply simp
   7.803 +done
   7.804 +(*
   7.805 +lemma bar:
   7.806 + "[| P;  !!x.  P ==> Q x  |] ==> Q x"
   7.807 +by simp
   7.808 +
   7.809 +lemma metaMP: "[| A ==> B; A |] ==> B"
   7.810 +by blast
   7.811 +
   7.812 +lemma True
   7.813 +proof- 
   7.814 +  presume t: "C  ==> E"
   7.815 +  thm metaMP [OF t]
   7.816 +
   7.817 +  presume r1: "\<And> B. P \<Longrightarrow> B"
   7.818 +  presume r2: "\<And> C. C \<Longrightarrow> P"
   7.819 +  thm r1 [OF r2]
   7.820 +
   7.821 +  thm metaMP [OF t]
   7.822 +
   7.823 +lemma ws_subcls1_induct4: "\<lbrakk>is_class G C; ws_prog G;  
   7.824 +  \<And>C c. \<lbrakk>C \<noteq> Object\<longrightarrow> P (super c)\<rbrakk> \<Longrightarrow> P C
   7.825 + \<rbrakk> \<Longrightarrow> P C"
   7.826 +proof -
   7.827 +  assume cls_C: "is_class G C"
   7.828 +  and       ws: "ws_prog G"
   7.829 +  and      hyp: "\<And>C c. \<lbrakk>C \<noteq> Object\<longrightarrow> P (super c)\<rbrakk> \<Longrightarrow> P C"
   7.830 +  thm ws_subcls1_induct [OF cls_C ws hyp]
   7.831 +
   7.832 +show
   7.833 +(\<And>C c. class G C = Some c \<and>
   7.834 +       (C \<noteq> Object \<longrightarrow> G\<turnstile>C\<prec>\<^sub>C\<^sub>1super c \<and> ?P (super c) \<and> is_class G (super c)) \<Longrightarrow>
   7.835 +       ?P C) \<Longrightarrow>
   7.836 +?P C
   7.837 +  show ?thesis
   7.838 +    thm "thm ws_subcls1_induct [OF cls_C ws hyp]"
   7.839 +    apply (rule ws_subcls1_induct)
   7.840 +  proof (rule ws_subcls1_induct)
   7.841 +    fix C c
   7.842 +    assume "class G C = Some c \<and>
   7.843 +            (C \<noteq> Object \<longrightarrow>
   7.844 +              G\<turnstile>C\<prec>\<^sub>C\<^sub>1super c \<and> P (super c) \<and> is_class G (super c))"
   7.845 +    show "C \<noteq> Object \<longrightarrow> P (super  (?c C c))" 
   7.846 +apply (erule ws_subcls1_induct)
   7.847 +apply assumption
   7.848 +apply (erule conjE)
   7.849 +apply (case_tac "C=Object")
   7.850 +apply blast
   7.851 +apply (erule impE)
   7.852 +apply assumption
   7.853 +apply (erule conjE)+
   7.854 +apply (rotate_tac 2)
   7.855 +sorry
   7.856 +
   7.857 +*)
   7.858 +
   7.859 +
   7.860 +constdefs
   7.861 +imethds:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
   7.862 +  (* methods of an interface, with overriding and inheritance, cf. 9.2 *)
   7.863 +"imethds G I 
   7.864 +  \<equiv> iface_rec (G,I)  
   7.865 +              (\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus> 
   7.866 +                        (o2s \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
   7.867 +	
   7.868 +
   7.869 +
   7.870 +end
     8.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     8.2 +++ b/src/HOL/Bali/DeclConcepts.thy	Mon Jan 28 17:00:19 2002 +0100
     8.3 @@ -0,0 +1,2540 @@
     8.4 +header {* Advanced concepts on Java declarations like overriding, inheritance,
     8.5 +dynamic method lookup*}
     8.6 +
     8.7 +theory DeclConcepts = TypeRel:
     8.8 +
     8.9 +section "access control (cf. 6.6), overriding and hiding (cf. 8.4.6.1)"
    8.10 +
    8.11 +constdefs
    8.12 +is_public :: "prog \<Rightarrow> qtname \<Rightarrow> bool"
    8.13 +"is_public G qn \<equiv> (case class G qn of
    8.14 +                     None       \<Rightarrow> (case iface G qn of
    8.15 +                                      None       \<Rightarrow> False
    8.16 +                                    | Some iface \<Rightarrow> access iface = Public)
    8.17 +                   | Some class \<Rightarrow> access class = Public)"
    8.18 +
    8.19 +subsection "accessibility of types (cf. 6.6.1)"
    8.20 +text {* 
    8.21 +Primitive types are always accessible, interfaces and classes are accessible
    8.22 +in their package or if they are defined public, an array type is accessible if
    8.23 +its element type is accessible *}
    8.24 + 
    8.25 +consts accessible_in   :: "prog \<Rightarrow> ty     \<Rightarrow> pname \<Rightarrow> bool"  
    8.26 +                                      ("_ \<turnstile> _ accessible'_in _" [61,61,61] 60)
    8.27 +       rt_accessible_in:: "prog \<Rightarrow> ref_ty \<Rightarrow> pname \<Rightarrow> bool" 
    8.28 +                                      ("_ \<turnstile> _ accessible'_in' _" [61,61,61] 60) 
    8.29 +primrec
    8.30 +"G\<turnstile>(PrimT p)   accessible_in pack  = True"
    8.31 +accessible_in_RefT_simp: 
    8.32 +"G\<turnstile>(RefT  r)   accessible_in pack  = G\<turnstile>r accessible_in' pack"
    8.33 +
    8.34 +"G\<turnstile>(NullT)     accessible_in' pack = True"
    8.35 +"G\<turnstile>(IfaceT I)  accessible_in' pack = ((pid I = pack) \<or> is_public G I)"
    8.36 +"G\<turnstile>(ClassT C)  accessible_in' pack = ((pid C = pack) \<or> is_public G C)"
    8.37 +"G\<turnstile>(ArrayT ty) accessible_in' pack = G\<turnstile>ty accessible_in pack"
    8.38 +
    8.39 +declare accessible_in_RefT_simp [simp del]
    8.40 +
    8.41 +constdefs
    8.42 +  is_acc_class :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
    8.43 +    "is_acc_class G P C \<equiv> is_class G C \<and> G\<turnstile>(Class C) accessible_in P"
    8.44 +  is_acc_iface :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
    8.45 +    "is_acc_iface G P I \<equiv> is_iface G I \<and> G\<turnstile>(Iface I) accessible_in P"
    8.46 +  is_acc_type  :: "prog \<Rightarrow> pname \<Rightarrow> ty     \<Rightarrow> bool"
    8.47 +    "is_acc_type  G P T \<equiv> is_type G T  \<and> G\<turnstile>T accessible_in P"
    8.48 +  is_acc_reftype  :: "prog \<Rightarrow> pname \<Rightarrow> ref_ty \<Rightarrow> bool"
    8.49 +  "is_acc_reftype  G P T \<equiv> isrtype G T  \<and> G\<turnstile>T accessible_in' P"
    8.50 +
    8.51 +lemma is_acc_classD:
    8.52 + "is_acc_class G P C \<Longrightarrow>  is_class G C \<and> G\<turnstile>(Class C) accessible_in P"
    8.53 +by (simp add: is_acc_class_def)
    8.54 +
    8.55 +lemma is_acc_class_is_class: "is_acc_class G P C \<Longrightarrow> is_class G C"
    8.56 +by (auto simp add: is_acc_class_def)
    8.57 +
    8.58 +lemma is_acc_ifaceD:
    8.59 +  "is_acc_iface G P I \<Longrightarrow> is_iface G I \<and> G\<turnstile>(Iface I) accessible_in P"
    8.60 +by (simp add: is_acc_iface_def)
    8.61 +
    8.62 +lemma is_acc_typeD:
    8.63 + "is_acc_type  G P T \<equiv> is_type G T  \<and> G\<turnstile>T accessible_in P"
    8.64 +by (simp add: is_acc_type_def)
    8.65 +
    8.66 +lemma is_acc_reftypeD:
    8.67 +"is_acc_reftype  G P T \<Longrightarrow> isrtype G T  \<and> G\<turnstile>T accessible_in' P"
    8.68 +by (simp add: is_acc_reftype_def)
    8.69 +
    8.70 +subsection "accessibility of members"
    8.71 +text {*
    8.72 +The accessibility of members is more involved as the accessibility of types.
    8.73 +We have to distinguish several cases to model the different effects of 
    8.74 +accessibility during inheritance, overriding and ordinary member access 
    8.75 +*}
    8.76 +
    8.77 +subsubsection {* Various technical conversion and selection functions *}
    8.78 +
    8.79 +text {* overloaded selector @{text accmodi} to select the access modifier 
    8.80 +out of various HOL types *}
    8.81 +
    8.82 +axclass has_accmodi < "type"
    8.83 +consts accmodi:: "'a::has_accmodi \<Rightarrow> acc_modi"
    8.84 +
    8.85 +instance acc_modi::has_accmodi
    8.86 +by (intro_classes)
    8.87 +
    8.88 +defs (overloaded)
    8.89 +acc_modi_accmodi_def: "accmodi (a::acc_modi) \<equiv> a"
    8.90 +
    8.91 +lemma acc_modi_accmodi_simp[simp]: "accmodi (a::acc_modi) = a"
    8.92 +by (simp add: acc_modi_accmodi_def)
    8.93 +
    8.94 +instance access_field_type:: ("type","type") has_accmodi
    8.95 +by (intro_classes)
    8.96 +
    8.97 +defs (overloaded)
    8.98 +decl_acc_modi_def: "accmodi (d::('a:: type) decl_scheme) \<equiv> access d"
    8.99 +
   8.100 +
   8.101 +lemma decl_acc_modi_simp[simp]: "accmodi (d::('a::type) decl_scheme) = access d"
   8.102 +by (simp add: decl_acc_modi_def)
   8.103 +
   8.104 +instance * :: ("type",has_accmodi) has_accmodi
   8.105 +by (intro_classes)
   8.106 +
   8.107 +defs (overloaded)
   8.108 +pair_acc_modi_def: "accmodi p \<equiv> (accmodi (snd p))"
   8.109 +
   8.110 +lemma pair_acc_modi_simp[simp]: "accmodi (x,a) = (accmodi a)"
   8.111 +by (simp add: pair_acc_modi_def)
   8.112 +
   8.113 +instance memberdecl :: has_accmodi
   8.114 +by (intro_classes)
   8.115 +
   8.116 +defs (overloaded)
   8.117 +memberdecl_acc_modi_def: "accmodi m \<equiv> (case m of
   8.118 +                                          fdecl f \<Rightarrow> accmodi f
   8.119 +                                        | mdecl m \<Rightarrow> accmodi m)"
   8.120 +
   8.121 +lemma memberdecl_fdecl_acc_modi_simp[simp]:
   8.122 + "accmodi (fdecl m) = accmodi m"
   8.123 +by (simp add: memberdecl_acc_modi_def)
   8.124 +
   8.125 +lemma memberdecl_mdecl_acc_modi_simp[simp]:
   8.126 + "accmodi (mdecl m) = accmodi m"
   8.127 +by (simp add: memberdecl_acc_modi_def)
   8.128 +
   8.129 +text {* overloaded selector @{text declclass} to select the declaring class 
   8.130 +out of various HOL types *}
   8.131 +
   8.132 +axclass has_declclass < "type"
   8.133 +consts declclass:: "'a::has_declclass \<Rightarrow> qtname"
   8.134 +
   8.135 +instance pid_field_type::("type","type") has_declclass
   8.136 +by (intro_classes)
   8.137 +
   8.138 +defs (overloaded)
   8.139 +qtname_declclass_def: "declclass (q::qtname) \<equiv> q"
   8.140 +
   8.141 +lemma qtname_declclass_simp[simp]: "declclass (q::qtname) = q"
   8.142 +by (simp add: qtname_declclass_def)
   8.143 +
   8.144 +instance * :: ("has_declclass","type") has_declclass
   8.145 +by (intro_classes)
   8.146 +
   8.147 +defs (overloaded)
   8.148 +pair_declclass_def: "declclass p \<equiv> declclass (fst p)"
   8.149 +
   8.150 +lemma pair_declclass_simp[simp]: "declclass (c,x) = declclass c" 
   8.151 +by (simp add: pair_declclass_def)
   8.152 +
   8.153 +text {* overloaded selector @{text is_static} to select the static modifier 
   8.154 +out of various HOL types *}
   8.155 +
   8.156 +
   8.157 +axclass has_static < "type"
   8.158 +consts is_static :: "'a::has_static \<Rightarrow> bool"
   8.159 +
   8.160 +(*
   8.161 +consts is_static :: "'a \<Rightarrow> bool"
   8.162 +*)
   8.163 +
   8.164 +instance access_field_type :: ("type","has_static") has_static
   8.165 +by (intro_classes) 
   8.166 +
   8.167 +defs (overloaded)
   8.168 +decl_is_static_def: 
   8.169 + "is_static (m::('a::has_static) decl_scheme) \<equiv> is_static (Decl.decl.more m)" 
   8.170 +
   8.171 +instance static_field_type :: ("type","type") has_static
   8.172 +by (intro_classes)
   8.173 +
   8.174 +defs (overloaded)
   8.175 +static_field_type_is_static_def: 
   8.176 + "is_static (m::(bool,'b::type) static_field_type) \<equiv> static_val m"
   8.177 +
   8.178 +lemma member_is_static_simp: "is_static (m::'a member_scheme) = static m"
   8.179 +apply (cases m)
   8.180 +apply (simp add: static_field_type_is_static_def 
   8.181 +                 decl_is_static_def Decl.member.dest_convs)
   8.182 +done
   8.183 +
   8.184 +instance * :: ("type","has_static") has_static
   8.185 +by (intro_classes)
   8.186 +
   8.187 +defs (overloaded)
   8.188 +pair_is_static_def: "is_static p \<equiv> is_static (snd p)"
   8.189 +
   8.190 +lemma pair_is_static_simp [simp]: "is_static (x,s) = is_static s"
   8.191 +by (simp add: pair_is_static_def)
   8.192 +
   8.193 +lemma pair_is_static_simp1: "is_static p = is_static (snd p)"
   8.194 +by (simp add: pair_is_static_def)
   8.195 +
   8.196 +instance memberdecl:: has_static
   8.197 +by (intro_classes)
   8.198 +
   8.199 +defs (overloaded)
   8.200 +memberdecl_is_static_def: 
   8.201 + "is_static m \<equiv> (case m of
   8.202 +                    fdecl f \<Rightarrow> is_static f
   8.203 +                  | mdecl m \<Rightarrow> is_static m)"
   8.204 +
   8.205 +lemma memberdecl_is_static_fdecl_simp[simp]:
   8.206 + "is_static (fdecl f) = is_static f"
   8.207 +by (simp add: memberdecl_is_static_def)
   8.208 +
   8.209 +lemma memberdecl_is_static_mdecl_simp[simp]:
   8.210 + "is_static (mdecl m) = is_static m"
   8.211 +by (simp add: memberdecl_is_static_def)
   8.212 +
   8.213 +lemma mhead_static_simp [simp]: "is_static (mhead m) = is_static m"
   8.214 +by (cases m) (simp add: mhead_def member_is_static_simp)
   8.215 +
   8.216 +constdefs  (* some mnemotic selectors for (qtname \<times> ('a::more) decl_scheme) 
   8.217 +            * the first component is a class or interface name
   8.218 +            * the second component is a method, field or method head *)
   8.219 +(* "declclass":: "(qtname \<times> ('a::more) decl_scheme) \<Rightarrow> qtname"*)
   8.220 +(* "declclass \<equiv> fst" *)          (* get the class component *)
   8.221 +
   8.222 + "decliface":: "(qtname \<times> ('a::type) decl_scheme) \<Rightarrow> qtname"
   8.223 + "decliface \<equiv> fst"          (* get the interface component *)
   8.224 +
   8.225 +(*
   8.226 + "member"::   "(qtname \<times> ('a::type) decl_scheme) \<Rightarrow> ('a::type) decl_scheme"
   8.227 +*)
   8.228 + "mbr"::   "(qtname \<times> memberdecl) \<Rightarrow> memberdecl"
   8.229 + "mbr \<equiv> snd"            (* get the memberdecl component *)
   8.230 +
   8.231 + "mthd"::   "('b \<times> 'a) \<Rightarrow> 'a"
   8.232 +                           (* also used for mdecl,mhead *)
   8.233 + "mthd \<equiv> snd"              (* get the method component *)
   8.234 +
   8.235 + "fld"::   "('b \<times> ('a::type) decl_scheme) \<Rightarrow> ('a::type) decl_scheme"
   8.236 +              (* also used for ((vname \<times> qtname)\<times> field) *)
   8.237 + "fld \<equiv> snd"               (* get the field component *)
   8.238 +
   8.239 +(* "accmodi" :: "('b \<times> ('a::type) decl_scheme) \<Rightarrow> acc_modi"*)
   8.240 +                           (* also used for mdecl *)
   8.241 +(* "accmodi \<equiv> access \<circ> snd"*)  (* get the access modifier *) 
   8.242 +(*
   8.243 + "is_static" ::"('b \<times> ('a::type) member_scheme) \<Rightarrow> bool" *)
   8.244 +                            (* also defined for emhead cf. WellType *)
   8.245 + (*"is_static \<equiv> static \<circ> snd"*) (* get the static modifier *)
   8.246 +
   8.247 +constdefs (* some mnemotic selectors for (vname \<times> qtname) *)
   8.248 + fname:: "(vname \<times> 'a) \<Rightarrow> vname" (* also used for fdecl *)
   8.249 + "fname \<equiv> fst"
   8.250 +  
   8.251 +  declclassf:: "(vname \<times> qtname) \<Rightarrow> qtname"
   8.252 + "declclassf \<equiv> snd"
   8.253 +
   8.254 +(*
   8.255 +lemma declclass_simp[simp]: "declclass (C,m) = C"
   8.256 +by (simp add: declclass_def)
   8.257 +*)
   8.258 +
   8.259 +lemma decliface_simp[simp]: "decliface (I,m) = I"
   8.260 +by (simp add: decliface_def) 
   8.261 +
   8.262 +lemma mbr_simp[simp]: "mbr (C,m) = m"
   8.263 +by (simp add: mbr_def)
   8.264 +
   8.265 +lemma access_mbr_simp [simp]: "(accmodi (mbr m)) = accmodi m"
   8.266 +by (cases m) (simp add:  mbr_def) 
   8.267 +
   8.268 +lemma mthd_simp[simp]: "mthd (C,m) = m"
   8.269 +by (simp add: mthd_def)
   8.270 +
   8.271 +lemma fld_simp[simp]: "fld (C,f) = f"
   8.272 +by (simp add: fld_def)
   8.273 +
   8.274 +lemma accmodi_simp[simp]: "accmodi (C,m) = access m"
   8.275 +by (simp )
   8.276 +
   8.277 +lemma access_mthd_simp [simp]: "(access (mthd m)) = accmodi m"
   8.278 +by (cases m) (simp add: mthd_def) 
   8.279 +
   8.280 +lemma access_fld_simp [simp]: "(access (fld f)) = accmodi f"
   8.281 +by (cases f) (simp add:  fld_def) 
   8.282 +
   8.283 +(*
   8.284 +lemma is_static_simp[simp]: "is_static (C,m) = static m"
   8.285 +by (simp add: is_static_def)
   8.286 +*)
   8.287 +
   8.288 +lemma static_mthd_simp[simp]: "static (mthd m) = is_static m"
   8.289 +by (cases m) (simp add:  mthd_def member_is_static_simp)
   8.290 +
   8.291 +lemma mthd_is_static_simp [simp]: "is_static (mthd m) = is_static m"
   8.292 +by (cases m) simp
   8.293 +
   8.294 +lemma static_fld_simp[simp]: "static (fld f) = is_static f"
   8.295 +by (cases f) (simp add:  fld_def member_is_static_simp)
   8.296 +
   8.297 +lemma ext_field_simp [simp]: "(declclass f,fld f) = f"
   8.298 +by (cases f) (simp add:  fld_def)
   8.299 +
   8.300 +lemma ext_method_simp [simp]: "(declclass m,mthd m) = m"
   8.301 +by (cases m) (simp add:  mthd_def)
   8.302 +
   8.303 +lemma ext_mbr_simp [simp]: "(declclass m,mbr m) = m"
   8.304 +by (cases m) (simp add: mbr_def)
   8.305 +
   8.306 +lemma fname_simp[simp]:"fname (n,c) = n"
   8.307 +by (simp add: fname_def)
   8.308 +
   8.309 +lemma declclassf_simp[simp]:"declclassf (n,c) = c"
   8.310 +by (simp add: declclassf_def)
   8.311 +
   8.312 +constdefs  (* some mnemotic selectors for (vname \<times> qtname) *)
   8.313 +  "fldname"  :: "(vname \<times> qtname) \<Rightarrow> vname" 
   8.314 +  "fldname \<equiv> fst"
   8.315 +
   8.316 +  "fldclass" :: "(vname \<times> qtname) \<Rightarrow> qtname"
   8.317 +  "fldclass \<equiv> snd"
   8.318 +
   8.319 +lemma fldname_simp[simp]: "fldname (n,c) = n"
   8.320 +by (simp add: fldname_def)
   8.321 +
   8.322 +lemma fldclass_simp[simp]: "fldclass (n,c) = c"
   8.323 +by (simp add: fldclass_def)
   8.324 +
   8.325 +lemma ext_fieldname_simp[simp]: "(fldname f,fldclass f) = f"
   8.326 +by (simp add: fldname_def fldclass_def)
   8.327 +
   8.328 +text {* Convert a qualified method declaration (qualified with its declaring 
   8.329 +class) to a qualified member declaration:  @{text methdMembr}  *}
   8.330 +
   8.331 +constdefs
   8.332 +methdMembr :: "(qtname \<times> mdecl) \<Rightarrow> (qtname \<times> memberdecl)"
   8.333 + "methdMembr m \<equiv> (fst m,mdecl (snd m))"
   8.334 +
   8.335 +lemma methdMembr_simp[simp]: "methdMembr (c,m) = (c,mdecl m)"
   8.336 +by (simp add: methdMembr_def)
   8.337 +
   8.338 +lemma accmodi_methdMembr_simp[simp]: "accmodi (methdMembr m) = accmodi m"
   8.339 +by (cases m) (simp add: methdMembr_def)
   8.340 +
   8.341 +lemma is_static_methdMembr_simp[simp]: "is_static (methdMembr m) = is_static m"
   8.342 +by (cases m) (simp add: methdMembr_def)
   8.343 +
   8.344 +lemma declclass_methdMembr_simp[simp]: "declclass (methdMembr m) = declclass m"
   8.345 +by (cases m) (simp add: methdMembr_def)
   8.346 +
   8.347 +text {* Convert a qualified method (qualified with its declaring 
   8.348 +class) to a qualified member declaration:  @{text method}  *}
   8.349 +
   8.350 +constdefs
   8.351 +method :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> memberdecl)" 
   8.352 +"method sig m \<equiv> (declclass m, mdecl (sig, mthd m))"
   8.353 +
   8.354 +lemma method_simp[simp]: "method sig (C,m) = (C,mdecl (sig,m))"
   8.355 +by (simp add: method_def)
   8.356 +
   8.357 +lemma accmodi_method_simp[simp]: "accmodi (method sig m) = accmodi m"
   8.358 +by (simp add: method_def)
   8.359 +
   8.360 +lemma declclass_method_simp[simp]: "declclass (method sig m) = declclass m"
   8.361 +by (simp add: method_def)
   8.362 +
   8.363 +lemma is_static_method_simp[simp]: "is_static (method sig m) = is_static m"
   8.364 +by (cases m) (simp add: method_def)
   8.365 +
   8.366 +lemma mbr_method_simp[simp]: "mbr (method sig m) = mdecl (sig,mthd m)"
   8.367 +by (simp add: mbr_def method_def)
   8.368 +
   8.369 +lemma memberid_method_simp[simp]:  "memberid (method sig m) = mid sig"
   8.370 +  by (simp add: method_def) 
   8.371 +
   8.372 +constdefs
   8.373 +fieldm :: "vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> (qtname \<times> memberdecl)" 
   8.374 +"fieldm n f \<equiv> (declclass f, fdecl (n, fld f))"
   8.375 +
   8.376 +lemma fieldm_simp[simp]: "fieldm n (C,f) = (C,fdecl (n,f))"
   8.377 +by (simp add: fieldm_def)
   8.378 +
   8.379 +lemma accmodi_fieldm_simp[simp]: "accmodi (fieldm n f) = accmodi f"
   8.380 +by (simp add: fieldm_def)
   8.381 +
   8.382 +lemma declclass_fieldm_simp[simp]: "declclass (fieldm n f) = declclass f"
   8.383 +by (simp add: fieldm_def)
   8.384 +
   8.385 +lemma is_static_fieldm_simp[simp]: "is_static (fieldm n f) = is_static f"
   8.386 +by (cases f) (simp add: fieldm_def)
   8.387 +
   8.388 +lemma mbr_fieldm_simp[simp]: "mbr (fieldm n f) = fdecl (n,fld f)"
   8.389 +by (simp add: mbr_def fieldm_def)
   8.390 +
   8.391 +lemma memberid_fieldm_simp[simp]:  "memberid (fieldm n f) = fid n"
   8.392 +by (simp add: fieldm_def) 
   8.393 +
   8.394 +text {* Select the signature out of a qualified method declaration:
   8.395 + @{text msig} *}
   8.396 +
   8.397 +constdefs msig:: "(qtname \<times> mdecl) \<Rightarrow> sig"
   8.398 +"msig m \<equiv> fst (snd m)"
   8.399 +
   8.400 +lemma msig_simp[simp]: "msig (c,(s,m)) = s"
   8.401 +by (simp add: msig_def)
   8.402 +
   8.403 +text {* Convert a qualified method (qualified with its declaring 
   8.404 +class) to a qualified method declaration:  @{text qmdecl}  *}
   8.405 +
   8.406 +constdefs qmdecl :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> mdecl)"
   8.407 +"qmdecl sig m \<equiv> (declclass m, (sig,mthd m))"
   8.408 +
   8.409 +lemma qmdecl_simp[simp]: "qmdecl sig (C,m) = (C,(sig,m))"
   8.410 +by (simp add: qmdecl_def)
   8.411 +
   8.412 +lemma declclass_qmdecl_simp[simp]: "declclass (qmdecl sig m) = declclass m"
   8.413 +by (simp add: qmdecl_def)
   8.414 +
   8.415 +lemma accmodi_qmdecl_simp[simp]: "accmodi (qmdecl sig m) = accmodi m"
   8.416 +by (simp add: qmdecl_def)
   8.417 +
   8.418 +lemma is_static_qmdecl_simp[simp]: "is_static (qmdecl sig m) = is_static m"
   8.419 +by (cases m) (simp add: qmdecl_def)
   8.420 +
   8.421 +lemma msig_qmdecl_simp[simp]: "msig (qmdecl sig m) = sig"
   8.422 +by (simp add: qmdecl_def)
   8.423 +
   8.424 +lemma mdecl_qmdecl_simp[simp]:  
   8.425 + "mdecl (mthd (qmdecl sig new)) = mdecl (sig, mthd new)" 
   8.426 +by (simp add: qmdecl_def) 
   8.427 +
   8.428 +lemma methdMembr_qmdecl_simp [simp]: 
   8.429 + "methdMembr (qmdecl sig old) = method sig old"
   8.430 +by (simp add: methdMembr_def qmdecl_def method_def)
   8.431 +
   8.432 +text {* overloaded selector @{text resTy} to select the result type 
   8.433 +out of various HOL types *}
   8.434 +
   8.435 +axclass has_resTy < "type"
   8.436 +consts resTy:: "'a::has_resTy \<Rightarrow> ty"
   8.437 +
   8.438 +instance access_field_type :: ("type","has_resTy") has_resTy
   8.439 +by (intro_classes)
   8.440 +
   8.441 +defs (overloaded)
   8.442 +decl_resTy_def: 
   8.443 + "resTy (m::('a::has_resTy) decl_scheme) \<equiv> resTy (Decl.decl.more m)" 
   8.444 +
   8.445 +instance static_field_type :: ("type","has_resTy") has_resTy
   8.446 +by (intro_classes)
   8.447 +
   8.448 +defs (overloaded)
   8.449 +static_field_type_resTy_def: 
   8.450 + "resTy (m::(bool,'b::has_resTy) static_field_type) 
   8.451 +  \<equiv> resTy (static_more m)" 
   8.452 +
   8.453 +instance pars_field_type :: ("type","has_resTy") has_resTy
   8.454 +by (intro_classes)
   8.455 +
   8.456 +defs (overloaded)
   8.457 +pars_field_type_resTy_def: 
   8.458 + "resTy (m::(vname list,'b::has_resTy) pars_field_type) 
   8.459 +  \<equiv> resTy (pars_more m)" 
   8.460 +
   8.461 +instance resT_field_type :: ("type","type") has_resTy
   8.462 +by (intro_classes)
   8.463 +
   8.464 +defs (overloaded)
   8.465 +resT_field_type_resTy_def: 
   8.466 + "resTy (m::(ty,'b::type) resT_field_type) 
   8.467 +  \<equiv> resT_val m" 
   8.468 +
   8.469 +lemma mhead_resTy_simp: "resTy (m::'a mhead_scheme) = resT m"
   8.470 +apply (cases m)
   8.471 +apply (simp add: decl_resTy_def static_field_type_resTy_def 
   8.472 +                 pars_field_type_resTy_def resT_field_type_resTy_def
   8.473 +                 member.dest_convs mhead.dest_convs)
   8.474 +done
   8.475 +
   8.476 +lemma resTy_mhead [simp]:"resTy (mhead m) = resTy m"
   8.477 +by (simp add: mhead_def mhead_resTy_simp)
   8.478 +
   8.479 +instance * :: ("type","has_resTy") has_resTy
   8.480 +by (intro_classes)
   8.481 +
   8.482 +defs (overloaded)
   8.483 +pair_resTy_def: "resTy p \<equiv> resTy (snd p)"
   8.484 +
   8.485 +lemma pair_resTy_simp[simp]: "resTy (x,m) = resTy m"
   8.486 +by (simp add: pair_resTy_def)
   8.487 +
   8.488 +lemma qmdecl_resTy_simp [simp]: "resTy (qmdecl sig m) = resTy m"
   8.489 +by (cases m) (simp)
   8.490 +
   8.491 +lemma resTy_mthd [simp]:"resTy (mthd m) = resTy m"
   8.492 +by (cases m) (simp add: mthd_def )
   8.493 +
   8.494 +subsubsection "inheritable-in"
   8.495 +text {*
   8.496 +@{text "G\<turnstile>m inheritable_in P"}: m can be inherited by
   8.497 +classes in package P if:
   8.498 +\begin{itemize} 
   8.499 +\item the declaration class of m is accessible in P and
   8.500 +\item the member m is declared with protected or public access or if it is
   8.501 +      declared with default (package) access, the package of the declaration 
   8.502 +      class of m is also P. If the member m is declared with private access
   8.503 +      it is not accessible for inheritance at all.
   8.504 +\end{itemize}
   8.505 +*}
   8.506 +constdefs
   8.507 +inheritable_in:: 
   8.508 + "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool"
   8.509 +                  ("_ \<turnstile> _ inheritable'_in _" [61,61,61] 60)
   8.510 +"G\<turnstile>membr inheritable_in pack 
   8.511 +  \<equiv> (case (accmodi membr) of
   8.512 +       Private   \<Rightarrow> False
   8.513 +     | Package   \<Rightarrow> (pid (declclass membr)) = pack
   8.514 +     | Protected \<Rightarrow> True
   8.515 +     | Public    \<Rightarrow> True)"
   8.516 +
   8.517 +syntax
   8.518 +Method_inheritable_in::
   8.519 + "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> pname \<Rightarrow> bool"
   8.520 +                ("_ \<turnstile>Method _ inheritable'_in _ " [61,61,61] 60)
   8.521 +
   8.522 +translations
   8.523 +"G\<turnstile>Method m inheritable_in p" == "G\<turnstile>methdMembr m inheritable_in p"
   8.524 +
   8.525 +syntax
   8.526 +Methd_inheritable_in::
   8.527 + "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> pname \<Rightarrow> bool"
   8.528 +                ("_ \<turnstile>Methd _ _ inheritable'_in _ " [61,61,61,61] 60)
   8.529 +
   8.530 +translations
   8.531 +"G\<turnstile>Methd s m inheritable_in p" == "G\<turnstile>(method s m) inheritable_in p"
   8.532 +
   8.533 +subsubsection "declared-in/undeclared-in"
   8.534 +
   8.535 +constdefs cdeclaredmethd:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,methd) table"
   8.536 +"cdeclaredmethd G C 
   8.537 +  \<equiv> (case class G C of
   8.538 +       None \<Rightarrow> \<lambda> sig. None
   8.539 +     | Some c \<Rightarrow> table_of (methods c)
   8.540 +    )"
   8.541 +
   8.542 +constdefs
   8.543 +cdeclaredfield:: "prog \<Rightarrow> qtname \<Rightarrow> (vname,field) table"
   8.544 +"cdeclaredfield G C 
   8.545 +  \<equiv> (case class G C of
   8.546 +       None \<Rightarrow> \<lambda> sig. None
   8.547 +     | Some c \<Rightarrow> table_of (cfields c)
   8.548 +    )"
   8.549 +
   8.550 +
   8.551 +constdefs
   8.552 +declared_in:: "prog  \<Rightarrow> memberdecl \<Rightarrow> qtname \<Rightarrow> bool"
   8.553 +                                 ("_\<turnstile> _ declared'_in _" [61,61,61] 60)
   8.554 +"G\<turnstile>m declared_in C \<equiv> (case m of
   8.555 +                        fdecl (fn,f ) \<Rightarrow> cdeclaredfield G C fn  = Some f
   8.556 +                      | mdecl (sig,m) \<Rightarrow> cdeclaredmethd G C sig = Some m)"
   8.557 +
   8.558 +syntax
   8.559 +method_declared_in:: "prog  \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
   8.560 +                                 ("_\<turnstile>Method _ declared'_in _" [61,61,61] 60)
   8.561 +translations
   8.562 +"G\<turnstile>Method m declared_in C" == "G\<turnstile>mdecl (mthd m) declared_in C"
   8.563 +
   8.564 +syntax
   8.565 +methd_declared_in:: "prog  \<Rightarrow> sig  \<Rightarrow>(qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
   8.566 +                               ("_\<turnstile>Methd _  _ declared'_in _" [61,61,61,61] 60)
   8.567 +translations
   8.568 +"G\<turnstile>Methd s m declared_in C" == "G\<turnstile>mdecl (s,mthd m) declared_in C"
   8.569 +
   8.570 +lemma declared_in_classD:
   8.571 + "G\<turnstile>m declared_in C \<Longrightarrow> is_class G C"
   8.572 +by (cases m) 
   8.573 +   (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)
   8.574 +
   8.575 +constdefs
   8.576 +undeclared_in:: "prog  \<Rightarrow> memberid \<Rightarrow> qtname \<Rightarrow> bool"
   8.577 +                                 ("_\<turnstile> _ undeclared'_in _" [61,61,61] 60)
   8.578 +
   8.579 +"G\<turnstile>m undeclared_in C \<equiv> (case m of
   8.580 +                            fid fn  \<Rightarrow> cdeclaredfield G C fn  = None
   8.581 +                          | mid sig \<Rightarrow> cdeclaredmethd G C sig = None)"
   8.582 +
   8.583 +lemma method_declared_inI:
   8.584 +  "\<lbrakk>class G C = Some c; table_of (methods c) sig = Some m\<rbrakk> 
   8.585 +   \<Longrightarrow> G\<turnstile>mdecl (sig,m) declared_in C"
   8.586 +by (auto simp add: declared_in_def cdeclaredmethd_def)
   8.587 +
   8.588 +
   8.589 +subsubsection "members"
   8.590 +
   8.591 +consts
   8.592 +members:: "prog \<Rightarrow> (qtname \<times> (qtname \<times> memberdecl)) set"
   8.593 +(* Can't just take a function: prog \<Rightarrow> qtname \<Rightarrow> memberdecl set because
   8.594 +   the class qtname changes to the superclass in the inductive definition
   8.595 +   below
   8.596 +*)
   8.597 +
   8.598 +syntax
   8.599 +member_of:: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
   8.600 +                           ("_ \<turnstile> _ member'_of _" [61,61,61] 60)
   8.601 +
   8.602 +translations
   8.603 + "G\<turnstile>m member_of C" \<rightleftharpoons> "(C,m) \<in> members G"
   8.604 +
   8.605 +inductive "members G"  intros
   8.606 +
   8.607 +Immediate: "\<lbrakk>G\<turnstile>mbr m declared_in C;declclass m = C\<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
   8.608 +Inherited: "\<lbrakk>G\<turnstile>m inheritable_in (pid C); G\<turnstile>memberid m undeclared_in C; 
   8.609 +             G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S 
   8.610 +            \<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
   8.611 +text {* Note that in the case of an inherited member only the members of the
   8.612 +direct superclass are concerned. If a member of a superclass of the direct
   8.613 +superclass isn't inherited in the direct superclass (not member of the
   8.614 +direct superclass) than it can't be a member of the class. E.g. If a
   8.615 +member of a class A is defined with package access it isn't member of a 
   8.616 +subclass S if S isn't in the same package as A. Any further subclasses 
   8.617 +of S will not inherit the member, regardless if they are in the same
   8.618 +package as A or not.*}
   8.619 +
   8.620 +syntax
   8.621 +method_member_of:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
   8.622 +                           ("_ \<turnstile>Method _ member'_of _" [61,61,61] 60)
   8.623 +
   8.624 +translations
   8.625 + "G\<turnstile>Method m member_of C" \<rightleftharpoons> "G\<turnstile>(methdMembr m) member_of C" 
   8.626 +
   8.627 +syntax
   8.628 +methd_member_of:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
   8.629 +                           ("_ \<turnstile>Methd _ _ member'_of _" [61,61,61,61] 60)
   8.630 +
   8.631 +translations
   8.632 + "G\<turnstile>Methd s m member_of C" \<rightleftharpoons> "G\<turnstile>(method s m) member_of C" 
   8.633 +
   8.634 +syntax
   8.635 +fieldm_member_of:: "prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> bool"
   8.636 +                           ("_ \<turnstile>Field _  _ member'_of _" [61,61,61] 60)
   8.637 +
   8.638 +translations
   8.639 + "G\<turnstile>Field n f member_of C" \<rightleftharpoons> "G\<turnstile>fieldm n f member_of C" 
   8.640 +
   8.641 +constdefs
   8.642 +inherits:: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> bool"
   8.643 +                           ("_ \<turnstile> _ inherits _" [61,61,61] 60)
   8.644 +"G\<turnstile>C inherits m 
   8.645 +  \<equiv> G\<turnstile>m inheritable_in (pid C) \<and> G\<turnstile>memberid m undeclared_in C \<and> 
   8.646 +    (\<exists> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<and> G\<turnstile>(Class S) accessible_in (pid C) \<and> G\<turnstile>m member_of S)"
   8.647 +
   8.648 +lemma inherits_member: "G\<turnstile>C inherits m \<Longrightarrow> G\<turnstile>m member_of C"
   8.649 +by (auto simp add: inherits_def intro: members.Inherited)
   8.650 +
   8.651 +
   8.652 +constdefs member_in::"prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
   8.653 +                           ("_ \<turnstile> _ member'_in _" [61,61,61] 60)
   8.654 +"G\<turnstile>m member_in C \<equiv> \<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC"
   8.655 +text {* A member is in a class if it is member of the class or a superclass.
   8.656 +If a member is in a class we can select this member. This additional notion
   8.657 +is necessary since not all members are inherited to subclasses. So such
   8.658 +members are not member-of the subclass but member-in the subclass.*}
   8.659 +
   8.660 +syntax
   8.661 +method_member_in:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
   8.662 +                           ("_ \<turnstile>Method _ member'_in _" [61,61,61] 60)
   8.663 +
   8.664 +translations
   8.665 + "G\<turnstile>Method m member_in C" \<rightleftharpoons> "G\<turnstile>(methdMembr m) member_in C" 
   8.666 +
   8.667 +syntax
   8.668 +methd_member_in:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
   8.669 +                           ("_ \<turnstile>Methd _ _ member'_in _" [61,61,61,61] 60)
   8.670 +
   8.671 +translations
   8.672 + "G\<turnstile>Methd s m member_in C" \<rightleftharpoons> "G\<turnstile>(method s m) member_in C" 
   8.673 +
   8.674 +consts stat_overridesR::
   8.675 +  "prog  \<Rightarrow> ((qtname \<times> mdecl) \<times> (qtname \<times> mdecl)) set"
   8.676 +
   8.677 +lemma member_inD: "G\<turnstile>m member_in C 
   8.678 + \<Longrightarrow> \<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC"
   8.679 +by (auto simp add: member_in_def)
   8.680 +
   8.681 +lemma member_inI: "\<lbrakk>G \<turnstile> m member_of provC;G\<turnstile> C \<preceq>\<^sub>C provC\<rbrakk> \<Longrightarrow>  G\<turnstile>m member_in C"
   8.682 +by (auto simp add: member_in_def)
   8.683 +
   8.684 +lemma member_of_to_member_in: "G \<turnstile> m member_of C \<Longrightarrow> G \<turnstile>m member_in C"
   8.685 +by (auto intro: member_inI)
   8.686 +
   8.687 +
   8.688 +subsubsection "overriding"
   8.689 +
   8.690 +text {* Unfortunately the static notion of overriding (used during the
   8.691 +typecheck of the compiler) and the dynamic notion of overriding (used during
   8.692 +execution in the JVM) are not exactly the same. 
   8.693 +*}
   8.694 +
   8.695 +text {* Static overriding (used during the typecheck of the compiler) *}
   8.696 +syntax
   8.697 +stat_overrides:: "prog  \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
   8.698 +                                  ("_ \<turnstile> _ overrides\<^sub>S _" [61,61,61] 60)
   8.699 +translations
   8.700 + "G\<turnstile>new overrides\<^sub>S  old" == "(new,old) \<in> stat_overridesR G "
   8.701 +
   8.702 +inductive "stat_overridesR G" intros
   8.703 +
   8.704 +Direct: "\<lbrakk>\<not> is_static new; msig new = msig old; 
   8.705 +         G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
   8.706 +         G\<turnstile>Method new declared_in (declclass new);  
   8.707 +         G\<turnstile>Method old declared_in (declclass old); 
   8.708 +         G\<turnstile>Method old inheritable_in pid (declclass new);
   8.709 +         G\<turnstile>(declclass new) \<prec>\<^sub>C\<^sub>1 superNew;
   8.710 +         G \<turnstile>Method old member_of superNew
   8.711 +         \<rbrakk> \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
   8.712 +
   8.713 +Indirect: "\<lbrakk>G\<turnstile>new overrides\<^sub>S inter; G\<turnstile>inter overrides\<^sub>S old\<rbrakk>
   8.714 +           \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
   8.715 +
   8.716 +text {* Dynamic overriding (used during the typecheck of the compiler) *}
   8.717 +consts overridesR::
   8.718 +  "prog  \<Rightarrow> ((qtname \<times> mdecl) \<times> (qtname \<times> mdecl)) set"
   8.719 +
   8.720 +
   8.721 +overrides:: "prog  \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
   8.722 +                                  ("_ \<turnstile> _ overrides _" [61,61,61] 60)
   8.723 +translations
   8.724 + "G\<turnstile>new overrides old" == "(new,old) \<in> overridesR G "
   8.725 +
   8.726 +inductive "overridesR G" intros
   8.727 +
   8.728 +Direct: "\<lbrakk>\<not> is_static new; \<not> is_static old; accmodi new \<noteq> Private;
   8.729 +         msig new = msig old; 
   8.730 +         G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
   8.731 +         G\<turnstile>Method new declared_in (declclass new);  
   8.732 +         G\<turnstile>Method old declared_in (declclass old);    
   8.733 +         G\<turnstile>Method old inheritable_in pid (declclass new);
   8.734 +         G\<turnstile>resTy new \<preceq> resTy old
   8.735 +         \<rbrakk> \<Longrightarrow> G\<turnstile>new overrides old"
   8.736 +
   8.737 +Indirect: "\<lbrakk>G\<turnstile>new overrides inter; G\<turnstile>inter overrides old\<rbrakk>
   8.738 +           \<Longrightarrow> G\<turnstile>new overrides old"
   8.739 +
   8.740 +syntax
   8.741 +sig_stat_overrides:: 
   8.742 + "prog  \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool" 
   8.743 +                                  ("_,_\<turnstile> _ overrides\<^sub>S _" [61,61,61,61] 60)
   8.744 +translations
   8.745 + "G,s\<turnstile>new overrides\<^sub>S old" \<rightharpoonup> "G\<turnstile>(qmdecl s new) overrides\<^sub>S (qmdecl s old)" 
   8.746 +
   8.747 +syntax
   8.748 +sig_overrides:: "prog  \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool" 
   8.749 +                                  ("_,_\<turnstile> _ overrides _" [61,61,61,61] 60)
   8.750 +translations
   8.751 + "G,s\<turnstile>new overrides old" \<rightharpoonup> "G\<turnstile>(qmdecl s new) overrides (qmdecl s old)" 
   8.752 +
   8.753 +subsubsection "Hiding"
   8.754 +
   8.755 +constdefs hides::
   8.756 +"prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
   8.757 +                                ("_\<turnstile> _ hides _" [61,61,61] 60)
   8.758 +"G\<turnstile>new hides old
   8.759 +  \<equiv> is_static new \<and> msig new = msig old \<and>
   8.760 +    G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
   8.761 +    G\<turnstile>Method new declared_in (declclass new) \<and>
   8.762 +    G\<turnstile>Method old declared_in (declclass old) \<and> 
   8.763 +    G\<turnstile>Method old inheritable_in pid (declclass new)"
   8.764 +
   8.765 +syntax
   8.766 +sig_hides:: "prog  \<Rightarrow> sig \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
   8.767 +                                  ("_,_\<turnstile> _ hides _" [61,61,61,61] 60)
   8.768 +translations
   8.769 + "G,s\<turnstile>new hides old" \<rightharpoonup> "G\<turnstile>(qmdecl s new) hides (qmdecl s old)" 
   8.770 +
   8.771 +lemma hidesI:
   8.772 +"\<lbrakk>is_static new; msig new = msig old;
   8.773 +  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
   8.774 +  G\<turnstile>Method new declared_in (declclass new);
   8.775 +  G\<turnstile>Method old declared_in (declclass old);
   8.776 +  G\<turnstile>Method old inheritable_in pid (declclass new)
   8.777 + \<rbrakk> \<Longrightarrow> G\<turnstile>new hides old"
   8.778 +by (auto simp add: hides_def)
   8.779 +
   8.780 +lemma hidesD:
   8.781 +"\<lbrakk>G\<turnstile>new hides old\<rbrakk> \<Longrightarrow>  
   8.782 +  declclass new \<noteq> Object \<and> is_static new \<and> msig new = msig old \<and> 
   8.783 +  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
   8.784 +  G\<turnstile>Method new declared_in (declclass new) \<and>   
   8.785 +  G\<turnstile>Method old declared_in (declclass old)"
   8.786 +by (auto simp add: hides_def)
   8.787 +
   8.788 +lemma overrides_commonD:
   8.789 +"\<lbrakk>G\<turnstile>new overrides old\<rbrakk> \<Longrightarrow>  
   8.790 +  declclass new \<noteq> Object \<and> \<not> is_static new \<and> \<not> is_static old \<and>
   8.791 +  accmodi new \<noteq> Private \<and> 
   8.792 +  msig new = msig old  \<and>
   8.793 +  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
   8.794 +  G\<turnstile>Method new declared_in (declclass new) \<and>   
   8.795 +  G\<turnstile>Method old declared_in (declclass old)"
   8.796 +by (induct set: overridesR) (auto intro: trancl_trans)
   8.797 +
   8.798 +lemma ws_overrides_commonD:
   8.799 +"\<lbrakk>G\<turnstile>new overrides old;ws_prog G\<rbrakk> \<Longrightarrow>  
   8.800 +  declclass new \<noteq> Object \<and> \<not> is_static new \<and> \<not> is_static old \<and>
   8.801 +  accmodi new \<noteq> Private \<and> G\<turnstile>resTy new \<preceq> resTy old \<and>
   8.802 +  msig new = msig old  \<and>
   8.803 +  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
   8.804 +  G\<turnstile>Method new declared_in (declclass new) \<and>   
   8.805 +  G\<turnstile>Method old declared_in (declclass old)"
   8.806 +by (induct set: overridesR) (auto intro: trancl_trans ws_widen_trans)
   8.807 +
   8.808 +lemma stat_overrides_commonD:
   8.809 +"\<lbrakk>G\<turnstile>new overrides\<^sub>S old\<rbrakk> \<Longrightarrow>  
   8.810 +  declclass new \<noteq> Object \<and> \<not> is_static new \<and> msig new = msig old \<and> 
   8.811 +  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
   8.812 +  G\<turnstile>Method new declared_in (declclass new) \<and>   
   8.813 +  G\<turnstile>Method old declared_in (declclass old)"
   8.814 +by (induct set: stat_overridesR) (auto intro: trancl_trans)
   8.815 +
   8.816 +lemma overrides_eq_sigD: 
   8.817 + "\<lbrakk>G\<turnstile>new overrides old\<rbrakk> \<Longrightarrow> msig old=msig new"
   8.818 +by (auto dest: overrides_commonD)
   8.819 +
   8.820 +lemma hides_eq_sigD: 
   8.821 + "\<lbrakk>G\<turnstile>new hides old\<rbrakk> \<Longrightarrow> msig old=msig new"
   8.822 +by (auto simp add: hides_def)
   8.823 +
   8.824 +subsubsection "permits access" 
   8.825 +constdefs 
   8.826 +permits_acc:: 
   8.827 + "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   8.828 +                   ("_ \<turnstile> _ in _ permits'_acc'_to _" [61,61,61,61] 60)
   8.829 +
   8.830 +"G\<turnstile>membr in class permits_acc_to accclass 
   8.831 +  \<equiv> (case (accmodi membr) of
   8.832 +       Private   \<Rightarrow> (declclass membr = accclass)
   8.833 +     | Package   \<Rightarrow> (pid (declclass membr) = pid accclass)
   8.834 +     | Protected \<Rightarrow> (pid (declclass membr) = pid accclass)
   8.835 +                    \<or>
   8.836 +                    (G\<turnstile>accclass \<prec>\<^sub>C declclass membr \<and> G\<turnstile>class \<preceq>\<^sub>C accclass) 
   8.837 +     | Public    \<Rightarrow> True)"
   8.838 +text {*
   8.839 +The subcondition of the @{term "Protected"} case: 
   8.840 +@{term "G\<turnstile>accclass \<prec>\<^sub>C declclass membr"} could also be relaxed to:
   8.841 +@{term "G\<turnstile>accclass \<preceq>\<^sub>C declclass membr"} since in case both classes are the
   8.842 +same the other condition @{term "(pid (declclass membr) = pid accclass)"}
   8.843 +holds anyway.
   8.844 +*} 
   8.845 +
   8.846 +text {* Like in case of overriding, the static and dynamic accessibility 
   8.847 +of members is not uniform.
   8.848 +\begin{itemize}
   8.849 +\item Statically the class/interface of the member must be accessible for the
   8.850 +      member to be accessible. During runtime this is not necessary. For
   8.851 +      Example, if a class is accessible and we are allowed to access a member
   8.852 +      of this class (statically) we expect that we can access this member in 
   8.853 +      an arbitrary subclass (during runtime). It's not intended to restrict
   8.854 +      the access to accessible subclasses during runtime.
   8.855 +\item Statically the member we want to access must be "member of" the class.
   8.856 +      Dynamically it must only be "member in" the class.
   8.857 +\end{itemize} 
   8.858 +*} 
   8.859 +
   8.860 +
   8.861 +consts
   8.862 +accessible_fromR:: 
   8.863 + "prog \<Rightarrow> qtname \<Rightarrow> ((qtname \<times> memberdecl) \<times>  qtname) set"
   8.864 +
   8.865 +syntax 
   8.866 +accessible_from:: 
   8.867 + "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   8.868 +                   ("_ \<turnstile> _ of _ accessible'_from _" [61,61,61,61] 60)
   8.869 +
   8.870 +translations
   8.871 +"G\<turnstile>membr of cls accessible_from accclass"  
   8.872 + \<rightleftharpoons> "(membr,cls) \<in> accessible_fromR G accclass"
   8.873 +
   8.874 +syntax 
   8.875 +method_accessible_from:: 
   8.876 + "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   8.877 +                   ("_ \<turnstile>Method _ of _ accessible'_from _" [61,61,61,61] 60)
   8.878 +
   8.879 +translations
   8.880 +"G\<turnstile>Method m of cls accessible_from accclass"  
   8.881 + \<rightleftharpoons> "G\<turnstile>methdMembr m of cls accessible_from accclass"  
   8.882 +
   8.883 +syntax 
   8.884 +methd_accessible_from:: 
   8.885 + "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   8.886 +                 ("_ \<turnstile>Methd _ _ of _ accessible'_from _" [61,61,61,61,61] 60)
   8.887 +
   8.888 +translations
   8.889 +"G\<turnstile>Methd s m of cls accessible_from accclass"  
   8.890 + \<rightleftharpoons> "G\<turnstile>(method s m) of cls accessible_from accclass"  
   8.891 +
   8.892 +syntax 
   8.893 +field_accessible_from:: 
   8.894 + "prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   8.895 +                 ("_ \<turnstile>Field _  _ of _ accessible'_from _" [61,61,61,61,61] 60)
   8.896 +
   8.897 +translations
   8.898 +"G\<turnstile>Field fn f of C accessible_from accclass"  
   8.899 + \<rightleftharpoons> "G\<turnstile>(fieldm fn f) of C accessible_from accclass" 
   8.900 +
   8.901 +inductive "accessible_fromR G accclass" intros
   8.902 +immediate:  "\<lbrakk>G\<turnstile>membr member_of class;
   8.903 +              G\<turnstile>(Class class) accessible_in (pid accclass);
   8.904 +              G\<turnstile>membr in class permits_acc_to accclass 
   8.905 +             \<rbrakk> \<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
   8.906 +
   8.907 +overriding: "\<lbrakk>G\<turnstile>membr member_of class;
   8.908 +              G\<turnstile>(Class class) accessible_in (pid accclass);
   8.909 +              membr=(C,mdecl new);
   8.910 +              G\<turnstile>(C,new) overrides\<^sub>S old; 
   8.911 +              G\<turnstile>class \<prec>\<^sub>C sup;
   8.912 +              G\<turnstile>Method old of sup accessible_from accclass
   8.913 +             \<rbrakk>\<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
   8.914 +
   8.915 +consts
   8.916 +dyn_accessible_fromR:: 
   8.917 + "prog \<Rightarrow> qtname \<Rightarrow> ((qtname \<times> memberdecl) \<times>  qtname) set"
   8.918 +
   8.919 +syntax 
   8.920 +dyn_accessible_from:: 
   8.921 + "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   8.922 +                   ("_ \<turnstile> _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
   8.923 +
   8.924 +translations
   8.925 +"G\<turnstile>membr in C dyn_accessible_from accC"  
   8.926 + \<rightleftharpoons> "(membr,C) \<in> dyn_accessible_fromR G accC"
   8.927 +
   8.928 +syntax 
   8.929 +method_dyn_accessible_from:: 
   8.930 + "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   8.931 +                 ("_ \<turnstile>Method _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
   8.932 +
   8.933 +translations
   8.934 +"G\<turnstile>Method m in C dyn_accessible_from accC"  
   8.935 + \<rightleftharpoons> "G\<turnstile>methdMembr m in C dyn_accessible_from accC"  
   8.936 +
   8.937 +syntax 
   8.938 +methd_dyn_accessible_from:: 
   8.939 + "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   8.940 +             ("_ \<turnstile>Methd _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60)
   8.941 +
   8.942 +translations
   8.943 +"G\<turnstile>Methd s m in C dyn_accessible_from accC"  
   8.944 + \<rightleftharpoons> "G\<turnstile>(method s m) in C dyn_accessible_from accC"  
   8.945 +
   8.946 +syntax 
   8.947 +field_dyn_accessible_from:: 
   8.948 + "prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
   8.949 +         ("_ \<turnstile>Field _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60)
   8.950 +
   8.951 +translations
   8.952 +"G\<turnstile>Field fn f in dynC dyn_accessible_from accC"  
   8.953 + \<rightleftharpoons> "G\<turnstile>(fieldm fn f) in dynC dyn_accessible_from accC"
   8.954 +  
   8.955 +(* #### Testet JVM noch über den Bytecodeverifier hinaus ob der
   8.956 + statische Typ accessible ist bevor es den Zugriff erlaubt 
   8.957 + \<longrightarrow> Test mit Reflektion\<dots>
   8.958 +*)
   8.959 +inductive "dyn_accessible_fromR G accclass" intros
   8.960 +immediate:  "\<lbrakk>G\<turnstile>membr member_in class;
   8.961 +              G\<turnstile>membr in class permits_acc_to accclass 
   8.962 +             \<rbrakk> \<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
   8.963 +
   8.964 +overriding: "\<lbrakk>G\<turnstile>membr member_in class;
   8.965 +              membr=(C,mdecl new);
   8.966 +              G\<turnstile>(C,new) overrides old; 
   8.967 +              G\<turnstile>class \<prec>\<^sub>C sup;
   8.968 +              G\<turnstile>Method old in sup dyn_accessible_from accclass
   8.969 +             \<rbrakk>\<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
   8.970 +
   8.971 +
   8.972 +lemma accessible_from_commonD: "G\<turnstile>m of C accessible_from S
   8.973 + \<Longrightarrow> G\<turnstile>m member_of C \<and> G\<turnstile>(Class C) accessible_in (pid S)"
   8.974 +by (auto elim: accessible_fromR.induct)
   8.975 +
   8.976 +lemma declared_not_undeclared:
   8.977 +  "G\<turnstile>m declared_in C \<Longrightarrow> \<not> G\<turnstile> memberid m undeclared_in C"
   8.978 +by (cases m) (auto simp add: declared_in_def undeclared_in_def)
   8.979 +
   8.980 +lemma not_undeclared_declared:
   8.981 +  "\<not> G\<turnstile> membr_id undeclared_in C \<Longrightarrow> (\<exists> m. G\<turnstile>m declared_in C \<and> 
   8.982 +                                           membr_id = memberid m)"
   8.983 +proof -
   8.984 +  assume not_undecl:"\<not> G\<turnstile> membr_id undeclared_in C"
   8.985 +  show ?thesis (is "?P membr_id")
   8.986 +  proof (cases membr_id)
   8.987 +    case (fid vname)
   8.988 +    with not_undecl
   8.989 +    obtain fld where
   8.990 +      "G\<turnstile>fdecl (vname,fld) declared_in C" 
   8.991 +      by (auto simp add: undeclared_in_def declared_in_def
   8.992 +                         cdeclaredfield_def)
   8.993 +    with fid show ?thesis 
   8.994 +      by auto
   8.995 +  next
   8.996 +    case (mid sig)
   8.997 +    with not_undecl
   8.998 +    obtain mthd where
   8.999 +      "G\<turnstile>mdecl (sig,mthd) declared_in C" 
  8.1000 +      by (auto simp add: undeclared_in_def declared_in_def
  8.1001 +                         cdeclaredmethd_def)
  8.1002 +    with mid show ?thesis 
  8.1003 +      by auto
  8.1004 +  qed
  8.1005 +qed
  8.1006 +
  8.1007 +lemma unique_declared_in:
  8.1008 + "\<lbrakk>G\<turnstile>m declared_in C; G\<turnstile>n declared_in C; memberid m = memberid n\<rbrakk>
  8.1009 + \<Longrightarrow> m = n"
  8.1010 +by (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def
  8.1011 +            split: memberdecl.splits)
  8.1012 +
  8.1013 +lemma unique_member_of: 
  8.1014 + (assumes n: "G\<turnstile>n member_of C" and  
  8.1015 +          m: "G\<turnstile>m member_of C" and
  8.1016 +       eqid: "memberid n = memberid m"
  8.1017 + ) "n=m"
  8.1018 +proof -
  8.1019 +  from n m eqid  
  8.1020 +  show "n=m"
  8.1021 +  proof (induct)
  8.1022 +    case (Immediate C n)
  8.1023 +    assume member_n: "G\<turnstile> mbr n declared_in C" "declclass n = C"
  8.1024 +    assume eqid: "memberid n = memberid m"
  8.1025 +    assume "G \<turnstile> m member_of C"
  8.1026 +    then show "n=m"
  8.1027 +    proof (cases)
  8.1028 +      case (Immediate _ m')
  8.1029 +      with eqid 
  8.1030 +      have "m=m'"
  8.1031 +           "memberid n = memberid m" 
  8.1032 +           "G\<turnstile> mbr m declared_in C" 
  8.1033 +           "declclass m = C"
  8.1034 +	by auto
  8.1035 +      with member_n   
  8.1036 +      show ?thesis
  8.1037 +	by (cases n, cases m) 
  8.1038 +           (auto simp add: declared_in_def 
  8.1039 +                           cdeclaredmethd_def cdeclaredfield_def
  8.1040 +                    split: memberdecl.splits)
  8.1041 +    next
  8.1042 +      case (Inherited _ _ m')
  8.1043 +      then have "G\<turnstile> memberid m undeclared_in C"
  8.1044 +	by simp
  8.1045 +      with eqid member_n
  8.1046 +      show ?thesis
  8.1047 +	by (cases n) (auto dest: declared_not_undeclared)
  8.1048 +    qed
  8.1049 +  next
  8.1050 +    case (Inherited C S n)
  8.1051 +    assume undecl: "G\<turnstile> memberid n undeclared_in C"
  8.1052 +    assume  super: "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S"
  8.1053 +    assume    hyp: "\<lbrakk>G \<turnstile> m member_of S; memberid n = memberid m\<rbrakk> \<Longrightarrow> n = m"
  8.1054 +    assume   eqid: "memberid n = memberid m"
  8.1055 +    assume "G \<turnstile> m member_of C"
  8.1056 +    then show "n=m"
  8.1057 +    proof (cases)
  8.1058 +      case Immediate
  8.1059 +      then have "G\<turnstile> mbr m declared_in C" by simp
  8.1060 +      with eqid undecl
  8.1061 +      show ?thesis
  8.1062 +	by (cases m) (auto dest: declared_not_undeclared)
  8.1063 +    next
  8.1064 +      case Inherited 
  8.1065 +      with super have "G \<turnstile> m member_of S"
  8.1066 +	by (auto dest!: subcls1D)
  8.1067 +      with eqid hyp
  8.1068 +      show ?thesis 
  8.1069 +	by blast
  8.1070 +    qed
  8.1071 +  qed
  8.1072 +qed
  8.1073 +
  8.1074 +lemma member_of_is_classD: "G\<turnstile>m member_of C \<Longrightarrow> is_class G C"
  8.1075 +proof (induct set: members)
  8.1076 +  case (Immediate C m)
  8.1077 +  assume "G\<turnstile> mbr m declared_in C"
  8.1078 +  then show "is_class G C"
  8.1079 +    by (cases "mbr m")
  8.1080 +       (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)
  8.1081 +next
  8.1082 +  case (Inherited C S m)  
  8.1083 +  assume "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S" and "is_class G S"
  8.1084 +  then show "is_class G C"
  8.1085 +    by - (rule subcls_is_class2,auto)
  8.1086 +qed
  8.1087 +
  8.1088 +lemma member_of_declC: 
  8.1089 + "G\<turnstile>m member_of C 
  8.1090 +  \<Longrightarrow> G\<turnstile>mbr m declared_in (declclass m)"
  8.1091 +by (induct set: members) auto
  8.1092 +
  8.1093 +lemma member_of_member_of_declC:
  8.1094 + "G\<turnstile>m member_of C 
  8.1095 +  \<Longrightarrow> G\<turnstile>m member_of (declclass m)"
  8.1096 +by (auto dest: member_of_declC intro: members.Immediate)
  8.1097 +
  8.1098 +lemma member_of_class_relation:
  8.1099 +  "G\<turnstile>m member_of C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
  8.1100 +proof (induct set: members)
  8.1101 +  case (Immediate C m)
  8.1102 +  then show "G\<turnstile>C \<preceq>\<^sub>C declclass m" by simp
  8.1103 +next
  8.1104 +  case (Inherited C S m)
  8.1105 +  then show "G\<turnstile>C \<preceq>\<^sub>C declclass m" 
  8.1106 +    by (auto dest: r_into_rtrancl intro: rtrancl_trans)
  8.1107 +qed
  8.1108 +
  8.1109 +lemma member_in_class_relation:
  8.1110 +  "G\<turnstile>m member_in C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
  8.1111 +by (auto dest: member_inD member_of_class_relation
  8.1112 +        intro: rtrancl_trans)
  8.1113 +
  8.1114 +lemma member_of_Package: 
  8.1115 + "\<lbrakk>G\<turnstile>m member_of C; accmodi m = Package\<rbrakk> 
  8.1116 +  \<Longrightarrow> pid (declclass m) = pid C" 
  8.1117 +proof -
  8.1118 +  assume   member: "G\<turnstile>m member_of C"
  8.1119 +  then show " accmodi m = Package \<Longrightarrow> ?thesis" (is "PROP ?P m C")
  8.1120 +  proof (induct rule: members.induct)
  8.1121 +    fix C m
  8.1122 +    assume     C: "declclass m = C"
  8.1123 +    then show "pid (declclass m) = pid C"
  8.1124 +      by simp
  8.1125 +  next
  8.1126 +    fix C S m  
  8.1127 +    assume inheritable: "G \<turnstile> m inheritable_in pid C"
  8.1128 +    assume         hyp: "PROP ?P m S" and
  8.1129 +           package_acc: "accmodi m = Package" 
  8.1130 +    with inheritable package_acc hyp
  8.1131 +    show "pid (declclass m) = pid C" 
  8.1132 +      by (auto simp add: inheritable_in_def)
  8.1133 +  qed
  8.1134 +qed
  8.1135 +
  8.1136 +lemma dyn_accessible_from_commonD: "G\<turnstile>m in C dyn_accessible_from S
  8.1137 + \<Longrightarrow> G\<turnstile>m member_in C"
  8.1138 +by (auto elim: dyn_accessible_fromR.induct)
  8.1139 +
  8.1140 +(* ### Gilt nicht für wf_progs!dynmaisches Override, 
  8.1141 +  da die accmodi Bedingung nur für stat override gilt! *)
  8.1142 +(*
  8.1143 +lemma override_Package:
  8.1144 + "\<lbrakk>G\<turnstile>new overrides old; 
  8.1145 +  \<And> new old. G\<turnstile>new overrides old \<Longrightarrow> accmodi old \<le> accmodi new;
  8.1146 +  accmodi old = Package; accmodi new = Package\<rbrakk>
  8.1147 +  \<Longrightarrow> pid (declclass old) = pid (declclass new)"
  8.1148 +proof - 
  8.1149 +  assume wf: "\<And> new old. G\<turnstile>new overrides old \<Longrightarrow> accmodi old \<le> accmodi new"
  8.1150 +  assume ovverride: "G\<turnstile>new overrides old"
  8.1151 +  then show "\<lbrakk>accmodi old = Package;accmodi new = Package\<rbrakk> \<Longrightarrow> ?thesis"
  8.1152 +    (is "?Pack old \<Longrightarrow> ?Pack new \<Longrightarrow> ?EqPid old new")
  8.1153 +  proof (induct rule: overridesR.induct)
  8.1154 +    case Direct
  8.1155 +    fix new old
  8.1156 +    assume "accmodi old = Package"
  8.1157 +           "G \<turnstile> methdMembr old inheritable_in pid (declclass new)"
  8.1158 +    then show "pid (declclass old) =  pid (declclass new)"
  8.1159 +      by (auto simp add: inheritable_in_def)
  8.1160 +  next
  8.1161 +    case (Indirect inter new old)
  8.1162 +    assume accmodi_old: "accmodi old = Package" and
  8.1163 +           accmodi_new: "accmodi new = Package" 
  8.1164 +    assume "G \<turnstile> new overrides inter"
  8.1165 +    with wf have le_inter_new: "accmodi inter \<le> accmodi new"
  8.1166 +      by blast
  8.1167 +    assume "G \<turnstile> inter overrides old"
  8.1168 +    with wf have le_old_inter: "accmodi old \<le> accmodi inter"
  8.1169 +      by blast
  8.1170 +    from accmodi_old accmodi_new le_inter_new le_old_inter
  8.1171 +    have "accmodi inter = Package"
  8.1172 +      by(auto simp add: le_acc_def less_acc_def)
  8.1173 +    with Indirect accmodi_old accmodi_new
  8.1174 +    show "?EqPid old new"
  8.1175 +      by auto
  8.1176 +  qed
  8.1177 +qed
  8.1178 +
  8.1179 +lemma stat_override_Package:
  8.1180 + "\<lbrakk>G\<turnstile>new overrides\<^sub>S old; 
  8.1181 +  \<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new;
  8.1182 +  accmodi old = Package; accmodi new = Package\<rbrakk>
  8.1183 +  \<Longrightarrow> pid (declclass old) = pid (declclass new)"
  8.1184 +proof - 
  8.1185 +  assume wf: "\<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new"
  8.1186 +  assume ovverride: "G\<turnstile>new overrides\<^sub>S old"
  8.1187 +  then show "\<lbrakk>accmodi old = Package;accmodi new = Package\<rbrakk> \<Longrightarrow> ?thesis"
  8.1188 +    (is "?Pack old \<Longrightarrow> ?Pack new \<Longrightarrow> ?EqPid old new")
  8.1189 +  proof (induct rule: stat_overridesR.induct)
  8.1190 +    case Direct
  8.1191 +    fix new old
  8.1192 +    assume "accmodi old = Package"
  8.1193 +           "G \<turnstile> methdMembr old inheritable_in pid (declclass new)"
  8.1194 +    then show "pid (declclass old) =  pid (declclass new)"
  8.1195 +      by (auto simp add: inheritable_in_def)
  8.1196 +  next
  8.1197 +    case (Indirect inter new old)
  8.1198 +    assume accmodi_old: "accmodi old = Package" and
  8.1199 +           accmodi_new: "accmodi new = Package" 
  8.1200 +    assume "G \<turnstile> new overrides\<^sub>S inter"
  8.1201 +    with wf have le_inter_new: "accmodi inter \<le> accmodi new"
  8.1202 +      by blast
  8.1203 +    assume "G \<turnstile> inter overrides\<^sub>S old"
  8.1204 +    with wf have le_old_inter: "accmodi old \<le> accmodi inter"
  8.1205 +      by blast
  8.1206 +    from accmodi_old accmodi_new le_inter_new le_old_inter
  8.1207 +    have "accmodi inter = Package"
  8.1208 +      by(auto simp add: le_acc_def less_acc_def)
  8.1209 +    with Indirect accmodi_old accmodi_new
  8.1210 +    show "?EqPid old new"
  8.1211 +      by auto
  8.1212 +  qed
  8.1213 +qed
  8.1214 +
  8.1215 +*)
  8.1216 +lemma no_Private_stat_override: 
  8.1217 + "\<lbrakk>G\<turnstile>new overrides\<^sub>S old\<rbrakk> \<Longrightarrow> accmodi old \<noteq> Private"
  8.1218 +by (induct set:  stat_overridesR) (auto simp add: inheritable_in_def)
  8.1219 +
  8.1220 +lemma no_Private_override: "\<lbrakk>G\<turnstile>new overrides old\<rbrakk> \<Longrightarrow> accmodi old \<noteq> Private"
  8.1221 +by (induct set: overridesR) (auto simp add: inheritable_in_def)
  8.1222 +
  8.1223 +lemma permits_acc_inheritance:
  8.1224 + "\<lbrakk>G\<turnstile>m in statC permits_acc_to accC; G\<turnstile>dynC \<preceq>\<^sub>C statC
  8.1225 +  \<rbrakk> \<Longrightarrow> G\<turnstile>m in dynC permits_acc_to accC"
  8.1226 +by (cases "accmodi m")
  8.1227 +   (auto simp add: permits_acc_def
  8.1228 +            intro: subclseq_trans) 
  8.1229 +
  8.1230 +lemma field_accessible_fromD:
  8.1231 + "\<lbrakk>G\<turnstile>membr of C accessible_from accC;is_field membr\<rbrakk> 
  8.1232 +  \<Longrightarrow> G\<turnstile>membr member_of C \<and>
  8.1233 +      G\<turnstile>(Class C) accessible_in (pid accC) \<and>
  8.1234 +      G\<turnstile>membr in C permits_acc_to accC"
  8.1235 +by (cases set: accessible_fromR)
  8.1236 +   (auto simp add: is_field_def split: memberdecl.splits) 
  8.1237 +
  8.1238 +lemma field_accessible_from_permits_acc_inheritance:
  8.1239 +"\<lbrakk>G\<turnstile>membr of statC accessible_from accC; is_field membr; G \<turnstile> dynC \<preceq>\<^sub>C statC\<rbrakk>
  8.1240 +\<Longrightarrow> G\<turnstile>membr in dynC permits_acc_to accC"
  8.1241 +by (auto dest: field_accessible_fromD intro: permits_acc_inheritance)
  8.1242 +
  8.1243 +
  8.1244 +(*
  8.1245 +lemma accessible_Package:
  8.1246 + "\<lbrakk>G \<turnstile> m of C accessible_from S;accmodi m = Package;
  8.1247 +   \<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new\<rbrakk>
  8.1248 +  \<Longrightarrow> pid S = pid C \<and> pid C = pid (declclass m)"
  8.1249 +proof -
  8.1250 +  assume wf: "\<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new"
  8.1251 +  assume "G \<turnstile> m of C accessible_from S"
  8.1252 +  then show "accmodi m = Package \<Longrightarrow> pid S = pid C \<and> pid C = pid (declclass m)"
  8.1253 +    (is "?Pack m \<Longrightarrow> ?P C m")
  8.1254 +  proof (induct rule: accessible_fromR.induct)
  8.1255 +    fix C m
  8.1256 +    assume "G\<turnstile>m member_of C"
  8.1257 +           "G \<turnstile> m in C permits_acc_to S"
  8.1258 +           "accmodi m = Package"      
  8.1259 +    then show "?P C m"
  8.1260 +      by (auto dest: member_of_Package simp add: permits_acc_def)
  8.1261 +  next
  8.1262 +    fix declC C new newm old Sup
  8.1263 +    assume member_new: "G \<turnstile> new member_of C" and 
  8.1264 +                acc_C: "G \<turnstile> Class C accessible_in pid S" and
  8.1265 +                  new: "new = (declC, mdecl newm)" and
  8.1266 +             override: "G \<turnstile> (declC, newm) overrides\<^sub>S old" and
  8.1267 +         subcls_C_Sup: "G\<turnstile>C \<prec>\<^sub>C Sup" and
  8.1268 +              acc_old: "G \<turnstile> methdMembr old of Sup accessible_from S" and
  8.1269 +                  hyp: "?Pack (methdMembr old) \<Longrightarrow> ?P Sup (methdMembr old)" and
  8.1270 +          accmodi_new: "accmodi new = Package"
  8.1271 +    from override wf 
  8.1272 +    have accmodi_weaken: "accmodi old \<le> accmodi newm"
  8.1273 +      by (cases old,cases newm) auto
  8.1274 +    from override new
  8.1275 +    have "accmodi old \<noteq> Private"
  8.1276 +      by (simp add: no_Private_stat_override)
  8.1277 +    with accmodi_weaken accmodi_new new
  8.1278 +    have accmodi_old: "accmodi old = Package"
  8.1279 +      by (cases "accmodi old") (auto simp add: le_acc_def less_acc_def) 
  8.1280 +    with hyp 
  8.1281 +    have P_sup: "?P Sup (methdMembr old)"
  8.1282 +      by (simp)
  8.1283 +    from wf override new accmodi_old accmodi_new
  8.1284 +    have eq_pid_new_old: "pid (declclass new) = pid (declclass old)"
  8.1285 +      by (auto dest: stat_override_Package) 
  8.1286 +    from member_new accmodi_new
  8.1287 +    have "pid (declclass new) = pid C"
  8.1288 +      by (auto dest: member_of_Package)
  8.1289 +    with eq_pid_new_old P_sup show "?P C new"
  8.1290 +      by auto
  8.1291 +  qed
  8.1292 +qed
  8.1293 +*)
  8.1294 +lemma accessible_fieldD: 
  8.1295 + "\<lbrakk>G\<turnstile>membr of C accessible_from accC; is_field membr\<rbrakk>
  8.1296 + \<Longrightarrow> G\<turnstile>membr member_of C \<and>
  8.1297 +     G\<turnstile>(Class C) accessible_in (pid accC) \<and>
  8.1298 +     G\<turnstile>membr in C permits_acc_to accC"
  8.1299 +by (induct rule: accessible_fromR.induct) (auto dest: is_fieldD)
  8.1300 +      
  8.1301 +(* lemmata:
  8.1302 + Wegen  G\<turnstile>Super accessible_in (pid C) folgt:
  8.1303 +  G\<turnstile>m declared_in C; G\<turnstile>m member_of D; accmodi m = Package (G\<turnstile>D \<preceq>\<^sub>C C)
  8.1304 +  \<Longrightarrow> pid C = pid D 
  8.1305 +
  8.1306 +  C package
  8.1307 +  m public in C
  8.1308 +  für alle anderen D: G\<turnstile>m undeclared_in C
  8.1309 +  m wird in alle subklassen vererbt, auch aus dem Package heraus!
  8.1310 +
  8.1311 +  G\<turnstile>m member_of C \<Longrightarrow> \<exists> D. G\<turnstile>C \<preceq>\<^sub>C D \<and> G\<turnstile>m declared_in D
  8.1312 +*)
  8.1313 +
  8.1314 +(* Begriff (C,m) overrides (D,m)
  8.1315 +    3 Fälle: Direkt,
  8.1316 +             Indirekt über eine Zwischenklasse (ohne weiteres override)
  8.1317 +             Indirekt über override
  8.1318 +*)
  8.1319 +   
  8.1320 +(*
  8.1321 +"G\<turnstile>m member_of C \<equiv> 
  8.1322 +constdefs declares_method:: "prog \<Rightarrow> sig \<Rightarrow> qtname \<Rightarrow> methd \<Rightarrow> bool"
  8.1323 +                                 ("_,_\<turnstile> _ declares'_method _" [61,61,61,61] 60)
  8.1324 +"G,sig\<turnstile>C declares_method m \<equiv> cdeclaredmethd G C sig = Some m" 
  8.1325 +
  8.1326 +constdefs is_declared:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool"
  8.1327 +"is_declared G sig em \<equiv> G,sig\<turnstile>declclass em declares_method mthd em"
  8.1328 +*)
  8.1329 +
  8.1330 +lemma member_of_Private:
  8.1331 +"\<lbrakk>G\<turnstile>m member_of C; accmodi m = Private\<rbrakk> \<Longrightarrow> declclass m = C"
  8.1332 +by (induct set: members) (auto simp add: inheritable_in_def)
  8.1333 +
  8.1334 +lemma member_of_subclseq_declC:
  8.1335 +  "G\<turnstile>m member_of C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
  8.1336 +by (induct set: members) (auto dest: r_into_rtrancl intro: rtrancl_trans)
  8.1337 +
  8.1338 +lemma member_of_inheritance:
  8.1339 +  (assumes    m: "G\<turnstile>m member_of D" and 
  8.1340 +     subclseq_D_C: "G\<turnstile>D \<preceq>\<^sub>C C" and
  8.1341 +     subclseq_C_m: "G\<turnstile>C \<preceq>\<^sub>C declclass m" and
  8.1342 +               ws: "ws_prog G" 
  8.1343 +  ) "G\<turnstile>m member_of C"
  8.1344 +proof -
  8.1345 +  from m subclseq_D_C subclseq_C_m
  8.1346 +  show ?thesis
  8.1347 +  proof (induct)
  8.1348 +    case (Immediate D m)
  8.1349 +    assume "declclass m = D" and
  8.1350 +           "G\<turnstile>D\<preceq>\<^sub>C C" and "G\<turnstile>C\<preceq>\<^sub>C declclass m"
  8.1351 +    with ws have "D=C" 
  8.1352 +      by (auto intro: subclseq_acyclic)
  8.1353 +    with Immediate 
  8.1354 +    show "G\<turnstile>m member_of C"
  8.1355 +      by (auto intro: members.Immediate)
  8.1356 +  next
  8.1357 +    case (Inherited D S m)
  8.1358 +    assume member_of_D_props: 
  8.1359 +            "G \<turnstile> m inheritable_in pid D" 
  8.1360 +            "G\<turnstile> memberid m undeclared_in D"  
  8.1361 +            "G \<turnstile> Class S accessible_in pid D" 
  8.1362 +            "G \<turnstile> m member_of S"
  8.1363 +    assume super: "G\<turnstile>D\<prec>\<^sub>C\<^sub>1S"
  8.1364 +    assume hyp: "\<lbrakk>G\<turnstile>S\<preceq>\<^sub>C C; G\<turnstile>C\<preceq>\<^sub>C declclass m\<rbrakk> \<Longrightarrow> G \<turnstile> m member_of C"
  8.1365 +    assume subclseq_C_m: "G\<turnstile>C\<preceq>\<^sub>C declclass m"
  8.1366 +    assume "G\<turnstile>D\<preceq>\<^sub>C C"
  8.1367 +    then show "G\<turnstile>m member_of C"
  8.1368 +    proof (cases rule: subclseq_cases)
  8.1369 +      case Eq
  8.1370 +      assume "D=C" 
  8.1371 +      with super member_of_D_props 
  8.1372 +      show ?thesis
  8.1373 +	by (auto intro: members.Inherited)
  8.1374 +    next
  8.1375 +      case Subcls
  8.1376 +      assume "G\<turnstile>D\<prec>\<^sub>C C"
  8.1377 +      with super 
  8.1378 +      have "G\<turnstile>S\<preceq>\<^sub>C C"
  8.1379 +	by (auto dest: subcls1D subcls_superD)
  8.1380 +      with subclseq_C_m hyp show ?thesis
  8.1381 +	by blast
  8.1382 +    qed
  8.1383 +  qed
  8.1384 +qed
  8.1385 +
  8.1386 +lemma member_of_subcls:
  8.1387 +  (assumes    old: "G\<turnstile>old member_of C" and 
  8.1388 +              new: "G\<turnstile>new member_of D" and
  8.1389 +             eqid: "memberid new = memberid old" and
  8.1390 +     subclseq_D_C: "G\<turnstile>D \<preceq>\<^sub>C C" and 
  8.1391 +   subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old" and
  8.1392 +               ws: "ws_prog G"
  8.1393 +  ) "G\<turnstile>D \<prec>\<^sub>C C"
  8.1394 +proof -
  8.1395 +  from old 
  8.1396 +  have subclseq_C_old: "G\<turnstile>C \<preceq>\<^sub>C declclass old"
  8.1397 +    by (auto dest: member_of_subclseq_declC)
  8.1398 +  from new 
  8.1399 +  have subclseq_D_new: "G\<turnstile>D \<preceq>\<^sub>C declclass new"
  8.1400 +    by (auto dest: member_of_subclseq_declC)
  8.1401 +  from subcls_new_old ws
  8.1402 +  have neq_new_old: "new\<noteq>old"
  8.1403 +    by (cases new,cases old) (auto dest: subcls_irrefl)
  8.1404 +  from subclseq_D_new subclseq_D_C
  8.1405 +  have "G\<turnstile>(declclass new) \<preceq>\<^sub>C C \<or> G\<turnstile>C \<preceq>\<^sub>C (declclass new)" 
  8.1406 +    by (rule subcls_compareable)
  8.1407 +  then have "G\<turnstile>(declclass new) \<preceq>\<^sub>C C"
  8.1408 +  proof
  8.1409 +    assume "G\<turnstile>declclass new\<preceq>\<^sub>C C" then show ?thesis .
  8.1410 +  next
  8.1411 +    assume "G\<turnstile>C \<preceq>\<^sub>C (declclass new)"
  8.1412 +    with new subclseq_D_C ws 
  8.1413 +    have "G\<turnstile>new member_of C"
  8.1414 +      by (blast intro: member_of_inheritance)
  8.1415 +    with eqid old 
  8.1416 +    have "new=old"
  8.1417 +      by (blast intro: unique_member_of)
  8.1418 +    with neq_new_old 
  8.1419 +    show ?thesis
  8.1420 +      by contradiction
  8.1421 +  qed
  8.1422 +  then show ?thesis
  8.1423 +  proof (cases rule: subclseq_cases)
  8.1424 +    case Eq
  8.1425 +    assume "declclass new = C"
  8.1426 +    with new have "G\<turnstile>new member_of C"
  8.1427 +      by (auto dest: member_of_member_of_declC)
  8.1428 +    with eqid old 
  8.1429 +    have "new=old"
  8.1430 +      by (blast intro: unique_member_of)
  8.1431 +    with neq_new_old 
  8.1432 +    show ?thesis
  8.1433 +      by contradiction
  8.1434 +  next
  8.1435 +    case Subcls
  8.1436 +    assume "G\<turnstile>declclass new\<prec>\<^sub>C C"
  8.1437 +    with subclseq_D_new
  8.1438 +    show "G\<turnstile>D\<prec>\<^sub>C C"
  8.1439 +      by (rule rtrancl_trancl_trancl)
  8.1440 +  qed
  8.1441 +qed
  8.1442 +
  8.1443 +corollary member_of_overrides_subcls:
  8.1444 + "\<lbrakk>G\<turnstile>Methd sig old member_of C; G\<turnstile>Methd sig new member_of D;G\<turnstile>D \<preceq>\<^sub>C C;
  8.1445 +   G,sig\<turnstile>new overrides old; ws_prog G\<rbrakk>
  8.1446 + \<Longrightarrow> G\<turnstile>D \<prec>\<^sub>C C"
  8.1447 +by (drule overrides_commonD) (auto intro: member_of_subcls)    
  8.1448 +
  8.1449 +corollary member_of_stat_overrides_subcls:
  8.1450 + "\<lbrakk>G\<turnstile>Methd sig old member_of C; G\<turnstile>Methd sig new member_of D;G\<turnstile>D \<preceq>\<^sub>C C;
  8.1451 +   G,sig\<turnstile>new overrides\<^sub>S old; ws_prog G\<rbrakk>
  8.1452 + \<Longrightarrow> G\<turnstile>D \<prec>\<^sub>C C"
  8.1453 +by (drule stat_overrides_commonD) (auto intro: member_of_subcls)    
  8.1454 +
  8.1455 +
  8.1456 +
  8.1457 +lemma inherited_field_access: 
  8.1458 + (assumes stat_acc: "G\<turnstile>membr of statC accessible_from accC" and
  8.1459 +          is_field: "is_field membr" and 
  8.1460 +          subclseq: "G \<turnstile> dynC \<preceq>\<^sub>C statC"
  8.1461 + ) "G\<turnstile>membr in dynC dyn_accessible_from accC"
  8.1462 +proof -
  8.1463 +  from stat_acc is_field subclseq 
  8.1464 +  show ?thesis
  8.1465 +    by (auto dest: accessible_fieldD 
  8.1466 +            intro: dyn_accessible_fromR.immediate 
  8.1467 +                   member_inI 
  8.1468 +                   permits_acc_inheritance)
  8.1469 +qed
  8.1470 +
  8.1471 +lemma accessible_inheritance:
  8.1472 + (assumes stat_acc: "G\<turnstile>m of statC accessible_from accC" and
  8.1473 +          subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
  8.1474 +       member_dynC: "G\<turnstile>m member_of dynC" and
  8.1475 +          dynC_acc: "G\<turnstile>(Class dynC) accessible_in (pid accC)"
  8.1476 + ) "G\<turnstile>m of dynC accessible_from accC"
  8.1477 +proof -
  8.1478 +  from stat_acc
  8.1479 +  have member_statC: "G\<turnstile>m member_of statC" 
  8.1480 +    by (auto dest: accessible_from_commonD)
  8.1481 +  from stat_acc
  8.1482 +  show ?thesis
  8.1483 +  proof (cases)
  8.1484 +    case immediate
  8.1485 +    with member_dynC member_statC subclseq dynC_acc
  8.1486 +    show ?thesis
  8.1487 +      by (auto intro: accessible_fromR.immediate permits_acc_inheritance)
  8.1488 +  next
  8.1489 +    case overriding
  8.1490 +    with member_dynC subclseq dynC_acc
  8.1491 +    show ?thesis
  8.1492 +      by (auto intro: accessible_fromR.overriding rtrancl_trancl_trancl)
  8.1493 +  qed
  8.1494 +qed
  8.1495 +
  8.1496 +section "fields and methods"
  8.1497 +
  8.1498 +
  8.1499 +types
  8.1500 +  fspec = "vname \<times> qtname"
  8.1501 +
  8.1502 +translations 
  8.1503 +  "fspec" <= (type) "vname \<times> qtname" 
  8.1504 +
  8.1505 +constdefs
  8.1506 +imethds:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
  8.1507 +"imethds G I 
  8.1508 +  \<equiv> iface_rec (G,I)  
  8.1509 +              (\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus> 
  8.1510 +                        (o2s \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
  8.1511 +text {* methods of an interface, with overriding and inheritance, cf. 9.2 *}
  8.1512 +
  8.1513 +constdefs
  8.1514 +accimethds:: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
  8.1515 +"accimethds G pack I
  8.1516 +  \<equiv> if G\<turnstile>Iface I accessible_in pack 
  8.1517 +       then imethds G I
  8.1518 +       else \<lambda> k. {}"
  8.1519 +text {* only returns imethds if the interface is accessible *}
  8.1520 +
  8.1521 +constdefs
  8.1522 +methd:: "prog \<Rightarrow> qtname  \<Rightarrow> (sig,qtname \<times> methd) table"
  8.1523 +
  8.1524 +"methd G C 
  8.1525 + \<equiv> class_rec (G,C) empty
  8.1526 +             (\<lambda>C c subcls_mthds. 
  8.1527 +               filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)
  8.1528 +                          subcls_mthds 
  8.1529 +               ++ 
  8.1530 +               table_of (map (\<lambda>(s,m). (s,C,m)) (methods c)))"
  8.1531 +text {* @{term "methd G C"}: methods of a class C (statically visible from C), 
  8.1532 +     with inheritance and hiding cf. 8.4.6;
  8.1533 +     Overriding is captured by @{text dynmethd}.
  8.1534 +     Every new method with the same signature coalesces the
  8.1535 +     method of a superclass. *}
  8.1536 +
  8.1537 +constdefs                      
  8.1538 +accmethd:: "prog \<Rightarrow> qtname \<Rightarrow> qtname  \<Rightarrow> (sig,qtname \<times> methd) table"
  8.1539 +"accmethd G S C 
  8.1540 + \<equiv> filter_tab (\<lambda>sig m. G\<turnstile>method sig m of C accessible_from S) 
  8.1541 +              (methd G C)"
  8.1542 +text {* @{term "accmethd G S C"}: only those methods of @{term "methd G C"}, 
  8.1543 +        accessible from S *}
  8.1544 +
  8.1545 +text {* Note the class component in the accessibility filter. The class where
  8.1546 +    method @{term m} is declared (@{term declC}) isn't necessarily accessible 
  8.1547 +    from the current scope @{term S}. The method can be made accessible 
  8.1548 +    through inheritance, too.
  8.1549 +    So we must test accessibility of method @{term m} of class @{term C} 
  8.1550 +    (not @{term "declclass m"}) *}
  8.1551 +
  8.1552 +constdefs 
  8.1553 +dynmethd:: "prog  \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
  8.1554 +"dynmethd G statC dynC  
  8.1555 +  \<equiv> \<lambda> sig. 
  8.1556 +     (if G\<turnstile>dynC \<preceq>\<^sub>C statC
  8.1557 +        then (case methd G statC sig of
  8.1558 +                None \<Rightarrow> None
  8.1559 +              | Some statM 
  8.1560 +                  \<Rightarrow> (class_rec (G,dynC) empty
  8.1561 +                       (\<lambda>C c subcls_mthds. 
  8.1562 +                          subcls_mthds
  8.1563 +                          ++
  8.1564 +                          (filter_tab 
  8.1565 +                            (\<lambda> _ dynM. G,sig\<turnstile>dynM overrides statM \<or> dynM=statM)
  8.1566 +                            (methd G C) ))
  8.1567 +                      ) sig
  8.1568 +              )
  8.1569 +        else None)"
  8.1570 +(*
  8.1571 +"dynmethd G statC dynC  
  8.1572 +  \<equiv> \<lambda> sig. 
  8.1573 +     (if G\<turnstile>dynC \<preceq>\<^sub>C statC
  8.1574 +        then (case methd G statC sig of
  8.1575 +                None \<Rightarrow> None
  8.1576 +              | Some statM 
  8.1577 +                    \<Rightarrow> (class_rec (G,statC) empty
  8.1578 +                         (\<lambda>C c subcls_mthds. 
  8.1579 +                            subcls_mthds
  8.1580 +                            ++
  8.1581 +                            (filter_tab 
  8.1582 +                              (\<lambda> _ dynM. G,sig\<turnstile>dynM overrides statM)  
  8.1583 +                              (table_of (map (\<lambda>(s,m). (s,C,m)) (methods c)))))
  8.1584 +                        ) sig
  8.1585 +              )
  8.1586 +        else None)"*)
  8.1587 +text {* @{term "dynmethd G statC dynC"}: dynamic method lookup of a reference 
  8.1588 +        with dynamic class @{term dynC} and static class @{term statC} *}
  8.1589 +text {* Note some kind of duality between @{term methd} and @{term dynmethd} 
  8.1590 +        in the @{term class_rec} arguments. Whereas @{term methd} filters the 
  8.1591 +        subclass methods (to get only the inherited ones), @{term dynmethd} 
  8.1592 +        filters the new methods (to get only those methods which actually
  8.1593 +        override the methods of the static class) *}
  8.1594 +
  8.1595 +constdefs 
  8.1596 +dynimethd:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
  8.1597 +"dynimethd G I dynC 
  8.1598 +  \<equiv> \<lambda> sig. if imethds G I sig \<noteq> {}
  8.1599 +               then methd G dynC sig
  8.1600 +               else dynmethd G Object dynC sig"
  8.1601 +text {* @{term "dynimethd G I dynC"}: dynamic method lookup of a reference with 
  8.1602 +        dynamic class dynC and static interface type I *}
  8.1603 +text {* 
  8.1604 +   When calling an interface method, we must distinguish if the method signature
  8.1605 +   was defined in the interface or if it must be an Object method in the other
  8.1606 +   case. If it was an interface method we search the class hierarchy
  8.1607 +   starting at the dynamic class of the object up to Object to find the 
  8.1608 +   first matching method (@{term methd}). Since all interface methods have 
  8.1609 +   public access the method can't be coalesced due to some odd visibility 
  8.1610 +   effects like in case of dynmethd. The method will be inherited or 
  8.1611 +   overridden in all classes from the first class implementing the interface 
  8.1612 +   down to the actual dynamic class.
  8.1613 + *}
  8.1614 +
  8.1615 +constdefs
  8.1616 +dynlookup::"prog  \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
  8.1617 +"dynlookup G statT dynC
  8.1618 +  \<equiv> (case statT of
  8.1619 +       NullT        \<Rightarrow> empty
  8.1620 +     | IfaceT I     \<Rightarrow> dynimethd G I      dynC
  8.1621 +     | ClassT statC \<Rightarrow> dynmethd  G statC  dynC
  8.1622 +     | ArrayT ty    \<Rightarrow> dynmethd  G Object dynC)"
  8.1623 +text {* @{term "dynlookup G statT dynC"}: dynamic lookup of a method within the 
  8.1624 +    static reference type statT and the dynamic class dynC. 
  8.1625 +    In a wellformd context statT will not be NullT and in case
  8.1626 +    statT is an array type, dynC=Object *}
  8.1627 +
  8.1628 +constdefs
  8.1629 +fields:: "prog \<Rightarrow> qtname \<Rightarrow> ((vname \<times> qtname) \<times> field) list"
  8.1630 +"fields G C 
  8.1631 +  \<equiv> class_rec (G,C) [] (\<lambda>C c ts. map (\<lambda>(n,t). ((n,C),t)) (cfields c) @ ts)"
  8.1632 +text {* @{term "fields G C"} 
  8.1633 +     list of fields of a class, including all the fields of the superclasses
  8.1634 +     (private, inherited and hidden ones) not only the accessible ones
  8.1635 +     (an instance of a object allocates all these fields *}
  8.1636 +
  8.1637 +constdefs
  8.1638 +accfield:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (vname, qtname  \<times>  field) table"
  8.1639 +"accfield G S C
  8.1640 +  \<equiv> let field_tab = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (fields G C))
  8.1641 +    in filter_tab (\<lambda>n (declC,f). G\<turnstile> (declC,fdecl (n,f)) of C accessible_from S)
  8.1642 +                  field_tab"
  8.1643 +text  {* @{term "accfield G C S"}: fields of a class @{term C} which are 
  8.1644 +         accessible from scope of class
  8.1645 +         @{term S} with inheritance and hiding, cf. 8.3 *}
  8.1646 +text {* note the class component in the accessibility filter (see also 
  8.1647 +        @{term methd}).
  8.1648 +   The class declaring field @{term f} (@{term declC}) isn't necessarily 
  8.1649 +   accessible from scope @{term S}. The field can be made visible through 
  8.1650 +   inheritance, too. So we must test accessibility of field @{term f} of class 
  8.1651 +   @{term C} (not @{term "declclass f"}) *} 
  8.1652 +
  8.1653 +constdefs
  8.1654 +
  8.1655 +  is_methd :: "prog \<Rightarrow> qtname  \<Rightarrow> sig \<Rightarrow> bool"
  8.1656 + "is_methd G \<equiv> \<lambda>C sig. is_class G C \<and> methd G C sig \<noteq> None"
  8.1657 +
  8.1658 +constdefs efname:: "((vname \<times> qtname) \<times> field) \<Rightarrow> (vname \<times> qtname)"
  8.1659 +"efname \<equiv> fst"
  8.1660 +
  8.1661 +lemma efname_simp[simp]:"efname (n,f) = n"
  8.1662 +by (simp add: efname_def) 
  8.1663 +
  8.1664 +
  8.1665 +subsection "imethds"
  8.1666 +
  8.1667 +lemma imethds_rec: "\<lbrakk>iface G I = Some i; ws_prog G\<rbrakk> \<Longrightarrow>  
  8.1668 +  imethds G I = Un_tables ((\<lambda>J. imethds  G J)`set (isuperIfs i)) \<oplus>\<oplus>  
  8.1669 +                      (o2s \<circ> table_of (map (\<lambda>(s,mh). (s,I,mh)) (imethods i)))"
  8.1670 +apply (unfold imethds_def)
  8.1671 +apply (rule iface_rec [THEN trans])
  8.1672 +apply auto
  8.1673 +done
  8.1674 +
  8.1675 +
  8.1676 +(* local lemma *)
  8.1677 +lemma imethds_norec:
  8.1678 +  "\<lbrakk>iface G md = Some i; ws_prog G; table_of (imethods i) sig = Some mh\<rbrakk> \<Longrightarrow>  
  8.1679 +  (md, mh) \<in> imethds G md sig"
  8.1680 +apply (subst imethds_rec)
  8.1681 +apply assumption+
  8.1682 +apply (rule iffD2)
  8.1683 +apply (rule overrides_t_Some_iff)
  8.1684 +apply (rule disjI1)
  8.1685 +apply (auto elim: table_of_map_SomeI)
  8.1686 +done
  8.1687 +
  8.1688 +lemma imethds_declI: "\<lbrakk>m \<in> imethds G I sig; ws_prog G; is_iface G I\<rbrakk> \<Longrightarrow> 
  8.1689 +  (\<exists>i. iface G (decliface m) = Some i \<and> 
  8.1690 +  table_of (imethods i) sig = Some (mthd m)) \<and>  
  8.1691 +  (I,decliface m) \<in> (subint1 G)^* \<and> m \<in> imethds G (decliface m) sig"
  8.1692 +apply (erule make_imp)
  8.1693 +apply (rule ws_subint1_induct, assumption, assumption)
  8.1694 +apply (subst imethds_rec, erule conjunct1, assumption)
  8.1695 +apply (force elim: imethds_norec intro: rtrancl_into_rtrancl2)
  8.1696 +done
  8.1697 +
  8.1698 +lemma imethds_cases [consumes 3, case_names NewMethod InheritedMethod]:
  8.1699 + (assumes im: "im \<in> imethds G I sig" and  
  8.1700 +         ifI: "iface G I = Some i" and
  8.1701 +          ws: "ws_prog G" and
  8.1702 +     hyp_new:  "table_of (map (\<lambda>(s, mh). (s, I, mh)) (imethods i)) sig 
  8.1703 +                = Some im \<Longrightarrow> P" and
  8.1704 +     hyp_inh: "\<And> J. \<lbrakk>J \<in> set (isuperIfs i); im \<in> imethds G J sig\<rbrakk> \<Longrightarrow> P"
  8.1705 +  ) "P"
  8.1706 +proof -
  8.1707 +  from ifI ws im hyp_new hyp_inh
  8.1708 +  show "P"
  8.1709 +    by (auto simp add: imethds_rec)
  8.1710 +qed
  8.1711 +
  8.1712 +subsection "accimethd"
  8.1713 +
  8.1714 +lemma accimethds_simp [simp]: 
  8.1715 +"G\<turnstile>Iface I accessible_in pack \<Longrightarrow> accimethds G pack I = imethds G I"
  8.1716 +by (simp add: accimethds_def)
  8.1717 +
  8.1718 +lemma accimethdsD:
  8.1719 + "im \<in> accimethds G pack I sig 
  8.1720 +  \<Longrightarrow> im \<in> imethds G I sig \<and> G\<turnstile>Iface I accessible_in pack"
  8.1721 +by (auto simp add: accimethds_def)
  8.1722 +
  8.1723 +lemma accimethdsI: 
  8.1724 +"\<lbrakk>im \<in> imethds G I sig;G\<turnstile>Iface I accessible_in pack\<rbrakk>
  8.1725 + \<Longrightarrow> im \<in> accimethds G pack I sig"
  8.1726 +by simp
  8.1727 +
  8.1728 +subsection "methd"
  8.1729 +
  8.1730 +lemma methd_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
  8.1731 +  methd G C 
  8.1732 +    = (if C = Object 
  8.1733 +          then empty 
  8.1734 +          else filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)
  8.1735 +                          (methd G (super c))) 
  8.1736 +      ++ table_of (map (\<lambda>(s,m). (s,C,m)) (methods c))"
  8.1737 +apply (unfold methd_def)
  8.1738 +apply (erule class_rec [THEN trans], assumption)
  8.1739 +apply (simp)
  8.1740 +done
  8.1741 +
  8.1742 +(* local lemma *)
  8.1743 +lemma methd_norec: 
  8.1744 + "\<lbrakk>class G declC = Some c; ws_prog G;table_of (methods c) sig = Some m\<rbrakk> 
  8.1745 +  \<Longrightarrow> methd G declC sig = Some (declC, m)"
  8.1746 +apply (simp only: methd_rec)
  8.1747 +apply (rule disjI1 [THEN override_Some_iff [THEN iffD2]])
  8.1748 +apply (auto elim: table_of_map_SomeI)
  8.1749 +done
  8.1750 +
  8.1751 +
  8.1752 +lemma methd_declC: 
  8.1753 +"\<lbrakk>methd G C sig = Some m; ws_prog G;is_class G C\<rbrakk> \<Longrightarrow>
  8.1754 + (\<exists>d. class G (declclass m)=Some d \<and> table_of (methods d) sig=Some (mthd m)) \<and> 
  8.1755 + G\<turnstile>C \<preceq>\<^sub>C (declclass m) \<and> methd G (declclass m) sig = Some m"   
  8.1756 +apply (erule make_imp)
  8.1757 +apply (rule ws_subcls1_induct, assumption, assumption)
  8.1758 +apply (subst methd_rec, assumption)
  8.1759 +apply (case_tac "Ca=Object")
  8.1760 +apply   (force elim: methd_norec )
  8.1761 +
  8.1762 +apply   simp
  8.1763 +apply   (case_tac "table_of (map (\<lambda>(s, m). (s, Ca, m)) (methods c)) sig")
  8.1764 +apply     (force intro: rtrancl_into_rtrancl2)
  8.1765 +
  8.1766 +apply     (auto intro: methd_norec)
  8.1767 +done
  8.1768 +
  8.1769 +lemma methd_inheritedD:
  8.1770 +  "\<lbrakk>class G C = Some c; ws_prog G;methd G C sig = Some m\<rbrakk>
  8.1771 +  \<Longrightarrow>  (declclass m \<noteq> C \<longrightarrow> G \<turnstile>C inherits method sig m)"
  8.1772 +by (auto simp add: methd_rec)
  8.1773 +
  8.1774 +lemma methd_diff_cls:
  8.1775 +"\<lbrakk>ws_prog G; is_class G C; is_class G D;
  8.1776 + methd G C sig = m; methd G D sig = n; m\<noteq>n
  8.1777 +\<rbrakk> \<Longrightarrow> C\<noteq>D"
  8.1778 +by (auto simp add: methd_rec)
  8.1779 +
  8.1780 +lemma method_declared_inI: 
  8.1781 + "\<lbrakk>table_of (methods c) sig = Some m; class G C = Some c\<rbrakk>
  8.1782 +  \<Longrightarrow> G\<turnstile>mdecl (sig,m) declared_in C"
  8.1783 +by (auto simp add: cdeclaredmethd_def declared_in_def)
  8.1784 +
  8.1785 +lemma methd_declared_in_declclass: 
  8.1786 + "\<lbrakk>methd G C sig = Some m; ws_prog G;is_class G C\<rbrakk> 
  8.1787 + \<Longrightarrow> G\<turnstile>Methd sig m declared_in (declclass m)"
  8.1788 +by (auto dest: methd_declC method_declared_inI)
  8.1789 +
  8.1790 +lemma member_methd:
  8.1791 + (assumes member_of: "G\<turnstile>Methd sig m member_of C" and
  8.1792 +                 ws: "ws_prog G"
  8.1793 + ) "methd G C sig = Some m"
  8.1794 +proof -
  8.1795 +  from member_of 
  8.1796 +  have iscls_C: "is_class G C" 
  8.1797 +    by (rule member_of_is_classD)
  8.1798 +  from iscls_C ws member_of
  8.1799 +  show ?thesis (is "?Methd C")
  8.1800 +  proof (induct rule: ws_class_induct')
  8.1801 +    case (Object co)
  8.1802 +    assume "G \<turnstile>Methd sig m member_of Object"
  8.1803 +    then have "G\<turnstile>Methd sig m declared_in Object \<and> declclass m = Object"
  8.1804 +      by (cases set: members) (cases m, auto dest: subcls1D)
  8.1805 +    with ws Object 
  8.1806 +    show "?Methd Object"
  8.1807 +      by (cases m)
  8.1808 +         (auto simp add: declared_in_def cdeclaredmethd_def methd_rec
  8.1809 +                  intro:  table_of_mapconst_SomeI)
  8.1810 +  next
  8.1811 +    case (Subcls C c)
  8.1812 +    assume clsC: "class G C = Some c" and
  8.1813 +      neq_C_Obj: "C \<noteq> Object" and
  8.1814 +            hyp: "G \<turnstile>Methd sig m member_of super c \<Longrightarrow> ?Methd (super c)" and
  8.1815 +      member_of: "G \<turnstile>Methd sig m member_of C"
  8.1816 +    from member_of
  8.1817 +    show "?Methd C"
  8.1818 +    proof (cases)
  8.1819 +      case (Immediate Ca membr)
  8.1820 +      then have "Ca=C" "membr = method sig m" and 
  8.1821 +                "G\<turnstile>Methd sig m declared_in C" "declclass m = C"
  8.1822 +	by (cases m,auto)
  8.1823 +      with clsC 
  8.1824 +      have "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = Some m"
  8.1825 +	by (cases m)
  8.1826 +	   (auto simp add: declared_in_def cdeclaredmethd_def
  8.1827 +	            intro: table_of_mapconst_SomeI)
  8.1828 +      with clsC neq_C_Obj ws 
  8.1829 +      show ?thesis
  8.1830 +	by (simp add: methd_rec)
  8.1831 +    next
  8.1832 +      case (Inherited Ca S membr)
  8.1833 +      with clsC
  8.1834 +      have eq_Ca_C: "Ca=C" and
  8.1835 +            undecl: "G\<turnstile>mid sig undeclared_in C" and
  8.1836 +             super: "G \<turnstile>Methd sig m member_of (super c)"
  8.1837 +	by (auto dest: subcls1D)
  8.1838 +      from eq_Ca_C clsC undecl 
  8.1839 +      have "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = None"
  8.1840 +	by (auto simp add: undeclared_in_def cdeclaredmethd_def
  8.1841 +	            intro: table_of_mapconst_NoneI)
  8.1842 +      moreover
  8.1843 +      from Inherited have "G \<turnstile> C inherits (method sig m)"
  8.1844 +	by (auto simp add: inherits_def)
  8.1845 +      moreover
  8.1846 +      note clsC neq_C_Obj ws super hyp 
  8.1847 +      ultimately
  8.1848 +      show ?thesis
  8.1849 +	by (auto simp add: methd_rec intro: filter_tab_SomeI)
  8.1850 +    qed
  8.1851 +  qed
  8.1852 +qed
  8.1853 +
  8.1854 +(*unused*)
  8.1855 +lemma finite_methd:"ws_prog G \<Longrightarrow> finite {methd G C sig |sig C. is_class G C}"
  8.1856 +apply (rule finite_is_class [THEN finite_SetCompr2])
  8.1857 +apply (intro strip)
  8.1858 +apply (erule_tac ws_subcls1_induct, assumption)
  8.1859 +apply (subst methd_rec)
  8.1860 +apply (assumption)
  8.1861 +apply (auto intro!: finite_range_map_of finite_range_filter_tab finite_range_map_of_override)
  8.1862 +done
  8.1863 +
  8.1864 +lemma finite_dom_methd:
  8.1865 + "\<lbrakk>ws_prog G; is_class G C\<rbrakk> \<Longrightarrow> finite (dom (methd G C))"
  8.1866 +apply (erule_tac ws_subcls1_induct)
  8.1867 +apply assumption
  8.1868 +apply (subst methd_rec)
  8.1869 +apply (assumption)
  8.1870 +apply (auto intro!: finite_dom_map_of finite_dom_filter_tab)
  8.1871 +done
  8.1872 +
  8.1873 +
  8.1874 +subsection "accmethd"
  8.1875 +
  8.1876 +lemma accmethd_SomeD:
  8.1877 +"accmethd G S C sig = Some m
  8.1878 + \<Longrightarrow> methd G C sig = Some m \<and> G\<turnstile>method sig m of C accessible_from S"
  8.1879 +by (auto simp add: accmethd_def dest: filter_tab_SomeD)
  8.1880 +
  8.1881 +lemma accmethd_SomeI:
  8.1882 +"\<lbrakk>methd G C sig = Some m; G\<turnstile>method sig m of C accessible_from S\<rbrakk> 
  8.1883 + \<Longrightarrow> accmethd G S C sig = Some m"
  8.1884 +by (auto simp add: accmethd_def intro: filter_tab_SomeI)
  8.1885 +
  8.1886 +lemma accmethd_declC: 
  8.1887 +"\<lbrakk>accmethd G S C sig = Some m; ws_prog G; is_class G C\<rbrakk> \<Longrightarrow>
  8.1888 + (\<exists>d. class G (declclass m)=Some d \<and> 
  8.1889 +  table_of (methods d) sig=Some (mthd m)) \<and> 
  8.1890 + G\<turnstile>C \<preceq>\<^sub>C (declclass m) \<and> methd G (declclass m) sig = Some m \<and> 
  8.1891 + G\<turnstile>method sig m of C accessible_from S"
  8.1892 +by (auto dest: accmethd_SomeD methd_declC accmethd_SomeI)
  8.1893 +
  8.1894 +
  8.1895 +lemma finite_dom_accmethd:
  8.1896 + "\<lbrakk>ws_prog G; is_class G C\<rbrakk> \<Longrightarrow> finite (dom (accmethd G S C))"
  8.1897 +by (auto simp add: accmethd_def intro: finite_dom_filter_tab finite_dom_methd)
  8.1898 +
  8.1899 +
  8.1900 +subsection "dynmethd"
  8.1901 +
  8.1902 +lemma dynmethd_rec:
  8.1903 +"\<lbrakk>class G dynC = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
  8.1904 + dynmethd G statC dynC sig
  8.1905 +   = (if G\<turnstile>dynC \<preceq>\<^sub>C statC
  8.1906 +        then (case methd G statC sig of
  8.1907 +                None \<Rightarrow> None
  8.1908 +              | Some statM 
  8.1909 +                  \<Rightarrow> (case methd G dynC sig of
  8.1910 +                        None \<Rightarrow> dynmethd G statC (super c) sig
  8.1911 +                      | Some dynM \<Rightarrow> 
  8.1912 +                          (if G,sig\<turnstile> dynM overrides statM \<or> dynM = statM 
  8.1913 +                              then Some dynM
  8.1914 +                              else (dynmethd G statC (super c) sig)
  8.1915 +                      )))
  8.1916 +         else None)" 
  8.1917 +   (is "_ \<Longrightarrow> _ \<Longrightarrow> ?Dynmethd_def dynC sig  = ?Dynmethd_rec dynC c sig") 
  8.1918 +proof -
  8.1919 +  assume clsDynC: "class G dynC = Some c" and 
  8.1920 +              ws: "ws_prog G"
  8.1921 +  then show "?Dynmethd_def dynC sig  = ?Dynmethd_rec dynC c sig" 
  8.1922 +  proof (induct rule: ws_class_induct'')
  8.1923 +    case (Object co)
  8.1924 +    show "?Dynmethd_def Object sig = ?Dynmethd_rec Object co sig"
  8.1925 +    proof (cases "G\<turnstile>Object \<preceq>\<^sub>C statC") 
  8.1926 +      case False
  8.1927 +      then show ?thesis by (simp add: dynmethd_def)
  8.1928 +    next
  8.1929 +      case True
  8.1930 +      then have eq_statC_Obj: "statC = Object" ..
  8.1931 +      show ?thesis 
  8.1932 +      proof (cases "methd G statC sig")
  8.1933 +	case None then show ?thesis by (simp add: dynmethd_def)
  8.1934 +      next
  8.1935 +	case Some
  8.1936 +	with True Object ws eq_statC_Obj 
  8.1937 +	show ?thesis
  8.1938 +	  by (auto simp add: dynmethd_def class_rec
  8.1939 +                      intro: filter_tab_SomeI)
  8.1940 +      qed
  8.1941 +    qed
  8.1942 +  next  
  8.1943 +    case (Subcls dynC c sc)
  8.1944 +    show "?Dynmethd_def dynC sig = ?Dynmethd_rec dynC c sig"
  8.1945 +    proof (cases "G\<turnstile>dynC \<preceq>\<^sub>C statC") 
  8.1946 +      case False
  8.1947 +      then show ?thesis by (simp add: dynmethd_def)
  8.1948 +    next
  8.1949 +      case True
  8.1950 +      note subclseq_dynC_statC = True
  8.1951 +      show ?thesis
  8.1952 +      proof (cases "methd G statC sig")
  8.1953 +	case None then show ?thesis by (simp add: dynmethd_def)
  8.1954 +      next
  8.1955 +	case (Some statM)
  8.1956 +	note statM = Some
  8.1957 +	let "?filter C" = 
  8.1958 +              "filter_tab
  8.1959 +                (\<lambda>_ dynM. G,sig \<turnstile> dynM overrides statM \<or> dynM = statM)
  8.1960 +                (methd G C)"
  8.1961 +        let "?class_rec C" =
  8.1962 +              "(class_rec (G, C) empty
  8.1963 +                           (\<lambda>C c subcls_mthds. subcls_mthds ++ (?filter C)))"
  8.1964 +	from statM Subcls ws subclseq_dynC_statC
  8.1965 +	have dynmethd_dynC_def:
  8.1966 +             "?Dynmethd_def dynC sig =
  8.1967 +                ((?class_rec (super c)) 
  8.1968 +                 ++
  8.1969 +                (?filter dynC)) sig"
  8.1970 +         by (simp (no_asm_simp) only: dynmethd_def class_rec)
  8.1971 +	    auto
  8.1972 +       show ?thesis
  8.1973 +       proof (cases "dynC = statC")
  8.1974 +	 case True
  8.1975 +	 with subclseq_dynC_statC statM dynmethd_dynC_def
  8.1976 +	 have "?Dynmethd_def dynC sig = Some statM"
  8.1977 +	   by (auto intro: override_find_right filter_tab_SomeI)
  8.1978 +	 with subclseq_dynC_statC True Some 
  8.1979 +	 show ?thesis
  8.1980 +	   by auto
  8.1981 +       next
  8.1982 +	 case False
  8.1983 +	 with subclseq_dynC_statC Subcls 
  8.1984 +	 have subclseq_super_statC: "G\<turnstile>(super c) \<preceq>\<^sub>C statC"
  8.1985 +	   by (blast dest: subclseq_superD)
  8.1986 +	 show ?thesis
  8.1987 +	 proof (cases "methd G dynC sig") 
  8.1988 +	   case None
  8.1989 +	   then have "?filter dynC sig = None"
  8.1990 +	     by (rule filter_tab_None)
  8.1991 +           then have "?Dynmethd_def dynC sig=?class_rec (super c) sig"
  8.1992 +	     by (simp add: dynmethd_dynC_def)
  8.1993 +	   with  subclseq_super_statC statM None
  8.1994 +	   have "?Dynmethd_def dynC sig = ?Dynmethd_def (super c) sig"
  8.1995 +	     by (auto simp add: empty_def dynmethd_def)
  8.1996 +           with None subclseq_dynC_statC statM
  8.1997 +	   show ?thesis 
  8.1998 +	     by simp
  8.1999 +	 next
  8.2000 +	   case (Some dynM)
  8.2001 +	   note dynM = Some
  8.2002 +	   let ?Termination = "G \<turnstile> qmdecl sig dynM overrides qmdecl sig statM \<or>
  8.2003 +                               dynM = statM"
  8.2004 +	   show ?thesis
  8.2005 +	   proof (cases "?filter dynC sig")
  8.2006 +	     case None
  8.2007 +	     with dynM 
  8.2008 +	     have no_termination: "\<not> ?Termination"
  8.2009 +	       by (simp add: filter_tab_def)
  8.2010 +	     from None 
  8.2011 +	     have "?Dynmethd_def dynC sig=?class_rec (super c) sig"
  8.2012 +	       by (simp add: dynmethd_dynC_def)
  8.2013 +	     with subclseq_super_statC statM dynM no_termination
  8.2014 +	     show ?thesis
  8.2015 +	       by (auto simp add: empty_def dynmethd_def)
  8.2016 +	   next
  8.2017 +	     case Some
  8.2018 +	     with dynM
  8.2019 +	     have termination: "?Termination"
  8.2020 +	       by (auto)
  8.2021 +	     with Some dynM
  8.2022 +	     have "?Dynmethd_def dynC sig=Some dynM"
  8.2023 +	      by (auto simp add: dynmethd_dynC_def)
  8.2024 +	     with subclseq_super_statC statM dynM termination
  8.2025 +	     show ?thesis
  8.2026 +	       by (auto simp add: dynmethd_def)
  8.2027 +	   qed
  8.2028 +	 qed
  8.2029 +       qed
  8.2030 +     qed
  8.2031 +   qed
  8.2032 + qed
  8.2033 +qed
  8.2034 +	       
  8.2035 +lemma dynmethd_C_C:"\<lbrakk>is_class G C; ws_prog G\<rbrakk> 
  8.2036 +\<Longrightarrow> dynmethd G C C sig = methd G C sig"          
  8.2037 +apply (auto simp add: dynmethd_rec)
  8.2038 +done
  8.2039 + 
  8.2040 +lemma dynmethdSomeD: 
  8.2041 + "\<lbrakk>dynmethd G statC dynC sig = Some dynM; is_class G dynC; ws_prog G\<rbrakk> 
  8.2042 +  \<Longrightarrow> G\<turnstile>dynC \<preceq>\<^sub>C statC \<and> (\<exists> statM. methd G statC sig = Some statM)"
  8.2043 +apply clarify
  8.2044 +apply rotate_tac
  8.2045 +by (auto simp add: dynmethd_rec)
  8.2046 + 
  8.2047 +lemma dynmethd_Some_cases [consumes 3, case_names Static Overrides]:
  8.2048 +  (assumes dynM: "dynmethd G statC dynC sig = Some dynM" and
  8.2049 +        is_cls_dynC: "is_class G dynC" and
  8.2050 +                 ws: "ws_prog G" and
  8.2051 +         hyp_static: "methd G statC sig = Some dynM \<Longrightarrow> P" and
  8.2052 +       hyp_override: "\<And> statM. \<lbrakk>methd G statC sig = Some statM;dynM\<noteq>statM;
  8.2053 +                       G,sig\<turnstile>dynM overrides statM\<rbrakk> \<Longrightarrow> P"
  8.2054 +   ) "P"
  8.2055 +proof -
  8.2056 +  from is_cls_dynC obtain dc where clsDynC: "class G dynC = Some dc" by blast
  8.2057 +  from clsDynC ws dynM hyp_static hyp_override
  8.2058 +  show "P"
  8.2059 +  proof (induct rule: ws_class_induct)
  8.2060 +    case (Object co)
  8.2061 +    with ws  have "statC = Object" 
  8.2062 +      by (auto simp add: dynmethd_rec)
  8.2063 +    with ws Object show ?thesis by (auto simp add: dynmethd_C_C)
  8.2064 +  next
  8.2065 +    case (Subcls C c)
  8.2066 +    with ws show ?thesis
  8.2067 +      by (auto simp add: dynmethd_rec)
  8.2068 +  qed
  8.2069 +qed
  8.2070 +
  8.2071 +lemma no_override_in_Object:
  8.2072 +  (assumes     dynM: "dynmethd G statC dynC sig = Some dynM" and
  8.2073 +            is_cls_dynC: "is_class G dynC" and
  8.2074 +                     ws: "ws_prog G" and
  8.2075 +                  statM: "methd G statC sig = Some statM" and
  8.2076 +         neq_dynM_statM: "dynM\<noteq>statM"
  8.2077 +   )
  8.2078 +   "dynC \<noteq> Object"
  8.2079 +proof -
  8.2080 +  from is_cls_dynC obtain dc where clsDynC: "class G dynC = Some dc" by blast
  8.2081 +  from clsDynC ws dynM statM neq_dynM_statM 
  8.2082 +  show ?thesis (is "?P dynC")
  8.2083 +  proof (induct rule: ws_class_induct)
  8.2084 +    case (Object co)
  8.2085 +    with ws  have "statC = Object" 
  8.2086 +      by (auto simp add: dynmethd_rec)
  8.2087 +    with ws Object show "?P Object" by (auto simp add: dynmethd_C_C)
  8.2088 +  next
  8.2089 +    case (Subcls dynC c)
  8.2090 +    with ws show "?P dynC"
  8.2091 +      by (auto simp add: dynmethd_rec)
  8.2092 +  qed
  8.2093 +qed
  8.2094 +
  8.2095 +
  8.2096 +lemma dynmethd_Some_rec_cases [consumes 3, 
  8.2097 +                               case_names Static Override  Recursion]:
  8.2098 +(assumes     dynM: "dynmethd G statC dynC sig = Some dynM" and
  8.2099 +                clsDynC: "class G dynC = Some c" and
  8.2100 +                     ws: "ws_prog G" and
  8.2101 +             hyp_static: "methd G statC sig = Some dynM \<Longrightarrow> P" and
  8.2102 +           hyp_override: "\<And> statM. \<lbrakk>methd G statC sig = Some statM;
  8.2103 +                                     methd G dynC sig = Some dynM; statM\<noteq>dynM;
  8.2104 +                                     G,sig\<turnstile> dynM overrides statM\<rbrakk> \<Longrightarrow> P" and
  8.2105 +
  8.2106 +          hyp_recursion: "\<lbrakk>dynC\<noteq>Object;
  8.2107 +                           dynmethd G statC (super c) sig = Some dynM\<rbrakk> \<Longrightarrow> P" 
  8.2108 +  ) "P"
  8.2109 +proof -
  8.2110 +  from clsDynC have "is_class G dynC" by simp
  8.2111 +  note no_override_in_Object' = no_override_in_Object [OF dynM this ws]
  8.2112 +  from ws clsDynC dynM hyp_static hyp_override hyp_recursion
  8.2113 +  show ?thesis
  8.2114 +    by (auto simp add: dynmethd_rec dest: no_override_in_Object')
  8.2115 +qed
  8.2116 +
  8.2117 +lemma dynmethd_declC: 
  8.2118 +"\<lbrakk>dynmethd G statC dynC sig = Some m;
  8.2119 +  is_class G statC;ws_prog G
  8.2120 + \<rbrakk> \<Longrightarrow>
  8.2121 +  (\<exists>d. class G (declclass m)=Some d \<and> table_of (methods d) sig=Some (mthd m)) \<and>
  8.2122 +  G\<turnstile>dynC \<preceq>\<^sub>C (declclass m) \<and> methd G (declclass m) sig = Some m"
  8.2123 +proof - 
  8.2124 +  assume  is_cls_statC: "is_class G statC" 
  8.2125 +  assume            ws: "ws_prog G"  
  8.2126 +  assume             m: "dynmethd G statC dynC sig = Some m"
  8.2127 +  from m 
  8.2128 +  have "G\<turnstile>dynC \<preceq>\<^sub>C statC" by (auto simp add: dynmethd_def)
  8.2129 +  from this is_cls_statC 
  8.2130 +  have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2)
  8.2131 +  from is_cls_dynC ws m  
  8.2132 +  show ?thesis (is "?P dynC") 
  8.2133 +  proof (induct rule: ws_class_induct')
  8.2134 +    case (Object co)
  8.2135 +    with ws have "statC=Object" by (auto simp add: dynmethd_rec)
  8.2136 +    with ws Object  
  8.2137 +    show "?P Object" 
  8.2138 +      by (auto simp add: dynmethd_C_C dest: methd_declC)
  8.2139 +  next
  8.2140 +    case (Subcls dynC c)
  8.2141 +    assume   hyp: "dynmethd G statC (super c) sig = Some m \<Longrightarrow> ?P (super c)" and
  8.2142 +         clsDynC: "class G dynC = Some c"  and
  8.2143 +              m': "dynmethd G statC dynC sig = Some m" and
  8.2144 +    neq_dynC_Obj: "dynC \<noteq> Object"
  8.2145 +    from ws this obtain statM where
  8.2146 +      subclseq_dynC_statC: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and 
  8.2147 +                     statM: "methd G statC sig = Some statM"
  8.2148 +      by (blast dest: dynmethdSomeD)
  8.2149 +    from clsDynC neq_dynC_Obj 
  8.2150 +    have subclseq_dynC_super: "G\<turnstile>dynC \<preceq>\<^sub>C (super c)" 
  8.2151 +      by (auto intro: subcls1I) 
  8.2152 +    from m' clsDynC ws 
  8.2153 +    show "?P dynC"
  8.2154 +    proof (cases rule: dynmethd_Some_rec_cases) 
  8.2155 +      case Static
  8.2156 +      with is_cls_statC ws subclseq_dynC_statC 
  8.2157 +      show ?thesis 
  8.2158 +	by (auto intro: rtrancl_trans dest: methd_declC)
  8.2159 +    next
  8.2160 +      case Override
  8.2161 +      with clsDynC ws 
  8.2162 +      show ?thesis 
  8.2163 +	by (auto dest: methd_declC)
  8.2164 +    next
  8.2165 +      case Recursion
  8.2166 +      with hyp subclseq_dynC_super 
  8.2167 +      show ?thesis 
  8.2168 +	by (auto intro: rtrancl_trans) 
  8.2169 +    qed
  8.2170 +  qed
  8.2171 +qed
  8.2172 +
  8.2173 +lemma methd_Some_dynmethd_Some:
  8.2174 +  (assumes    statM: "methd G statC sig = Some statM" and 
  8.2175 +           subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
  8.2176 +       is_cls_statC: "is_class G statC" and
  8.2177 +                 ws: "ws_prog G"
  8.2178 +   ) "\<exists> dynM. dynmethd G statC dynC sig = Some dynM"
  8.2179 +   (is "?P dynC")
  8.2180 +proof -
  8.2181 +  from subclseq is_cls_statC 
  8.2182 +  have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2)
  8.2183 +  then obtain dc where
  8.2184 +    clsDynC: "class G dynC = Some dc" by blast
  8.2185 +  from clsDynC ws subclseq 
  8.2186 +  show "?thesis"
  8.2187 +  proof (induct rule: ws_class_induct)
  8.2188 +    case (Object co)
  8.2189 +    with ws  have "statC = Object" 
  8.2190 +      by (auto)
  8.2191 +    with ws Object statM
  8.2192 +    show "?P Object"  
  8.2193 +      by (auto simp add: dynmethd_C_C)
  8.2194 +  next
  8.2195 +    case (Subcls dynC dc)
  8.2196 +    assume clsDynC': "class G dynC = Some dc"
  8.2197 +    assume neq_dynC_Obj: "dynC \<noteq> Object"
  8.2198 +    assume hyp: "G\<turnstile>super dc\<preceq>\<^sub>C statC \<Longrightarrow> ?P (super dc)"
  8.2199 +    assume subclseq': "G\<turnstile>dynC\<preceq>\<^sub>C statC"
  8.2200 +    then
  8.2201 +    show "?P dynC"
  8.2202 +    proof (cases rule: subclseq_cases)
  8.2203 +      case Eq
  8.2204 +      with ws statM clsDynC' 
  8.2205 +      show ?thesis
  8.2206 +	by (auto simp add: dynmethd_rec)
  8.2207 +    next
  8.2208 +      case Subcls
  8.2209 +      assume "G\<turnstile>dynC\<prec>\<^sub>C statC"
  8.2210 +      from this clsDynC' 
  8.2211 +      have "G\<turnstile>super dc\<preceq>\<^sub>C statC" by (rule subcls_superD)
  8.2212 +      with hyp ws clsDynC' subclseq' statM
  8.2213 +      show ?thesis
  8.2214 +	by (auto simp add: dynmethd_rec)
  8.2215 +    qed
  8.2216 +  qed
  8.2217 +qed
  8.2218 +
  8.2219 +lemma dynmethd_cases [consumes 4, case_names Static Overrides]:
  8.2220 +  (assumes    statM: "methd G statC sig = Some statM" and 
  8.2221 +           subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
  8.2222 +       is_cls_statC: "is_class G statC" and
  8.2223 +                 ws: "ws_prog G" and
  8.2224 +         hyp_static: "dynmethd G statC dynC sig = Some statM \<Longrightarrow> P" and
  8.2225 +       hyp_override: "\<And> dynM. \<lbrakk>dynmethd G statC dynC sig = Some dynM;
  8.2226 +                                 dynM\<noteq>statM;
  8.2227 +                           G,sig\<turnstile>dynM overrides statM\<rbrakk> \<Longrightarrow> P"
  8.2228 +   ) "P"
  8.2229 +proof -
  8.2230 +  from subclseq is_cls_statC 
  8.2231 +  have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2)
  8.2232 +  then obtain dc where
  8.2233 +    clsDynC: "class G dynC = Some dc" by blast
  8.2234 +  from statM subclseq is_cls_statC ws 
  8.2235 +  obtain dynM
  8.2236 +    where dynM: "dynmethd G statC dynC sig = Some dynM"
  8.2237 +    by (blast dest: methd_Some_dynmethd_Some)
  8.2238 +  from dynM is_cls_dynC ws 
  8.2239 +  show ?thesis
  8.2240 +  proof (cases rule: dynmethd_Some_cases)
  8.2241 +    case Static
  8.2242 +    with hyp_static dynM statM show ?thesis by simp
  8.2243 +  next
  8.2244 +    case Overrides
  8.2245 +    with hyp_override dynM statM show ?thesis by simp
  8.2246 +  qed
  8.2247 +qed
  8.2248 +
  8.2249 +lemma ws_dynmethd:
  8.2250 +  (assumes statM: "methd G statC sig = Some statM" and
  8.2251 +        subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
  8.2252 +    is_cls_statC: "is_class G statC" and
  8.2253 +              ws: "ws_prog G"
  8.2254 +   )
  8.2255 +   "\<exists> dynM. dynmethd G statC dynC sig = Some dynM \<and>
  8.2256 +            is_static dynM = is_static statM \<and> G\<turnstile>resTy dynM\<preceq>resTy statM"
  8.2257 +proof - 
  8.2258 +  from statM subclseq is_cls_statC ws
  8.2259 +  show ?thesis
  8.2260 +  proof (cases rule: dynmethd_cases)
  8.2261 +    case Static
  8.2262 +    with statM 
  8.2263 +    show ?thesis
  8.2264 +      by simp
  8.2265 +  next
  8.2266 +    case Overrides
  8.2267 +    with ws
  8.2268 +    show ?thesis
  8.2269 +      by (auto dest: ws_overrides_commonD)
  8.2270 +  qed
  8.2271 +qed
  8.2272 +
  8.2273 +(*
  8.2274 +lemma dom_dynmethd: 
  8.2275 +  "dom (dynmethd G statC dynC) \<subseteq> dom (methd G statC) \<union> dom (methd G dynC)"
  8.2276 +by (auto simp add: dynmethd_def dom_def)
  8.2277 +
  8.2278 +lemma finite_dom_dynmethd:
  8.2279 + "\<lbrakk>ws_prog G; is_class G statC; is_class G dynC\<rbrakk> 
  8.2280 +  \<Longrightarrow> finite (dom (dynmethd G statC dynC))"
  8.2281 +apply (rule_tac B="dom (methd G statC) \<union> dom (methd G dynC)" in finite_subset)
  8.2282 +apply (rule dom_dynmethd)
  8.2283 +apply (rule finite_UnI)
  8.2284 +apply (drule (2) finite_dom_methd)+
  8.2285 +done
  8.2286 +*)
  8.2287 +(*
  8.2288 +lemma dynmethd_SomeD: 
  8.2289 +"\<lbrakk>ws_prog G; is_class G statC; is_class G dynC;
  8.2290 + methd G statC sig = Some sm; dynmethd G statC dynC sig = Some dm; sm \<noteq> dm
  8.2291 + \<rbrakk> \<Longrightarrow> G\<turnstile>dynC \<prec>\<^sub>C statC \<and> 
  8.2292 +       (declclass dm \<noteq> dynC \<longrightarrow> G \<turnstile> dm accessible_through_inheritance_in dynC)"
  8.2293 +by (auto simp add: dynmethd_def 
  8.2294 +         dest: methd_inheritedD methd_diff_cls
  8.2295 +         intro: rtrancl_into_trancl3)
  8.2296 +*)
  8.2297 +
  8.2298 +subsection "dynlookup"
  8.2299 +
  8.2300 +lemma dynlookup_cases [consumes 1, case_names NullT IfaceT ClassT ArrayT]:
  8.2301 +"\<lbrakk>dynlookup G statT dynC sig = x;
  8.2302 +           \<lbrakk>statT = NullT       ; empty sig = x                  \<rbrakk> \<Longrightarrow> P;
  8.2303 +  \<And> I.    \<lbrakk>statT = IfaceT I    ; dynimethd G I      dynC sig = x\<rbrakk> \<Longrightarrow> P;
  8.2304 +  \<And> statC.\<lbrakk>statT = ClassT statC; dynmethd  G statC  dynC sig = x\<rbrakk> \<Longrightarrow> P;
  8.2305 +  \<And> ty.   \<lbrakk>statT = ArrayT ty   ; dynmethd  G Object dynC sig = x\<rbrakk> \<Longrightarrow> P
  8.2306 + \<rbrakk> \<Longrightarrow> P"
  8.2307 +by (cases statT) (auto simp add: dynlookup_def)
  8.2308 +
  8.2309 +subsection "fields"
  8.2310 +
  8.2311 +lemma fields_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
  8.2312 +  fields G C = map (\<lambda>(fn,ft). ((fn,C),ft)) (cfields c) @  
  8.2313 +  (if C = Object then [] else fields G (super c))"
  8.2314 +apply (simp only: fields_def)
  8.2315 +apply (erule class_rec [THEN trans])
  8.2316 +apply assumption
  8.2317 +apply clarsimp
  8.2318 +done
  8.2319 +
  8.2320 +(* local lemma *)
  8.2321 +lemma fields_norec: 
  8.2322 +"\<lbrakk>class G fd = Some c; ws_prog G;  table_of (cfields c) fn = Some f\<rbrakk> 
  8.2323 + \<Longrightarrow> table_of (fields G fd) (fn,fd) = Some f"
  8.2324 +apply (subst fields_rec)
  8.2325 +apply assumption+
  8.2326 +apply (subst map_of_override [symmetric])
  8.2327 +apply (rule disjI1 [THEN override_Some_iff [THEN iffD2]])
  8.2328 +apply (auto elim: table_of_map2_SomeI)
  8.2329 +done
  8.2330 +
  8.2331 +(* local lemma *)
  8.2332 +lemma table_of_fieldsD:
  8.2333 +"table_of (map (\<lambda>(fn,ft). ((fn,C),ft)) (cfields c)) efn = Some f
  8.2334 +\<Longrightarrow> (declclassf efn) = C \<and> table_of (cfields c) (fname efn) = Some f"
  8.2335 +apply (case_tac "efn")
  8.2336 +by auto
  8.2337 +
  8.2338 +lemma fields_declC: 
  8.2339 + "\<lbrakk>table_of (fields G C) efn = Some f; ws_prog G; is_class G C\<rbrakk> \<Longrightarrow>  
  8.2340 +  (\<exists>d. class G (declclassf efn) = Some d \<and> 
  8.2341 +                    table_of (cfields d) (fname efn)=Some f) \<and> 
  8.2342 +  G\<turnstile>C \<preceq>\<^sub>C (declclassf efn)  \<and> table_of (fields G (declclassf efn)) efn = Some f"
  8.2343 +apply (erule make_imp)
  8.2344 +apply (rule ws_subcls1_induct, assumption, assumption)
  8.2345 +apply (subst fields_rec, assumption)
  8.2346 +apply clarify
  8.2347 +apply (simp only: map_of_override [symmetric] del: map_of_override)
  8.2348 +apply (case_tac "table_of (map (split (\<lambda>fn. Pair (fn, Ca))) (cfields c)) efn") 
  8.2349 +apply   (force intro:rtrancl_into_rtrancl2 simp add: override_def)
  8.2350 +
  8.2351 +apply   (frule_tac fd="Ca" in fields_norec)
  8.2352 +apply     assumption
  8.2353 +apply     blast
  8.2354 +apply   (frule table_of_fieldsD)  
  8.2355 +apply   (frule_tac n="table_of (map (split (\<lambda>fn. Pair (fn, Ca))) (cfields c))"
  8.2356 +              and  m="table_of (if Ca = Object then [] else fields G (super c))"
  8.2357 +         in override_find_right)
  8.2358 +apply   (case_tac "efn")
  8.2359 +apply   (simp)
  8.2360 +done
  8.2361 +
  8.2362 +lemma fields_emptyI: "\<And>y. \<lbrakk>ws_prog G; class G C = Some c;cfields c = [];  
  8.2363 +  C \<noteq> Object \<longrightarrow> class G (super c) = Some y \<and> fields G (super c) = []\<rbrakk> \<Longrightarrow>  
  8.2364 +  fields G C = []"
  8.2365 +apply (subst fields_rec)
  8.2366 +apply assumption
  8.2367 +apply auto
  8.2368 +done
  8.2369 +
  8.2370 +(* easier than with table_of *)
  8.2371 +lemma fields_mono_lemma: 
  8.2372 +"\<lbrakk>x \<in> set (fields G C); G\<turnstile>D \<preceq>\<^sub>C C; ws_prog G\<rbrakk> 
  8.2373 + \<Longrightarrow> x \<in> set (fields G D)"
  8.2374 +apply (erule make_imp)
  8.2375 +apply (erule converse_rtrancl_induct)
  8.2376 +apply  fast
  8.2377 +apply (drule subcls1D)
  8.2378 +apply clarsimp
  8.2379 +apply (subst fields_rec)
  8.2380 +apply   auto
  8.2381 +done
  8.2382 +
  8.2383 +
  8.2384 +lemma ws_unique_fields_lemma: 
  8.2385 + "\<lbrakk>(efn,fd)  \<in> set (fields G (super c)); fc \<in> set (cfields c); ws_prog G;  
  8.2386 +   fname efn = fname fc; declclassf efn = C;
  8.2387 +   class G C = Some c; C \<noteq> Object; class G (super c) = Some d\<rbrakk> \<Longrightarrow> R"
  8.2388 +apply (frule_tac ws_prog_cdeclD [THEN conjunct2], assumption, assumption)
  8.2389 +apply (drule_tac weak_map_of_SomeI)
  8.2390 +apply (frule_tac subcls1I [THEN subcls1_irrefl], assumption, assumption)
  8.2391 +apply (auto dest: fields_declC [THEN conjunct2 [THEN conjunct1[THEN rtranclD]]])
  8.2392 +done
  8.2393 +
  8.2394 +lemma ws_unique_fields: "\<lbrakk>is_class G C; ws_prog G; 
  8.2395 +       \<And>C c. \<lbrakk>class G C = Some c\<rbrakk> \<Longrightarrow> unique (cfields c) \<rbrakk> \<Longrightarrow>
  8.2396 +      unique (fields G C)" 
  8.2397 +apply (rule ws_subcls1_induct, assumption, assumption)
  8.2398 +apply (subst fields_rec, assumption)            
  8.2399 +apply (auto intro!: unique_map_inj injI 
  8.2400 +            elim!: unique_append ws_unique_fields_lemma fields_norec
  8.2401 +            )
  8.2402 +done
  8.2403 +
  8.2404 +
  8.2405 +subsection "accfield"
  8.2406 +
  8.2407 +lemma accfield_fields: 
  8.2408 + "accfield G S C fn = Some f 
  8.2409 +  \<Longrightarrow> table_of (fields G C) (fn, declclass f) = Some (fld f)"
  8.2410 +apply (simp only: accfield_def Let_def)
  8.2411 +apply (rule table_of_remap_SomeD)
  8.2412 +apply (auto dest: filter_tab_SomeD)
  8.2413 +done
  8.2414 +
  8.2415 +
  8.2416 +lemma accfield_declC_is_class: 
  8.2417 + "\<lbrakk>is_class G C; accfield G S C en = Some (fd, f); ws_prog G\<rbrakk> \<Longrightarrow>  
  8.2418 +   is_class G fd"
  8.2419 +apply (drule accfield_fields)
  8.2420 +apply (drule fields_declC [THEN conjunct1], assumption)
  8.2421 +apply auto
  8.2422 +done
  8.2423 +
  8.2424 +lemma accfield_accessibleD: 
  8.2425 +  "accfield G S C fn = Some f \<Longrightarrow> G\<turnstile>Field fn f of C accessible_from S"
  8.2426 +by (auto simp add: accfield_def Let_def)
  8.2427 +
  8.2428 +subsection "is methd"
  8.2429 +
  8.2430 +lemma is_methdI: 
  8.2431 +"\<lbrakk>class G C = Some y; methd G C sig = Some b\<rbrakk> \<Longrightarrow> is_methd G C sig"
  8.2432 +apply (unfold is_methd_def)
  8.2433 +apply auto
  8.2434 +done
  8.2435 +
  8.2436 +lemma is_methdD: 
  8.2437 +"is_methd G C sig \<Longrightarrow> class G C \<noteq> None \<and> methd G C sig \<noteq> None"
  8.2438 +apply (unfold is_methd_def)
  8.2439 +apply auto
  8.2440 +done
  8.2441 +
  8.2442 +lemma finite_is_methd: 
  8.2443 + "ws_prog G \<Longrightarrow> finite (Collect (split (is_methd G)))"
  8.2444 +apply (unfold is_methd_def)
  8.2445 +apply (subst SetCompr_Sigma_eq)
  8.2446 +apply (rule finite_is_class [THEN finite_SigmaI])
  8.2447 +apply (simp only: mem_Collect_eq)
  8.2448 +apply (fold dom_def)
  8.2449 +apply (erule finite_dom_methd)
  8.2450 +apply assumption
  8.2451 +done
  8.2452 +
  8.2453 +section "calculation of the superclasses of a class"
  8.2454 +
  8.2455 +constdefs 
  8.2456 + superclasses:: "prog \<Rightarrow> qtname \<Rightarrow> qtname set"
  8.2457 + "superclasses G C \<equiv> class_rec (G,C) {} 
  8.2458 +                       (\<lambda> C c superclss. (if C=Object 
  8.2459 +                                            then {} 
  8.2460 +                                            else insert (super c) superclss))"
  8.2461 +   
  8.2462 +lemma superclasses_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
  8.2463 + superclasses G C 
  8.2464 + = (if (C=Object) 
  8.2465 +       then {}
  8.2466 +       else insert (super c) (superclasses G (super c)))"
  8.2467 +apply (unfold superclasses_def)
  8.2468 +apply (erule class_rec [THEN trans], assumption)
  8.2469 +apply (simp)
  8.2470 +done
  8.2471 +
  8.2472 +lemma superclasses_mono:
  8.2473 +"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D;ws_prog G; class G C = Some c;
  8.2474 +  \<And> C c. \<lbrakk>class G C = Some c;C\<noteq>Object\<rbrakk> \<Longrightarrow> \<exists> sc. class G (super c) = Some sc;
  8.2475 +  x\<in>superclasses G D 
  8.2476 +\<rbrakk> \<Longrightarrow> x\<in>superclasses G C" 
  8.2477 +proof -
  8.2478 +  
  8.2479 +  assume     ws: "ws_prog G"          and 
  8.2480 +          cls_C: "class G C = Some c" and
  8.2481 +             wf: "\<And>C c. \<lbrakk>class G C = Some c; C \<noteq> Object\<rbrakk>
  8.2482 +                  \<Longrightarrow> \<exists>sc. class G (super c) = Some sc"
  8.2483 +  assume clsrel: "G\<turnstile>C\<prec>\<^sub>C D"           
  8.2484 +  thus "\<And> c. \<lbrakk>class G C = Some c; x\<in>superclasses G D\<rbrakk>\<Longrightarrow>
  8.2485 +        x\<in>superclasses G C" (is "PROP ?P C"  
  8.2486 +                             is "\<And> c. ?CLS C c \<Longrightarrow> ?SUP D \<Longrightarrow> ?SUP C")
  8.2487 +  proof (induct ?P C  rule: converse_trancl_induct)
  8.2488 +    fix C c
  8.2489 +    assume "G\<turnstile>C\<prec>\<^sub>C\<^sub>1D" "class G C = Some c" "x \<in> superclasses G D"
  8.2490 +    with wf ws show "?SUP C"
  8.2491 +      by (auto    intro: no_subcls1_Object 
  8.2492 +               simp add: superclasses_rec subcls1_def)
  8.2493 +  next
  8.2494 +    fix C S c
  8.2495 +    assume clsrel': "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" "G\<turnstile>S \<prec>\<^sub>C D"
  8.2496 +       and    hyp : "\<And> s. \<lbrakk>class G S = Some s; x \<in> superclasses G D\<rbrakk>
  8.2497 +                           \<Longrightarrow> x \<in> superclasses G S"
  8.2498 +       and  cls_C': "class G C = Some c"
  8.2499 +       and       x: "x \<in> superclasses G D"
  8.2500 +    moreover note wf ws
  8.2501 +    moreover from calculation 
  8.2502 +    have "?SUP S" 
  8.2503 +      by (force intro: no_subcls1_Object simp add: subcls1_def)
  8.2504 +    moreover from calculation 
  8.2505 +    have "super c = S" 
  8.2506 +      by (auto intro: no_subcls1_Object simp add: subcls1_def)
  8.2507 +    ultimately show "?SUP C" 
  8.2508 +      by (auto intro: no_subcls1_Object simp add: superclasses_rec)
  8.2509 +  qed
  8.2510 +qed
  8.2511 +
  8.2512 +lemma subclsEval:
  8.2513 +"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D;ws_prog G; class G C = Some c;
  8.2514 +  \<And> C c. \<lbrakk>class G C = Some c;C\<noteq>Object\<rbrakk> \<Longrightarrow> \<exists> sc. class G (super c) = Some sc 
  8.2515 + \<rbrakk> \<Longrightarrow> D\<in>superclasses G C" 
  8.2516 +proof -
  8.2517 +  note converse_trancl_induct 
  8.2518 +       = converse_trancl_induct [consumes 1,case_names Single Step]
  8.2519 +  assume 
  8.2520 +             ws: "ws_prog G"          and 
  8.2521 +          cls_C: "class G C = Some c" and
  8.2522 +             wf: "\<And>C c. \<lbrakk>class G C = Some c; C \<noteq> Object\<rbrakk>
  8.2523 +                  \<Longrightarrow> \<exists>sc. class G (super c) = Some sc"
  8.2524 +  assume clsrel: "G\<turnstile>C\<prec>\<^sub>C D"           
  8.2525 +  thus "\<And> c. class G C = Some c\<Longrightarrow> D\<in>superclasses G C" 
  8.2526 +    (is "PROP ?P C"  is "\<And> c. ?CLS C c  \<Longrightarrow> ?SUP C")
  8.2527 +  proof (induct ?P C  rule: converse_trancl_induct)
  8.2528 +    fix C c
  8.2529 +    assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" "class G C = Some c"
  8.2530 +    with ws wf show "?SUP C"
  8.2531 +      by (auto intro: no_subcls1_Object simp add: superclasses_rec subcls1_def)
  8.2532 +  next
  8.2533 +    fix C S c
  8.2534 +    assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" "G\<turnstile>S\<prec>\<^sub>C D" 
  8.2535 +           "\<And>s. class G S = Some s \<Longrightarrow> D \<in> superclasses G S"
  8.2536 +           "class G C = Some c" 
  8.2537 +    with ws wf show "?SUP C"
  8.2538 +      by - (rule superclasses_mono,
  8.2539 +            auto dest: no_subcls1_Object simp add: subcls1_def )
  8.2540 +  qed
  8.2541 +qed
  8.2542 +
  8.2543 +end
  8.2544 \ No newline at end of file
     9.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     9.2 +++ b/src/HOL/Bali/Eval.thy	Mon Jan 28 17:00:19 2002 +0100
     9.3 @@ -0,0 +1,993 @@
     9.4 +(*  Title:      isabelle/Bali/Eval.thy
     9.5 +    ID:         $Id$
     9.6 +    Author:     David von Oheimb
     9.7 +    Copyright   1997 Technische Universitaet Muenchen
     9.8 +*)
     9.9 +header {* Operational evaluation (big-step) semantics of Java expressions and 
    9.10 +          statements
    9.11 +*}
    9.12 +
    9.13 +theory Eval = State + DeclConcepts:
    9.14 +
    9.15 +text {*
    9.16 +
    9.17 +improvements over Java Specification 1.0:
    9.18 +\begin{itemize}
    9.19 +\item dynamic method lookup does not need to consider the return type 
    9.20 +      (cf.15.11.4.4)
    9.21 +\item throw raises a NullPointer exception if a null reference is given, and 
    9.22 +      each throw of a standard exception yield a fresh exception object 
    9.23 +      (was not specified)
    9.24 +\item if there is not enough memory even to allocate an OutOfMemory exception,
    9.25 +  evaluation/execution fails, i.e. simply stops (was not specified)
    9.26 +\item array assignment checks lhs (and may throw exceptions) before evaluating 
    9.27 +      rhs
    9.28 +\item fixed exact positions of class initializations 
    9.29 +      (immediate at first active use)
    9.30 +\end{itemize}
    9.31 +
    9.32 +design issues:
    9.33 +\begin{itemize}
    9.34 +\item evaluation vs. (single-step) transition semantics
    9.35 +  evaluation semantics chosen, because:
    9.36 +  \begin{itemize} 
    9.37 +  \item[++] less verbose and therefore easier to read (and to handle in proofs)
    9.38 +  \item[+]  more abstract
    9.39 +  \item[+]  intermediate values (appearing in recursive rules) need not be 
    9.40 +     stored explicitly, e.g. no call body construct or stack of invocation 
    9.41 +     frames containing local variables and return addresses for method calls 
    9.42 +     needed
    9.43 +  \item[+]  convenient rule induction for subject reduction theorem
    9.44 +  \item[-]  no interleaving (for parallelism) can be described
    9.45 +  \item[-]  stating a property of infinite executions requires the meta-level 
    9.46 +     argument that this property holds for any finite prefixes of it 
    9.47 +     (e.g. stopped using a counter that is decremented to zero and then 
    9.48 +     throwing an exception)
    9.49 +  \end{itemize}
    9.50 +\item unified evaluation for variables, expressions, expression lists, 
    9.51 +      statements
    9.52 +\item the value entry in statement rules is redundant 
    9.53 +\item the value entry in rules is irrelevant in case of exceptions, but its full
    9.54 +  inclusion helps to make the rule structure independent of exception occurence.
    9.55 +\item as irrelevant value entries are ignored, it does not matter if they are 
    9.56 +      unique.
    9.57 +  For simplicity, (fixed) arbitrary values are preferred over "free" values.
    9.58 +\item the rule format is such that the start state may contain an exception.
    9.59 +  \begin{itemize}
    9.60 +  \item[++] faciliates exception handling
    9.61 +  \item[+]  symmetry
    9.62 +  \end{itemize}
    9.63 +\item the rules are defined carefully in order to be applicable even in not
    9.64 +  type-correct situations (yielding undefined values),
    9.65 +  e.g. @{text "the_Addr (Val (Bool b)) = arbitrary"}.
    9.66 +  \begin{itemize}
    9.67 +  \item[++] fewer rules 
    9.68 +  \item[-]  less readable because of auxiliary functions like @{text the_Addr}
    9.69 +  \end{itemize}
    9.70 +  Alternative: "defensive" evaluation throwing some InternalError exception
    9.71 +               in case of (impossible, for correct programs) type mismatches
    9.72 +\item there is exactly one rule per syntactic construct
    9.73 +  \begin{itemize}
    9.74 +  \item[+] no redundancy in case distinctions
    9.75 +  \end{itemize}
    9.76 +\item halloc fails iff there is no free heap address. When there is
    9.77 +  only one free heap address left, it returns an OutOfMemory exception.
    9.78 +  In this way it is guaranteed that when an OutOfMemory exception is thrown for
    9.79 +  the first time, there is a free location on the heap to allocate it.
    9.80 +\item the allocation of objects that represent standard exceptions is deferred 
    9.81 +      until execution of any enclosing catch clause, which is transparent to 
    9.82 +      the program.
    9.83 +  \begin{itemize}
    9.84 +  \item[-]  requires an auxiliary execution relation
    9.85 +  \item[++] avoids copies of allocation code and awkward case distinctions 
    9.86 +           (whether there is enough memory to allocate the exception) in 
    9.87 +            evaluation rules
    9.88 +  \end{itemize}
    9.89 +\item unfortunately @{text new_Addr} is not directly executable because of 
    9.90 +      Hilbert operator.
    9.91 +\end{itemize}
    9.92 +simplifications:
    9.93 +\begin{itemize}
    9.94 +\item local variables are initialized with default values 
    9.95 +      (no definite assignment)
    9.96 +\item garbage collection not considered, therefore also no finalizers
    9.97 +\item stack overflow and memory overflow during class initialization not 
    9.98 +      modelled
    9.99 +\item exceptions in initializations not replaced by ExceptionInInitializerError
   9.100 +\end{itemize}
   9.101 +*}
   9.102 +
   9.103 +types vvar  =         "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
   9.104 +      vals  =        "(val, vvar, val list) sum3"
   9.105 +translations
   9.106 +     "vvar" <= (type) "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
   9.107 +     "vals" <= (type)"(val, vvar, val list) sum3"
   9.108 +
   9.109 +syntax (xsymbols)
   9.110 +  dummy_res :: "vals" ("\<diamondsuit>")
   9.111 +translations
   9.112 +  "\<diamondsuit>" == "In1 Unit"
   9.113 +
   9.114 +constdefs
   9.115 +  arbitrary3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals"
   9.116 + "arbitrary3 \<equiv> sum3_case (In1 \<circ> sum_case (\<lambda>x. arbitrary) (\<lambda>x. Unit))
   9.117 +                     (\<lambda>x. In2 arbitrary) (\<lambda>x. In3 arbitrary)"
   9.118 +
   9.119 +lemma [simp]: "arbitrary3 (In1l x) = In1 arbitrary"
   9.120 +by (simp add: arbitrary3_def)
   9.121 +
   9.122 +lemma [simp]: "arbitrary3 (In1r x) = \<diamondsuit>"
   9.123 +by (simp add: arbitrary3_def)
   9.124 +
   9.125 +lemma [simp]: "arbitrary3 (In2  x) = In2 arbitrary"
   9.126 +by (simp add: arbitrary3_def)
   9.127 +
   9.128 +lemma [simp]: "arbitrary3 (In3  x) = In3 arbitrary"
   9.129 +by (simp add: arbitrary3_def)
   9.130 +
   9.131 +
   9.132 +section "exception throwing and catching"
   9.133 +
   9.134 +constdefs
   9.135 +  throw :: "val \<Rightarrow> abopt \<Rightarrow> abopt"
   9.136 + "throw a' x \<equiv> abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)"
   9.137 +
   9.138 +lemma throw_def2: 
   9.139 + "throw a' x = abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)"
   9.140 +apply (unfold throw_def)
   9.141 +apply (simp (no_asm))
   9.142 +done
   9.143 +
   9.144 +constdefs
   9.145 +  fits    :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60)
   9.146 + "G,s\<turnstile>a' fits T  \<equiv> (\<exists>rt. T=RefT rt) \<longrightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"
   9.147 +
   9.148 +lemma fits_Null [simp]: "G,s\<turnstile>Null fits T"
   9.149 +by (simp add: fits_def)
   9.150 +
   9.151 +
   9.152 +lemma fits_Addr_RefT [simp]:
   9.153 +  "G,s\<turnstile>Addr a fits RefT t = G\<turnstile>obj_ty (the (heap s a))\<preceq>RefT t"
   9.154 +by (simp add: fits_def)
   9.155 +
   9.156 +lemma fitsD: "\<And>X. G,s\<turnstile>a' fits T \<Longrightarrow> (\<exists>pt. T = PrimT pt) \<or>  
   9.157 +  (\<exists>t. T = RefT t) \<and> a' = Null \<or>  
   9.158 +  (\<exists>t. T = RefT t) \<and> a' \<noteq> Null \<and>  G\<turnstile>obj_ty (lookup_obj s a')\<preceq>T"
   9.159 +apply (unfold fits_def)
   9.160 +apply (case_tac "\<exists>pt. T = PrimT pt")
   9.161 +apply  simp_all
   9.162 +apply (case_tac "T")
   9.163 +defer 
   9.164 +apply (case_tac "a' = Null")
   9.165 +apply  simp_all
   9.166 +done
   9.167 +
   9.168 +constdefs
   9.169 +  catch ::"prog \<Rightarrow> state \<Rightarrow> qtname \<Rightarrow> bool"      ("_,_\<turnstile>catch _"[61,61,61]60)
   9.170 + "G,s\<turnstile>catch C\<equiv>\<exists>xc. abrupt s=Some (Xcpt xc) \<and> 
   9.171 +                    G,store s\<turnstile>Addr (the_Loc xc) fits Class C"
   9.172 +
   9.173 +lemma catch_Norm [simp]: "\<not>G,Norm s\<turnstile>catch tn"
   9.174 +apply (unfold catch_def)
   9.175 +apply (simp (no_asm))
   9.176 +done
   9.177 +
   9.178 +lemma catch_XcptLoc [simp]: 
   9.179 +  "G,(Some (Xcpt (Loc a)),s)\<turnstile>catch C = G,s\<turnstile>Addr a fits Class C"
   9.180 +apply (unfold catch_def)
   9.181 +apply (simp (no_asm))
   9.182 +done
   9.183 +
   9.184 +constdefs
   9.185 +  new_xcpt_var :: "vname \<Rightarrow> state \<Rightarrow> state"
   9.186 + "new_xcpt_var vn \<equiv> 
   9.187 +     \<lambda>(x,s). Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)"
   9.188 +
   9.189 +lemma new_xcpt_var_def2 [simp]: 
   9.190 + "new_xcpt_var vn (x,s) = 
   9.191 +    Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)"
   9.192 +apply (unfold new_xcpt_var_def)
   9.193 +apply (simp (no_asm))
   9.194 +done
   9.195 +
   9.196 +
   9.197 +
   9.198 +section "misc"
   9.199 +
   9.200 +constdefs
   9.201 +
   9.202 +  assign     :: "('a \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> 'a \<Rightarrow> state \<Rightarrow> state"
   9.203 + "assign f v \<equiv> \<lambda>(x,s). let (x',s') = (if x = None then f v else id) (x,s)
   9.204 +		   in  (x',if x' = None then s' else s)"
   9.205 +
   9.206 +(*
   9.207 +lemma assign_Norm_Norm [simp]: 
   9.208 +"f v \<lparr>abrupt=None,store=s\<rparr> = \<lparr>abrupt=None,store=s'\<rparr> 
   9.209 + \<Longrightarrow> assign f v \<lparr>abrupt=None,store=s\<rparr> = \<lparr>abrupt=None,store=s'\<rparr>"
   9.210 +by (simp add: assign_def Let_def)
   9.211 +*)
   9.212 +
   9.213 +lemma assign_Norm_Norm [simp]: 
   9.214 +"f v (Norm s) = Norm s' \<Longrightarrow> assign f v (Norm s) = Norm s'"
   9.215 +by (simp add: assign_def Let_def)
   9.216 +
   9.217 +(*
   9.218 +lemma assign_Norm_Some [simp]: 
   9.219 +  "\<lbrakk>abrupt (f v \<lparr>abrupt=None,store=s\<rparr>) = Some y\<rbrakk> 
   9.220 +   \<Longrightarrow> assign f v \<lparr>abrupt=None,store=s\<rparr> = \<lparr>abrupt=Some y,store =s\<rparr>"
   9.221 +by (simp add: assign_def Let_def split_beta)
   9.222 +*)
   9.223 +
   9.224 +lemma assign_Norm_Some [simp]: 
   9.225 +  "\<lbrakk>abrupt (f v (Norm s)) = Some y\<rbrakk> 
   9.226 +   \<Longrightarrow> assign f v (Norm s) = (Some y,s)"
   9.227 +by (simp add: assign_def Let_def split_beta)
   9.228 +
   9.229 +
   9.230 +lemma assign_Some [simp]: 
   9.231 +"assign f v (Some x,s) = (Some x,s)" 
   9.232 +by (simp add: assign_def Let_def split_beta)
   9.233 +
   9.234 +lemma assign_supd [simp]: 
   9.235 +"assign (\<lambda>v. supd (f v)) v (x,s)  
   9.236 +  = (x, if x = None then f v s else s)"
   9.237 +apply auto
   9.238 +done
   9.239 +
   9.240 +lemma assign_raise_if [simp]: 
   9.241 +  "assign (\<lambda>v (x,s). ((raise_if (b s v) xcpt) x, f v s)) v (x, s) =  
   9.242 +  (raise_if (b s v) xcpt x, if x=None \<and> \<not>b s v then f v s else s)"
   9.243 +apply (case_tac "x = None")
   9.244 +apply auto
   9.245 +done
   9.246 +
   9.247 +(*
   9.248 +lemma assign_raise_if [simp]: 
   9.249 +  "assign (\<lambda>v s. \<lparr>abrupt=(raise_if (b (store s) v) xcpt) (abrupt s),
   9.250 +                  store = f v (store s)\<rparr>) v s =  
   9.251 +  \<lparr>abrupt=raise_if (b (store s) v) xcpt (abrupt s),
   9.252 +   store= if (abrupt s)=None \<and> \<not>b (store s) v 
   9.253 +             then f v (store s) else (store s)\<rparr>"
   9.254 +apply (case_tac "abrupt s = None")
   9.255 +apply auto
   9.256 +done
   9.257 +*)
   9.258 +
   9.259 +constdefs
   9.260 +
   9.261 +  init_comp_ty :: "ty \<Rightarrow> stmt"
   9.262 + "init_comp_ty T \<equiv> if (\<exists>C. T = Class C) then Init (the_Class T) else Skip"
   9.263 +
   9.264 +lemma init_comp_ty_PrimT [simp]: "init_comp_ty (PrimT pt) = Skip"
   9.265 +apply (unfold init_comp_ty_def)
   9.266 +apply (simp (no_asm))
   9.267 +done
   9.268 +
   9.269 +constdefs
   9.270 +
   9.271 +(*
   9.272 +  target  :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname"
   9.273 + "target m s a' t 
   9.274 +    \<equiv> if m = IntVir
   9.275 +	 then obj_class (lookup_obj s a') 
   9.276 +         else the_Class (RefT t)"
   9.277 +*)
   9.278 +
   9.279 + invocation_class  :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname"
   9.280 + "invocation_class m s a' statT 
   9.281 +    \<equiv> (case m of
   9.282 +         Static \<Rightarrow> if (\<exists> statC. statT = ClassT statC) 
   9.283 +                      then the_Class (RefT statT) 
   9.284 +                      else Object
   9.285 +       | SuperM \<Rightarrow> the_Class (RefT statT)
   9.286 +       | IntVir \<Rightarrow> obj_class (lookup_obj s a'))"
   9.287 +
   9.288 +invocation_declclass::"prog \<Rightarrow> inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> qtname"
   9.289 +"invocation_declclass G m s a' statT sig 
   9.290 +   \<equiv> declclass (the (dynlookup G statT 
   9.291 +                                (invocation_class m s a' statT)
   9.292 +                                sig))" 
   9.293 +  
   9.294 +lemma invocation_class_IntVir [simp]: 
   9.295 +"invocation_class IntVir s a' statT = obj_class (lookup_obj s a')"
   9.296 +by (simp add: invocation_class_def)
   9.297 +
   9.298 +lemma dynclass_SuperM [simp]: 
   9.299 + "invocation_class SuperM s a' statT = the_Class (RefT statT)"
   9.300 +by (simp add: invocation_class_def)
   9.301 +(*
   9.302 +lemma invocation_class_notIntVir [simp]: 
   9.303 + "m \<noteq> IntVir \<Longrightarrow> invocation_class m s a' statT = the_Class (RefT statT)"
   9.304 +by (simp add: invocation_class_def)
   9.305 +*)
   9.306 +
   9.307 +lemma invocation_class_Static [simp]: 
   9.308 +  "invocation_class Static s a' statT = (if (\<exists> statC. statT = ClassT statC) 
   9.309 +                                            then the_Class (RefT statT) 
   9.310 +                                            else Object)"
   9.311 +by (simp add: invocation_class_def)
   9.312 +
   9.313 +constdefs
   9.314 +  init_lvars :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> inv_mode \<Rightarrow> val \<Rightarrow> val list \<Rightarrow>
   9.315 +		   state \<Rightarrow> state"
   9.316 + "init_lvars G C sig mode a' pvs 
   9.317 +   \<equiv> \<lambda> (x,s). 
   9.318 +      let m = mthd (the (methd G C sig));
   9.319 +          l = \<lambda> k. 
   9.320 +              (case k of
   9.321 +                 EName e 
   9.322 +                   \<Rightarrow> (case e of 
   9.323 +                         VNam v \<Rightarrow> (init_vals (table_of (lcls (mbody m)))
   9.324 +                                                     ((pars m)[\<mapsto>]pvs)) v
   9.325 +                       | Res    \<Rightarrow> Some (default_val (resTy m)))
   9.326 +               | This 
   9.327 +                   \<Rightarrow> (if mode=Static then None else Some a'))
   9.328 +      in set_lvars l (if mode = Static then x else np a' x,s)"
   9.329 +
   9.330 +
   9.331 +
   9.332 +lemma init_lvars_def2: "init_lvars G C sig mode a' pvs (x,s) =  
   9.333 +  set_lvars 
   9.334 +    (\<lambda> k. 
   9.335 +       (case k of
   9.336 +          EName e 
   9.337 +            \<Rightarrow> (case e of 
   9.338 +                  VNam v 
   9.339 +                  \<Rightarrow> (init_vals 
   9.340 +                       (table_of (lcls (mbody (mthd (the (methd G C sig))))))
   9.341 +                                 ((pars (mthd (the (methd G C sig))))[\<mapsto>]pvs)) v
   9.342 +               | Res \<Rightarrow> Some (default_val (resTy (mthd (the (methd G C sig))))))
   9.343 +        | This 
   9.344 +            \<Rightarrow> (if mode=Static then None else Some a')))
   9.345 +    (if mode = Static then x else np a' x,s)"
   9.346 +apply (unfold init_lvars_def)
   9.347 +apply (simp (no_asm) add: Let_def)
   9.348 +done
   9.349 +
   9.350 +constdefs
   9.351 +  body :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> expr"
   9.352 + "body G C sig \<equiv> let m = the (methd G C sig) 
   9.353 +                 in Body (declclass m) (stmt (mbody (mthd m)))"
   9.354 +
   9.355 +lemma body_def2: 
   9.356 +"body G C sig = Body  (declclass (the (methd G C sig))) 
   9.357 +                      (stmt (mbody (mthd (the (methd G C sig)))))"
   9.358 +apply (unfold body_def Let_def)
   9.359 +apply auto
   9.360 +done
   9.361 +
   9.362 +section "variables"
   9.363 +
   9.364 +constdefs
   9.365 +
   9.366 +  lvar :: "lname \<Rightarrow> st \<Rightarrow> vvar"
   9.367 + "lvar vn s \<equiv> (the (locals s vn), \<lambda>v. supd (lupd(vn\<mapsto>v)))"
   9.368 +
   9.369 +  fvar :: "qtname \<Rightarrow> bool \<Rightarrow> vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
   9.370 + "fvar C stat fn a' s 
   9.371 +    \<equiv> let (oref,xf) = if stat then (Stat C,id)
   9.372 +                              else (Heap (the_Addr a'),np a');
   9.373 +	          n = Inl (fn,C); 
   9.374 +                  f = (\<lambda>v. supd (upd_gobj oref n v)) 
   9.375 +      in ((the (values (the (globs (store s) oref)) n),f),abupd xf s)"
   9.376 +
   9.377 +  avar :: "prog \<Rightarrow> val \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
   9.378 + "avar G i' a' s 
   9.379 +    \<equiv> let   oref = Heap (the_Addr a'); 
   9.380 +               i = the_Intg i'; 
   9.381 +               n = Inr i;
   9.382 +        (T,k,cs) = the_Arr (globs (store s) oref); 
   9.383 +               f = (\<lambda>v (x,s). (raise_if (\<not>G,s\<turnstile>v fits T) 
   9.384 +                                           ArrStore x
   9.385 +                              ,upd_gobj oref n v s)) 
   9.386 +      in ((the (cs n),f)
   9.387 +         ,abupd (raise_if (\<not>i in_bounds k) IndOutBound \<circ> np a') s)"
   9.388 +
   9.389 +lemma fvar_def2: "fvar C stat fn a' s =  
   9.390 +  ((the 
   9.391 +     (values 
   9.392 +      (the (globs (store s) (if stat then Stat C else Heap (the_Addr a')))) 
   9.393 +      (Inl (fn,C)))
   9.394 +   ,(\<lambda>v. supd (upd_gobj (if stat then Stat C else Heap (the_Addr a')) 
   9.395 +                        (Inl (fn,C)) 
   9.396 +                        v)))
   9.397 +  ,abupd (if stat then id else np a') s)
   9.398 +  "
   9.399 +apply (unfold fvar_def)
   9.400 +apply (simp (no_asm) add: Let_def split_beta)
   9.401 +done
   9.402 +
   9.403 +lemma avar_def2: "avar G i' a' s =  
   9.404 +  ((the ((snd(snd(the_Arr (globs (store s) (Heap (the_Addr a')))))) 
   9.405 +           (Inr (the_Intg i')))
   9.406 +   ,(\<lambda>v (x,s').  (raise_if (\<not>G,s'\<turnstile>v fits (fst(the_Arr (globs (store s)
   9.407 +                                                   (Heap (the_Addr a')))))) 
   9.408 +                            ArrStore x
   9.409 +                 ,upd_gobj  (Heap (the_Addr a')) 
   9.410 +                               (Inr (the_Intg i')) v s')))
   9.411 +  ,abupd (raise_if (\<not>(the_Intg i') in_bounds (fst(snd(the_Arr (globs (store s) 
   9.412 +                   (Heap (the_Addr a'))))))) IndOutBound \<circ> np a')
   9.413 +          s)"
   9.414 +apply (unfold avar_def)
   9.415 +apply (simp (no_asm) add: Let_def split_beta)
   9.416 +done
   9.417 +
   9.418 +
   9.419 +section "evaluation judgments"
   9.420 +
   9.421 +consts
   9.422 +  eval   :: "prog \<Rightarrow> (state \<times> term    \<times> vals \<times> state) set"
   9.423 +  halloc::  "prog \<Rightarrow> (state \<times> obj_tag \<times> loc  \<times> state) set"
   9.424 +  sxalloc:: "prog \<Rightarrow> (state                  \<times> state) set"
   9.425 +
   9.426 +
   9.427 +syntax
   9.428 +eval ::"[prog,state,term,vals*state]=>bool"("_|-_ -_>-> _"  [61,61,80,   61]60)
   9.429 +exec ::"[prog,state,stmt      ,state]=>bool"("_|-_ -_-> _"   [61,61,65,   61]60)
   9.430 +evar ::"[prog,state,var  ,vvar,state]=>bool"("_|-_ -_=>_-> _"[61,61,90,61,61]60)
   9.431 +eval_::"[prog,state,expr ,val, state]=>bool"("_|-_ -_->_-> _"[61,61,80,61,61]60)
   9.432 +evals::"[prog,state,expr list ,
   9.433 +		    val  list ,state]=>bool"("_|-_ -_#>_-> _"[61,61,61,61,61]60)
   9.434 +hallo::"[prog,state,obj_tag,
   9.435 +	             loc,state]=>bool"("_|-_ -halloc _>_-> _"[61,61,61,61,61]60)
   9.436 +sallo::"[prog,state        ,state]=>bool"("_|-_ -sxalloc-> _"[61,61,      61]60)
   9.437 +
   9.438 +syntax (xsymbols)
   9.439 +eval ::"[prog,state,term,vals\<times>state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>_\<succ>\<rightarrow> _"  [61,61,80,   61]60)
   9.440 +exec ::"[prog,state,stmt      ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<rightarrow> _"   [61,61,65,   61]60)
   9.441 +evar ::"[prog,state,var  ,vvar,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_=\<succ>_\<rightarrow> _"[61,61,90,61,61]60)
   9.442 +eval_::"[prog,state,expr ,val ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_-\<succ>_\<rightarrow> _"[61,61,80,61,61]60)
   9.443 +evals::"[prog,state,expr list ,
   9.444 +		    val  list ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
   9.445 +hallo::"[prog,state,obj_tag,
   9.446 +	             loc,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>halloc _\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
   9.447 +sallo::"[prog,state,        state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>sxalloc\<rightarrow> _"[61,61,      61]60)
   9.448 +
   9.449 +translations
   9.450 +  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow>  w___s' " == "(s,t,w___s') \<in> eval G"
   9.451 +  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow> (w,  s')" <= "(s,t,w,  s') \<in> eval G"
   9.452 +  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow> (w,x,s')" <= "(s,t,w,x,s') \<in> eval G"
   9.453 +  "G\<turnstile>s \<midarrow>c    \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In1r c\<succ>\<rightarrow> (\<diamondsuit>,x,s')"
   9.454 +  "G\<turnstile>s \<midarrow>c    \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In1r c\<succ>\<rightarrow> (\<diamondsuit>  ,  s')"
   9.455 +  "G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In1l e\<succ>\<rightarrow> (In1 v ,x,s')"
   9.456 +  "G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In1l e\<succ>\<rightarrow> (In1 v ,  s')"
   9.457 +  "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In2  e\<succ>\<rightarrow> (In2 vf,x,s')"
   9.458 +  "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In2  e\<succ>\<rightarrow> (In2 vf,  s')"
   9.459 +  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In3  e\<succ>\<rightarrow> (In3 v ,x,s')"
   9.460 +  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In3  e\<succ>\<rightarrow> (In3 v ,  s')"
   9.461 +  "G\<turnstile>s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,s')" <= "(s,oi,a,x,s') \<in> halloc G"
   9.462 +  "G\<turnstile>s \<midarrow>halloc oi\<succ>a\<rightarrow>    s' " == "(s,oi,a,  s') \<in> halloc G"
   9.463 +  "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>     (x,s')" <= "(s     ,x,s') \<in> sxalloc G"
   9.464 +  "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>        s' " == "(s     ,  s') \<in> sxalloc G"
   9.465 +
   9.466 +inductive "halloc G" intros (* allocating objects on the heap, cf. 12.5 *)
   9.467 +
   9.468 +  Abrupt: 
   9.469 +  "G\<turnstile>(Some x,s) \<midarrow>halloc oi\<succ>arbitrary\<rightarrow> (Some x,s)"
   9.470 +
   9.471 +  New:	  "\<lbrakk>new_Addr (heap s) = Some a; 
   9.472 +	    (x,oi') = (if atleast_free (heap s) (Suc (Suc 0)) then (None,oi)
   9.473 +		       else (Some (Xcpt (Loc a)),CInst (SXcpt OutOfMemory)))\<rbrakk>
   9.474 +            \<Longrightarrow>
   9.475 +	    G\<turnstile>Norm s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,init_obj G oi' (Heap a) s)"
   9.476 +
   9.477 +inductive "sxalloc G" intros (* allocating exception objects for
   9.478 +	 	 	      standard exceptions (other than OutOfMemory) *)
   9.479 +
   9.480 +  Norm:	 "G\<turnstile> Norm              s   \<midarrow>sxalloc\<rightarrow>  Norm             s"
   9.481 +
   9.482 +  XcptL: "G\<turnstile>(Some (Xcpt (Loc a) ),s)  \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s)"
   9.483 +
   9.484 +  SXcpt: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>halloc (CInst (SXcpt xn))\<succ>a\<rightarrow> (x,s1)\<rbrakk> \<Longrightarrow>
   9.485 +	  G\<turnstile>(Some (Xcpt (Std xn)),s0) \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s1)"
   9.486 +
   9.487 +
   9.488 +inductive "eval G" intros
   9.489 +
   9.490 +(* propagation of abrupt completion *)
   9.491 +
   9.492 +  (* cf. 14.1, 15.5 *)
   9.493 +  Abrupt: 
   9.494 +   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (arbitrary3 t,(Some xc,s))"
   9.495 +
   9.496 +
   9.497 +(* execution of statements *)
   9.498 +
   9.499 +  (* cf. 14.5 *)
   9.500 +  Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<rightarrow> Norm s"
   9.501 +
   9.502 +  (* cf. 14.7 *)
   9.503 +  Expr:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
   9.504 +				  G\<turnstile>Norm s0 \<midarrow>Expr e\<rightarrow> s1"
   9.505 +
   9.506 +  Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<rightarrow> s1\<rbrakk> \<Longrightarrow>
   9.507 +                                G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<rightarrow> abupd (absorb (Break l)) s1"
   9.508 +  (* cf. 14.2 *)
   9.509 +  Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<rightarrow> s1;
   9.510 +	  G\<turnstile>     s1 \<midarrow>c2 \<rightarrow> s2\<rbrakk> \<Longrightarrow>
   9.511 +				 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<rightarrow> s2"
   9.512 +
   9.513 +
   9.514 +  (* cf. 14.8.2 *)
   9.515 +  If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
   9.516 +	  G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<rbrakk> \<Longrightarrow>
   9.517 +		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<rightarrow> s2"
   9.518 +
   9.519 +  (* cf. 14.10, 14.10.1 *)
   9.520 +  (*      G\<turnstile>Norm s0 \<midarrow>If(e) (c;; While(e) c) Else Skip\<rightarrow> s3 *)
   9.521 +  (* A "continue jump" from the while body c is handled by 
   9.522 +     this rule. If a continue jump with the proper label was invoked inside c
   9.523 +     this label (Cont l) is deleted out of the abrupt component of the state 
   9.524 +     before the iterative evaluation of the while statement.
   9.525 +     A "break jump" is handled by the Lab Statement (Lab l (while\<dots>).
   9.526 +  *)
   9.527 +  Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
   9.528 +	  if normal s1 \<and> the_Bool b 
   9.529 +             then (G\<turnstile>s1 \<midarrow>c\<rightarrow> s2 \<and> 
   9.530 +                   G\<turnstile>(abupd (absorb (Cont l)) s2) \<midarrow>l\<bullet> While(e) c\<rightarrow> s3)
   9.531 +	     else s3 = s1\<rbrakk> \<Longrightarrow>
   9.532 +			      G\<turnstile>Norm s0 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
   9.533 +
   9.534 +  Do: "G\<turnstile>Norm s \<midarrow>Do j\<rightarrow> (Some (Jump j), s)"
   9.535 +   
   9.536 +  (* cf. 14.16 *)
   9.537 +  Throw: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1\<rbrakk> \<Longrightarrow>
   9.538 +				 G\<turnstile>Norm s0 \<midarrow>Throw e\<rightarrow> abupd (throw a') s1"
   9.539 +
   9.540 +  (* cf. 14.18.1 *)
   9.541 +  Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; 
   9.542 +	  if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2\<rbrakk> \<Longrightarrow>
   9.543 +		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(C vn) c2\<rightarrow> s3"
   9.544 +
   9.545 +  (* cf. 14.18.2 *)
   9.546 +  Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1,s1);
   9.547 +	  G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2\<rbrakk> \<Longrightarrow>
   9.548 +               G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<rightarrow> abupd (abrupt_if (x1\<noteq>None) x1) s2"
   9.549 +
   9.550 +  (* cf. 12.4.2, 8.5 *)
   9.551 +  Init:	"\<lbrakk>the (class G C) = c;
   9.552 +	  if inited C (globs s0) then s3 = Norm s0
   9.553 +	  else (G\<turnstile>Norm (init_class_obj G C s0) 
   9.554 +		  \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1 \<and>
   9.555 +	       G\<turnstile>set_lvars empty s1 \<midarrow>init c\<rightarrow> s2 \<and> s3 = restore_lvars s1 s2)\<rbrakk> 
   9.556 +              \<Longrightarrow>
   9.557 +		 G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s3"
   9.558 +
   9.559 +(* evaluation of expressions *)
   9.560 +
   9.561 +  (* cf. 15.8.1, 12.4.1 *)
   9.562 +  NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1;
   9.563 +	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
   9.564 +	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
   9.565 +
   9.566 +  (* cf. 15.9.1, 12.4.1 *)
   9.567 +  NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<rightarrow> s2; 
   9.568 +	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
   9.569 +	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<rightarrow> s3"
   9.570 +
   9.571 +  (* cf. 15.15 *)
   9.572 +  Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
   9.573 +	  s2 = abupd (raise_if (\<not>G,store s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
   9.574 +			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<rightarrow> s2"
   9.575 +
   9.576 +  (* cf. 15.19.2 *)
   9.577 +  Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
   9.578 +	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
   9.579 +			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<rightarrow> s1"
   9.580 +
   9.581 +  (* cf. 15.7.1 *)
   9.582 +  Lit:	"G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<rightarrow> Norm s"
   9.583 +
   9.584 +
   9.585 +
   9.586 +  (* cf. 15.10.2 *)
   9.587 +  Super: "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<rightarrow> Norm s"
   9.588 +
   9.589 +  (* cf. 15.2 *)
   9.590 +  Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<rightarrow> s1\<rbrakk> \<Longrightarrow>
   9.591 +	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<rightarrow> s1"
   9.592 +
   9.593 +  (* cf. 15.25.1 *)
   9.594 +  Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<rightarrow> s1;
   9.595 +          G\<turnstile>     s1 \<midarrow>e-\<succ>v  \<rightarrow> s2\<rbrakk> \<Longrightarrow>
   9.596 +				   G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v\<rightarrow> assign f v s2"
   9.597 +
   9.598 +  (* cf. 15.24 *)
   9.599 +  Cond:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1;
   9.600 +          G\<turnstile>     s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2\<rbrakk> \<Longrightarrow>
   9.601 +			    G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<rightarrow> s2"
   9.602 +
   9.603 +
   9.604 +  (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5 *)
   9.605 +  Call:	
   9.606 +  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1; G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2;
   9.607 +    D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>;
   9.608 +    G\<turnstile>init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs s2 
   9.609 +         \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<rightarrow> s3\<rbrakk>
   9.610 +   \<Longrightarrow>
   9.611 +       G\<turnstile>Norm s0 \<midarrow>{statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<rightarrow> (restore_lvars s2 s3)"
   9.612 +
   9.613 +  Methd:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
   9.614 +				G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<rightarrow> s1"
   9.615 +
   9.616 +  (* cf. 14.15, 12.4.1 *)
   9.617 +  Body:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<rightarrow> s2\<rbrakk> \<Longrightarrow>
   9.618 + G\<turnstile>Norm s0 \<midarrow>Body D c -\<succ>the (locals (store s2) Result)\<rightarrow>abupd (absorb Ret) s2"
   9.619 +
   9.620 +(* evaluation of variables *)
   9.621 +
   9.622 +  (* cf. 15.13.1, 15.7.2 *)
   9.623 +  LVar:	"G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<rightarrow> Norm s"
   9.624 +
   9.625 +  (* cf. 15.10.1, 12.4.1 *)
   9.626 +  FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2;
   9.627 +	  (v,s2') = fvar C stat fn a s2\<rbrakk> \<Longrightarrow>
   9.628 +	  G\<turnstile>Norm s0 \<midarrow>{C,stat}e..fn=\<succ>v\<rightarrow> s2'"
   9.629 +
   9.630 +  (* cf. 15.12.1, 15.25.1 *)
   9.631 +  AVar:	"\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2;
   9.632 +	  (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
   9.633 +	              G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<rightarrow> s2'"
   9.634 +
   9.635 +
   9.636 +(* evaluation of expression lists *)
   9.637 +
   9.638 +  (* cf. 15.11.4.2 *)
   9.639 +  Nil:
   9.640 +				    "G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<rightarrow> Norm s0"
   9.641 +
   9.642 +  (* cf. 15.6.4 *)
   9.643 +  Cons:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<rightarrow> s1;
   9.644 +          G\<turnstile>     s1 \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s2\<rbrakk> \<Longrightarrow>
   9.645 +				   G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<rightarrow> s2"
   9.646 +
   9.647 +(* Rearrangement of premisses:
   9.648 +[0,1(Abrupt),2(Skip),8(Do),4(Lab),28(Nil),29(Cons),25(LVar),15(Cast),16(Inst),
   9.649 + 17(Lit),18(Super),19(Acc),3(Expr),5(Comp),23(Methd),24(Body),21(Cond),6(If),
   9.650 + 7(Loop),11(Fin),9(Throw),13(NewC),14(NewA),12(Init),20(Ass),10(Try),26(FVar),
   9.651 + 27(AVar),22(Call)]
   9.652 +*)
   9.653 +ML {*
   9.654 +bind_thm ("eval_induct_", rearrange_prems 
   9.655 +[0,1,2,8,4,28,29,25,15,16,
   9.656 + 17,18,19,3,5,23,24,21,6,
   9.657 + 7,11,9,13,14,12,20,10,26,
   9.658 + 27,22] (thm "eval.induct"))
   9.659 +*}
   9.660 +
   9.661 +lemmas eval_induct = eval_induct_ [split_format and and and and and and and and
   9.662 +   and and and and s1 (* Acc *) and and s2 (* Comp *) and and and and and and 
   9.663 +   s2 (* Fin *) and and s2 (* NewC *)] 
   9.664 +
   9.665 +declare split_if     [split del] split_if_asm     [split del] 
   9.666 +        option.split [split del] option.split_asm [split del]
   9.667 +inductive_cases halloc_elim_cases: 
   9.668 +  "G\<turnstile>(Some xc,s) \<midarrow>halloc oi\<succ>a\<rightarrow> s'"
   9.669 +  "G\<turnstile>(Norm    s) \<midarrow>halloc oi\<succ>a\<rightarrow> s'"
   9.670 +
   9.671 +inductive_cases sxalloc_elim_cases:
   9.672 + 	"G\<turnstile> Norm                 s  \<midarrow>sxalloc\<rightarrow> s'"
   9.673 + 	"G\<turnstile>(Some (Xcpt (Loc a )),s) \<midarrow>sxalloc\<rightarrow> s'"
   9.674 + 	"G\<turnstile>(Some (Xcpt (Std xn)),s) \<midarrow>sxalloc\<rightarrow> s'"
   9.675 +inductive_cases sxalloc_cases: "G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'"
   9.676 +
   9.677 +lemma sxalloc_elim_cases2: "\<lbrakk>G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s';  
   9.678 +       \<And>s.   \<lbrakk>s' = Norm s\<rbrakk> \<Longrightarrow> P;  
   9.679 +       \<And>a s. \<lbrakk>s' = (Some (Xcpt (Loc a)),s)\<rbrakk> \<Longrightarrow> P  
   9.680 +      \<rbrakk> \<Longrightarrow> P"
   9.681 +apply cut_tac 
   9.682 +apply (erule sxalloc_cases)
   9.683 +apply blast+
   9.684 +done
   9.685 +
   9.686 +declare not_None_eq [simp del] (* IntDef.Zero_def [simp del] *)
   9.687 +declare split_paired_All [simp del] split_paired_Ex [simp del]
   9.688 +ML_setup {*
   9.689 +simpset_ref() := simpset() delloop "split_all_tac"
   9.690 +*}
   9.691 +inductive_cases eval_cases: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> vs'"
   9.692 +
   9.693 +inductive_cases eval_elim_cases:
   9.694 +        "G\<turnstile>(Some xc,s) \<midarrow>t                         \<succ>\<rightarrow> vs'"
   9.695 +	"G\<turnstile>Norm s \<midarrow>In1r Skip                      \<succ>\<rightarrow> xs'"
   9.696 +        "G\<turnstile>Norm s \<midarrow>In1r (Do j)                    \<succ>\<rightarrow> xs'"
   9.697 +        "G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> c)                    \<succ>\<rightarrow> xs'"
   9.698 +	"G\<turnstile>Norm s \<midarrow>In3  ([])                      \<succ>\<rightarrow> vs'"
   9.699 +	"G\<turnstile>Norm s \<midarrow>In3  (e#es)                    \<succ>\<rightarrow> vs'"
   9.700 +	"G\<turnstile>Norm s \<midarrow>In1l (Lit w)                   \<succ>\<rightarrow> vs'"
   9.701 +	"G\<turnstile>Norm s \<midarrow>In2  (LVar vn)                 \<succ>\<rightarrow> vs'"
   9.702 +	"G\<turnstile>Norm s \<midarrow>In1l (Cast T e)                \<succ>\<rightarrow> vs'"
   9.703 +	"G\<turnstile>Norm s \<midarrow>In1l (e InstOf T)              \<succ>\<rightarrow> vs'"
   9.704 +	"G\<turnstile>Norm s \<midarrow>In1l (Super)                   \<succ>\<rightarrow> vs'"
   9.705 +	"G\<turnstile>Norm s \<midarrow>In1l (Acc va)                  \<succ>\<rightarrow> vs'"
   9.706 +	"G\<turnstile>Norm s \<midarrow>In1r (Expr e)                  \<succ>\<rightarrow> xs'"
   9.707 +	"G\<turnstile>Norm s \<midarrow>In1r (c1;; c2)                 \<succ>\<rightarrow> xs'"
   9.708 +	"G\<turnstile>Norm s \<midarrow>In1l (Methd C sig)             \<succ>\<rightarrow> xs'"
   9.709 +	"G\<turnstile>Norm s \<midarrow>In1l (Body D c)                \<succ>\<rightarrow> xs'"
   9.710 +	"G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2)            \<succ>\<rightarrow> vs'"
   9.711 +	"G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2)        \<succ>\<rightarrow> xs'"
   9.712 +	"G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> While(e) c)           \<succ>\<rightarrow> xs'"
   9.713 +	"G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2)           \<succ>\<rightarrow> xs'"
   9.714 +	"G\<turnstile>Norm s \<midarrow>In1r (Throw e)                 \<succ>\<rightarrow> xs'"
   9.715 +	"G\<turnstile>Norm s \<midarrow>In1l (NewC C)                  \<succ>\<rightarrow> vs'"
   9.716 +	"G\<turnstile>Norm s \<midarrow>In1l (New T[e])                \<succ>\<rightarrow> vs'"
   9.717 +	"G\<turnstile>Norm s \<midarrow>In1l (Ass va e)                \<succ>\<rightarrow> vs'"
   9.718 +	"G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2)  \<succ>\<rightarrow> xs'"
   9.719 +	"G\<turnstile>Norm s \<midarrow>In2  ({C,stat}e..fn)           \<succ>\<rightarrow> vs'"
   9.720 +	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                 \<succ>\<rightarrow> vs'"
   9.721 +	"G\<turnstile>Norm s \<midarrow>In1l ({statT,mode}e\<cdot>mn({pT}p)) \<succ>\<rightarrow> vs'"
   9.722 +	"G\<turnstile>Norm s \<midarrow>In1r (Init C)                  \<succ>\<rightarrow> xs'"
   9.723 +declare not_None_eq [simp]  (* IntDef.Zero_def [simp] *)
   9.724 +declare split_paired_All [simp] split_paired_Ex [simp]
   9.725 +ML_setup {*
   9.726 +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
   9.727 +*}
   9.728 +declare split_if     [split] split_if_asm     [split] 
   9.729 +        option.split [split] option.split_asm [split]
   9.730 +
   9.731 +lemma eval_Inj_elim: 
   9.732 + "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') 
   9.733 + \<Longrightarrow> case t of 
   9.734 +       In1 ec \<Rightarrow> (case ec of 
   9.735 +                    Inl e \<Rightarrow> (\<exists>v. w = In1 v) 
   9.736 +                  | Inr c \<Rightarrow> w = \<diamondsuit>)  
   9.737 +     | In2 e \<Rightarrow> (\<exists>v. w = In2 v) 
   9.738 +     | In3 e \<Rightarrow> (\<exists>v. w = In3 v)"
   9.739 +apply (erule eval_cases)
   9.740 +apply auto
   9.741 +apply (induct_tac "t")
   9.742 +apply (induct_tac "a")
   9.743 +apply auto
   9.744 +done
   9.745 +
   9.746 +ML_setup {*
   9.747 +fun eval_fun nam inj rhs =
   9.748 +let
   9.749 +  val name = "eval_" ^ nam ^ "_eq"
   9.750 +  val lhs = "G\<turnstile>s \<midarrow>" ^ inj ^ " t\<succ>\<rightarrow> (w, s')"
   9.751 +  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
   9.752 +	(K [Auto_tac, ALLGOALS (ftac (thm "eval_Inj_elim")) THEN Auto_tac])
   9.753 +  fun is_Inj (Const (inj,_) $ _) = true
   9.754 +    | is_Inj _                   = false
   9.755 +  fun pred (_ $ (Const ("Pair",_) $ _ $ 
   9.756 +      (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ x $ _ ))) $ _ ) = is_Inj x
   9.757 +in
   9.758 +  make_simproc name lhs pred (thm name)
   9.759 +end
   9.760 +
   9.761 +val eval_expr_proc =eval_fun "expr" "In1l" "\<exists>v.  w=In1 v   \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<rightarrow> s'"
   9.762 +val eval_var_proc  =eval_fun "var"  "In2"  "\<exists>vf. w=In2 vf  \<and> G\<turnstile>s \<midarrow>t=\<succ>vf\<rightarrow> s'"
   9.763 +val eval_exprs_proc=eval_fun "exprs""In3"  "\<exists>vs. w=In3 vs  \<and> G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs\<rightarrow> s'"
   9.764 +val eval_stmt_proc =eval_fun "stmt" "In1r" "     w=\<diamondsuit> \<and> G\<turnstile>s \<midarrow>t    \<rightarrow> s'";
   9.765 +Addsimprocs [eval_expr_proc,eval_var_proc,eval_exprs_proc,eval_stmt_proc];
   9.766 +bind_thms ("AbruptIs", sum3_instantiate (thm "eval.Abrupt"))
   9.767 +*}
   9.768 +
   9.769 +declare halloc.Abrupt [intro!] eval.Abrupt [intro!]  AbruptIs [intro!] 
   9.770 +
   9.771 +
   9.772 +lemma eval_no_abrupt_lemma: 
   9.773 +   "\<And>s s'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<Longrightarrow> normal s' \<longrightarrow> normal s"
   9.774 +by (erule eval_cases, auto)
   9.775 +
   9.776 +lemma eval_no_abrupt: 
   9.777 +  "G\<turnstile>(x,s) \<midarrow>t\<succ>\<rightarrow> (w,Norm s') = 
   9.778 +        (x = None \<and> G\<turnstile>Norm s \<midarrow>t\<succ>\<rightarrow> (w,Norm s'))"
   9.779 +apply auto
   9.780 +apply (frule eval_no_abrupt_lemma, auto)+
   9.781 +done
   9.782 +
   9.783 +ML {*
   9.784 +local
   9.785 +  fun is_None (Const ("Option.option.None",_)) = true
   9.786 +    | is_None _ = false
   9.787 +  fun pred (t as (_ $ (Const ("Pair",_) $
   9.788 +     (Const ("Pair", _) $ x $ _) $ _ ) $ _)) = is_None x
   9.789 +in
   9.790 +  val eval_no_abrupt_proc = 
   9.791 +  make_simproc "eval_no_abrupt" "G\<turnstile>(x,s) \<midarrow>e\<succ>\<rightarrow> (w,Norm s')" pred 
   9.792 +          (thm "eval_no_abrupt")
   9.793 +end;
   9.794 +Addsimprocs [eval_no_abrupt_proc]
   9.795 +*}
   9.796 +
   9.797 +
   9.798 +lemma eval_abrupt_lemma: 
   9.799 +  "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s') \<Longrightarrow> abrupt s=Some xc \<longrightarrow> s'= s \<and> v = arbitrary3 t"
   9.800 +by (erule eval_cases, auto)
   9.801 +
   9.802 +lemma eval_abrupt: 
   9.803 + " G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (w,s') =  
   9.804 +     (s'=(Some xc,s) \<and> w=arbitrary3 t \<and> 
   9.805 +     G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (arbitrary3 t,(Some xc,s)))"
   9.806 +apply auto
   9.807 +apply (frule eval_abrupt_lemma, auto)+
   9.808 +done
   9.809 +
   9.810 +ML {*
   9.811 +local
   9.812 +  fun is_Some (Const ("Pair",_) $ (Const ("Option.option.Some",_) $ _)$ _) =true
   9.813 +    | is_Some _ = false
   9.814 +  fun pred (_ $ (Const ("Pair",_) $
   9.815 +     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
   9.816 +       x))) $ _ ) = is_Some x
   9.817 +in
   9.818 +  val eval_abrupt_proc = 
   9.819 +  make_simproc "eval_abrupt" 
   9.820 +               "G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<rightarrow> (w,s')" pred (thm "eval_abrupt")
   9.821 +end;
   9.822 +Addsimprocs [eval_abrupt_proc]
   9.823 +*}
   9.824 +
   9.825 +
   9.826 +lemma LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else arbitrary)\<rightarrow> s"
   9.827 +apply (case_tac "s", case_tac "a = None")
   9.828 +by (auto intro!: eval.Lit)
   9.829 +
   9.830 +lemma SkipI [intro!]: "G\<turnstile>s \<midarrow>Skip\<rightarrow> s"
   9.831 +apply (case_tac "s", case_tac "a = None")
   9.832 +by (auto intro!: eval.Skip)
   9.833 +
   9.834 +lemma ExprI: "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>Expr e\<rightarrow> s'"
   9.835 +apply (case_tac "s", case_tac "a = None")
   9.836 +by (auto intro!: eval.Expr)
   9.837 +
   9.838 +lemma CompI: "\<lbrakk>G\<turnstile>s \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2\<rbrakk> \<Longrightarrow> G\<turnstile>s \<midarrow>c1;; c2\<rightarrow> s2"
   9.839 +apply (case_tac "s", case_tac "a = None")
   9.840 +by (auto intro!: eval.Comp)
   9.841 +
   9.842 +lemma CondI: 
   9.843 +  "\<And>s1. \<lbrakk>G\<turnstile>s \<midarrow>e-\<succ>b\<rightarrow> s1; G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2\<rbrakk> \<Longrightarrow> 
   9.844 +         G\<turnstile>s \<midarrow>e ? e1 : e2-\<succ>(if normal s1 then v else arbitrary)\<rightarrow> s2"
   9.845 +apply (case_tac "s", case_tac "a = None")
   9.846 +by (auto intro!: eval.Cond)
   9.847 +
   9.848 +lemma IfI: "\<lbrakk>G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s1; G\<turnstile>s1 \<midarrow>(if the_Bool v then c1 else c2)\<rightarrow> s2\<rbrakk>
   9.849 +                 \<Longrightarrow> G\<turnstile>s \<midarrow>If(e) c1 Else c2\<rightarrow> s2"
   9.850 +apply (case_tac "s", case_tac "a = None")
   9.851 +by (auto intro!: eval.If)
   9.852 +
   9.853 +lemma MethdI: "G\<turnstile>s \<midarrow>body G C sig-\<succ>v\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>Methd C sig-\<succ>v\<rightarrow> s'"
   9.854 +apply (case_tac "s", case_tac "a = None")
   9.855 +by (auto intro!: eval.Methd)
   9.856 +
   9.857 +lemma eval_Call: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1; G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>pvs\<rightarrow> s2;  
   9.858 +       D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>; 
   9.859 +       G\<turnstile>init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' pvs s2 
   9.860 +        \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> v\<rightarrow> s3; 
   9.861 +       s3' = restore_lvars s2 s3\<rbrakk> \<Longrightarrow>  
   9.862 +       G\<turnstile>Norm s0 \<midarrow>{statT,mode}e\<cdot>mn({pTs}ps)-\<succ>v\<rightarrow> s3'"
   9.863 +apply (drule eval.Call, assumption)
   9.864 +apply (rule HOL.refl)
   9.865 +apply simp+
   9.866 +done
   9.867 +
   9.868 +lemma eval_Init: 
   9.869 +"\<lbrakk>if inited C (globs s0) then s3 = Norm s0 
   9.870 +  else G\<turnstile>Norm (init_class_obj G C s0)  
   9.871 +         \<midarrow>(if C = Object then Skip else Init (super (the (class G C))))\<rightarrow> s1 \<and>
   9.872 +       G\<turnstile>set_lvars empty s1 \<midarrow>(init (the (class G C)))\<rightarrow> s2 \<and> 
   9.873 +      s3 = restore_lvars s1 s2\<rbrakk> \<Longrightarrow>  
   9.874 +  G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s3"
   9.875 +apply (rule eval.Init)
   9.876 +apply auto
   9.877 +done
   9.878 +
   9.879 +lemma init_done: "initd C s \<Longrightarrow> G\<turnstile>s \<midarrow>Init C\<rightarrow> s"
   9.880 +apply (case_tac "s", simp)
   9.881 +apply (case_tac "a")
   9.882 +apply  safe
   9.883 +apply (rule eval_Init)
   9.884 +apply   auto
   9.885 +done
   9.886 +
   9.887 +lemma eval_StatRef: 
   9.888 +"G\<turnstile>s \<midarrow>StatRef rt-\<succ>(if abrupt s=None then Null else arbitrary)\<rightarrow> s"
   9.889 +apply (case_tac "s", simp)
   9.890 +apply (case_tac "a = None")
   9.891 +apply (auto del: eval.Abrupt intro!: eval.intros)
   9.892 +done
   9.893 +
   9.894 +
   9.895 +lemma SkipD [dest!]: "G\<turnstile>s \<midarrow>Skip\<rightarrow> s' \<Longrightarrow> s' = s" 
   9.896 +apply (erule eval_cases)
   9.897 +by auto
   9.898 +
   9.899 +lemma Skip_eq [simp]: "G\<turnstile>s \<midarrow>Skip\<rightarrow> s' = (s = s')"
   9.900 +by auto
   9.901 +
   9.902 +(*unused*)
   9.903 +lemma init_retains_locals [rule_format (no_asm)]: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<Longrightarrow>  
   9.904 +  (\<forall>C. t=In1r (Init C) \<longrightarrow> locals (store s) = locals (store s'))"
   9.905 +apply (erule eval.induct)
   9.906 +apply (simp (no_asm_use) split del: split_if_asm option.split_asm)+
   9.907 +apply auto
   9.908 +done
   9.909 +
   9.910 +lemma halloc_xcpt [dest!]: 
   9.911 +  "\<And>s'. G\<turnstile>(Some xc,s) \<midarrow>halloc oi\<succ>a\<rightarrow> s' \<Longrightarrow> s'=(Some xc,s)"
   9.912 +apply (erule_tac halloc_elim_cases)
   9.913 +by auto
   9.914 +
   9.915 +(*
   9.916 +G\<turnstile>(x,(h,l)) \<midarrow>e\<succ>v\<rightarrow> (x',(h',l'))) \<Longrightarrow> l This = l' This"
   9.917 +G\<turnstile>(x,(h,l)) \<midarrow>s  \<rightarrow> (x',(h',l'))) \<Longrightarrow> l This = l' This"
   9.918 +*)
   9.919 +
   9.920 +lemma eval_Methd: 
   9.921 +  "G\<turnstile>s \<midarrow>In1l(body G C sig)\<succ>\<rightarrow> (w,s') \<Longrightarrow> G\<turnstile>s \<midarrow>In1l(Methd C sig)\<succ>\<rightarrow> (w,s')"
   9.922 +apply (case_tac "s")
   9.923 +apply (case_tac "a")
   9.924 +apply clarsimp+
   9.925 +apply (erule eval.Methd)
   9.926 +apply (drule eval_abrupt_lemma)
   9.927 +apply force
   9.928 +done
   9.929 +
   9.930 +
   9.931 +section "single valued"
   9.932 +
   9.933 +lemma unique_halloc [rule_format (no_asm)]: 
   9.934 +  "\<And>s as as'. (s,oi,as)\<in>halloc G \<Longrightarrow> (s,oi,as')\<in>halloc G \<longrightarrow> as'=as"
   9.935 +apply (simp (no_asm_simp) only: split_tupled_all)
   9.936 +apply (erule halloc.induct)
   9.937 +apply  (auto elim!: halloc_elim_cases split del: split_if split_if_asm)
   9.938 +apply (drule trans [THEN sym], erule sym) 
   9.939 +defer
   9.940 +apply (drule trans [THEN sym], erule sym)
   9.941 +apply auto
   9.942 +done
   9.943 +
   9.944 +
   9.945 +lemma single_valued_halloc: 
   9.946 +  "single_valued {((s,oi),(a,s')). G\<turnstile>s \<midarrow>halloc oi\<succ>a \<rightarrow> s'}"
   9.947 +apply (unfold single_valued_def)
   9.948 +by (clarsimp, drule (1) unique_halloc, auto)
   9.949 +
   9.950 +
   9.951 +lemma unique_sxalloc [rule_format (no_asm)]: 
   9.952 +  "\<And>s s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'' \<longrightarrow> s'' = s'"
   9.953 +apply (simp (no_asm_simp) only: split_tupled_all)
   9.954 +apply (erule sxalloc.induct)
   9.955 +apply   (auto dest: unique_halloc elim!: sxalloc_elim_cases 
   9.956 +              split del: split_if split_if_asm)
   9.957 +done
   9.958 +
   9.959 +lemma single_valued_sxalloc: "single_valued {(s,s'). G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'}"
   9.960 +apply (unfold single_valued_def)
   9.961 +apply (blast dest: unique_sxalloc)
   9.962 +done
   9.963 +
   9.964 +lemma split_pairD: "(x,y) = p \<Longrightarrow> x = fst p & y = snd p"
   9.965 +by auto
   9.966 +
   9.967 +lemma unique_eval [rule_format (no_asm)]: 
   9.968 +  "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws \<Longrightarrow> (\<forall>ws'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws' \<longrightarrow> ws' = ws)"
   9.969 +apply (case_tac "ws")
   9.970 +apply (simp only:)
   9.971 +apply (erule thin_rl)
   9.972 +apply (erule eval_induct)
   9.973 +apply (tactic {* ALLGOALS (EVERY'
   9.974 +      [strip_tac, rotate_tac ~1, eresolve_tac (thms "eval_elim_cases")]) *})
   9.975 +(* 29 subgoals *)
   9.976 +prefer 26 (* Try *) 
   9.977 +apply (simp (no_asm_use) only: split add: split_if_asm)
   9.978 +(* 32 subgoals *)
   9.979 +prefer 28 (* Init *)
   9.980 +apply (case_tac "inited C (globs s0)", (simp only: if_True if_False)+)
   9.981 +prefer 24 (* While *)
   9.982 +apply (simp (no_asm_use) only: split add: split_if_asm, blast)
   9.983 +apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp)
   9.984 +apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp)
   9.985 +apply blast
   9.986 +(* 31 subgoals *)
   9.987 +apply (blast dest: unique_sxalloc unique_halloc split_pairD)+
   9.988 +done
   9.989 +
   9.990 +(* unused *)
   9.991 +lemma single_valued_eval: 
   9.992 + "single_valued {((s,t),vs'). G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> vs'}"
   9.993 +apply (unfold single_valued_def)
   9.994 +by (clarify, drule (1) unique_eval, auto)
   9.995 +
   9.996 +end
    10.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    10.2 +++ b/src/HOL/Bali/Evaln.thy	Mon Jan 28 17:00:19 2002 +0100
    10.3 @@ -0,0 +1,373 @@
    10.4 +(*  Title:      isabelle/Bali/Evaln.thy
    10.5 +    ID:         $Id$
    10.6 +    Author:     David von Oheimb
    10.7 +    Copyright   1999 Technische Universitaet Muenchen
    10.8 +*)
    10.9 +header {* Operational evaluation (big-step) semantics of Java expressions and 
   10.10 +          statements
   10.11 +*}
   10.12 +
   10.13 +theory Evaln = Eval:
   10.14 +
   10.15 +text {*
   10.16 +Variant of eval relation with counter for bounded recursive depth
   10.17 +Evaln could completely replace Eval.
   10.18 +*}
   10.19 +
   10.20 +consts
   10.21 +
   10.22 +  evaln	:: "prog \<Rightarrow> (state \<times> term \<times> nat \<times> vals \<times> state) set"
   10.23 +
   10.24 +syntax
   10.25 +
   10.26 +  evaln	:: "[prog, state, term,        nat, vals * state] => bool"
   10.27 +				("_|-_ -_>-_-> _"   [61,61,80,   61,61] 60)
   10.28 +  evarn	:: "[prog, state, var  , vvar        , nat, state] => bool"
   10.29 +				("_|-_ -_=>_-_-> _" [61,61,90,61,61,61] 60)
   10.30 +  eval_n:: "[prog, state, expr , val         , nat, state] => bool"
   10.31 +				("_|-_ -_->_-_-> _" [61,61,80,61,61,61] 60)
   10.32 +  evalsn:: "[prog, state, expr list, val list, nat, state] => bool"
   10.33 +				("_|-_ -_#>_-_-> _" [61,61,61,61,61,61] 60)
   10.34 +  execn	:: "[prog, state, stmt ,               nat, state] => bool"
   10.35 +				("_|-_ -_-_-> _"    [61,61,65,   61,61] 60)
   10.36 +
   10.37 +syntax (xsymbols)
   10.38 +
   10.39 +  evaln	:: "[prog, state, term,         nat, vals \<times> state] \<Rightarrow> bool"
   10.40 +				("_\<turnstile>_ \<midarrow>_\<succ>\<midarrow>_\<rightarrow> _"   [61,61,80,   61,61] 60)
   10.41 +  evarn	:: "[prog, state, var  , vvar         , nat, state] \<Rightarrow> bool"
   10.42 +				("_\<turnstile>_ \<midarrow>_=\<succ>_\<midarrow>_\<rightarrow> _" [61,61,90,61,61,61] 60)
   10.43 +  eval_n:: "[prog, state, expr , val ,          nat, state] \<Rightarrow> bool"
   10.44 +				("_\<turnstile>_ \<midarrow>_-\<succ>_\<midarrow>_\<rightarrow> _" [61,61,80,61,61,61] 60)
   10.45 +  evalsn:: "[prog, state, expr list, val  list, nat, state] \<Rightarrow> bool"
   10.46 +				("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<midarrow>_\<rightarrow> _" [61,61,61,61,61,61] 60)
   10.47 +  execn	:: "[prog, state, stmt ,                nat, state] \<Rightarrow> bool"
   10.48 +				("_\<turnstile>_ \<midarrow>_\<midarrow>_\<rightarrow> _"     [61,61,65,   61,61] 60)
   10.49 +
   10.50 +translations
   10.51 +
   10.52 +  "G\<turnstile>s \<midarrow>t    \<succ>\<midarrow>n\<rightarrow>  w___s' " == "(s,t,n,w___s') \<in> evaln G"
   10.53 +  "G\<turnstile>s \<midarrow>t    \<succ>\<midarrow>n\<rightarrow> (w,  s')" <= "(s,t,n,w,  s') \<in> evaln G"
   10.54 +  "G\<turnstile>s \<midarrow>t    \<succ>\<midarrow>n\<rightarrow> (w,x,s')" <= "(s,t,n,w,x,s') \<in> evaln G"
   10.55 +  "G\<turnstile>s \<midarrow>c     \<midarrow>n\<rightarrow> (x,s')" <= "G\<turnstile>s \<midarrow>In1r  c\<succ>\<midarrow>n\<rightarrow> (\<diamondsuit>    ,x,s')"
   10.56 +  "G\<turnstile>s \<midarrow>c     \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In1r  c\<succ>\<midarrow>n\<rightarrow> (\<diamondsuit>    ,  s')"
   10.57 +  "G\<turnstile>s \<midarrow>e-\<succ>v  \<midarrow>n\<rightarrow> (x,s')" <= "G\<turnstile>s \<midarrow>In1l e\<succ>\<midarrow>n\<rightarrow> (In1 v ,x,s')"
   10.58 +  "G\<turnstile>s \<midarrow>e-\<succ>v  \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In1l e\<succ>\<midarrow>n\<rightarrow> (In1 v ,  s')"
   10.59 +  "G\<turnstile>s \<midarrow>e=\<succ>vf \<midarrow>n\<rightarrow> (x,s')" <= "G\<turnstile>s \<midarrow>In2  e\<succ>\<midarrow>n\<rightarrow> (In2 vf,x,s')"
   10.60 +  "G\<turnstile>s \<midarrow>e=\<succ>vf \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In2  e\<succ>\<midarrow>n\<rightarrow> (In2 vf,  s')"
   10.61 +  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v  \<midarrow>n\<rightarrow> (x,s')" <= "G\<turnstile>s \<midarrow>In3  e\<succ>\<midarrow>n\<rightarrow> (In3 v ,x,s')"
   10.62 +  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v  \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In3  e\<succ>\<midarrow>n\<rightarrow> (In3 v ,  s')"
   10.63 +
   10.64 +
   10.65 +inductive "evaln G" intros
   10.66 +
   10.67 +(* propagation of abrupt completion *)
   10.68 +
   10.69 +  Abrupt:   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (arbitrary3 t,(Some xc,s))"
   10.70 +
   10.71 +
   10.72 +(* evaluation of variables *)
   10.73 +
   10.74 +  LVar:	"G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<midarrow>n\<rightarrow> Norm s"
   10.75 +
   10.76 +  FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s2;
   10.77 +	  (v,s2') = fvar C stat fn a' s2\<rbrakk> \<Longrightarrow>
   10.78 +	  G\<turnstile>Norm s0 \<midarrow>{C,stat}e..fn=\<succ>v\<midarrow>n\<rightarrow> s2'"
   10.79 +
   10.80 +  AVar:	"\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1 ; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2; 
   10.81 +	  (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
   10.82 +	              G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<midarrow>n\<rightarrow> s2'"
   10.83 +
   10.84 +
   10.85 +
   10.86 +
   10.87 +(* evaluation of expressions *)
   10.88 +
   10.89 +  NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1;
   10.90 +	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
   10.91 +	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<midarrow>n\<rightarrow> s2"
   10.92 +
   10.93 +  NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<midarrow>n\<rightarrow> s2; 
   10.94 +	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
   10.95 +	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<midarrow>n\<rightarrow> s3"
   10.96 +
   10.97 +  Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
   10.98 +	  s2 = abupd (raise_if (\<not>G,snd s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
   10.99 +			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<midarrow>n\<rightarrow> s2"
  10.100 +
  10.101 +  Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
  10.102 +	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
  10.103 +			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<midarrow>n\<rightarrow> s1"
  10.104 +
  10.105 +  Lit:			   "G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<midarrow>n\<rightarrow> Norm s"
  10.106 +
  10.107 +  Super:		   "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<midarrow>n\<rightarrow> Norm s"
  10.108 +
  10.109 +  Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
  10.110 +	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<midarrow>n\<rightarrow> s1"
  10.111 +
  10.112 +  Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<midarrow>n\<rightarrow> s1;
  10.113 +          G\<turnstile>     s1 \<midarrow>e-\<succ>v     \<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
  10.114 +				   G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v\<midarrow>n\<rightarrow> assign f v s2"
  10.115 +
  10.116 +  Cond:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1;
  10.117 +          G\<turnstile>     s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
  10.118 +			    G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<midarrow>n\<rightarrow> s2"
  10.119 +
  10.120 +  Call:	
  10.121 +  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2;
  10.122 +    D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>; 
  10.123 +    G\<turnstile>init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs s2
  10.124 +            \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<midarrow>n\<rightarrow> s3\<rbrakk>
  10.125 +   \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>{statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<midarrow>n\<rightarrow> (restore_lvars s2 s3)"
  10.126 +
  10.127 +  Methd:"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
  10.128 +				G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<midarrow>Suc n\<rightarrow> s1"
  10.129 +
  10.130 +  Body:	"\<lbrakk>G\<turnstile>Norm s0\<midarrow>Init D\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2\<rbrakk>\<Longrightarrow>
  10.131 + G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)\<midarrow>n\<rightarrow>abupd (absorb Ret) s2"
  10.132 +
  10.133 +(* evaluation of expression lists *)
  10.134 +
  10.135 +  Nil:
  10.136 +				"G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<midarrow>n\<rightarrow> Norm s0"
  10.137 +
  10.138 +  Cons:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<midarrow>n\<rightarrow> s1;
  10.139 +          G\<turnstile>     s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
  10.140 +			     G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<midarrow>n\<rightarrow> s2"
  10.141 +
  10.142 +
  10.143 +(* execution of statements *)
  10.144 +
  10.145 +  Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<midarrow>n\<rightarrow> Norm s"
  10.146 +
  10.147 +  Expr:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
  10.148 +				  G\<turnstile>Norm s0 \<midarrow>Expr e\<midarrow>n\<rightarrow> s1"
  10.149 +
  10.150 +  Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
  10.151 +                             G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<midarrow>n\<rightarrow> abupd (absorb (Break l)) s1"
  10.152 +
  10.153 +  Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1;
  10.154 +	  G\<turnstile>     s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
  10.155 +				 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<midarrow>n\<rightarrow> s2"
  10.156 +
  10.157 +  If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
  10.158 +	  G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
  10.159 +		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<midarrow>n\<rightarrow> s2"
  10.160 +
  10.161 +  Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
  10.162 +	  if normal s1 \<and> the_Bool b 
  10.163 +             then (G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2 \<and> 
  10.164 +                   G\<turnstile>(abupd (absorb (Cont l)) s2) \<midarrow>l\<bullet> While(e) c\<midarrow>n\<rightarrow> s3)
  10.165 +	     else s3 = s1\<rbrakk> \<Longrightarrow>
  10.166 +			      G\<turnstile>Norm s0 \<midarrow>l\<bullet> While(e) c\<midarrow>n\<rightarrow> s3"
  10.167 +  
  10.168 +  Do: "G\<turnstile>Norm s \<midarrow>Do j\<midarrow>n\<rightarrow> (Some (Jump j), s)"
  10.169 +  
  10.170 +  Throw:"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
  10.171 +				 G\<turnstile>Norm s0 \<midarrow>Throw e\<midarrow>n\<rightarrow> abupd (throw a') s1"
  10.172 +
  10.173 +  Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2;
  10.174 +	  if G,s2\<turnstile>catch tn then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3 else s3 = s2\<rbrakk>
  10.175 +          \<Longrightarrow>
  10.176 +		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(tn vn) c2\<midarrow>n\<rightarrow> s3"
  10.177 +
  10.178 +  Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> (x1,s1);
  10.179 +	  G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
  10.180 +              G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<midarrow>n\<rightarrow> abupd (abrupt_if (x1\<noteq>None) x1) s2"
  10.181 +  
  10.182 +  Init:	"\<lbrakk>the (class G C) = c;
  10.183 +	  if inited C (globs s0) then s3 = Norm s0
  10.184 +	  else (G\<turnstile>Norm (init_class_obj G C s0)
  10.185 +	          \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1 \<and>
  10.186 +	        G\<turnstile>set_lvars empty s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2 \<and> 
  10.187 +                s3 = restore_lvars s1 s2)\<rbrakk>
  10.188 +          \<Longrightarrow>
  10.189 +		 G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s3"
  10.190 +monos
  10.191 +  if_def2
  10.192 +
  10.193 +lemma evaln_eval: "\<And>ws. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> ws \<Longrightarrow> G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws"
  10.194 +apply (simp (no_asm_simp) only: split_tupled_all)
  10.195 +apply (erule evaln.induct)
  10.196 +apply (rule eval.intros, (assumption+)?,(force split del: split_if)?)+
  10.197 +done
  10.198 +
  10.199 +
  10.200 +lemma Suc_le_D_lemma: "\<lbrakk>Suc n <= m'; (\<And>m. n <= m \<Longrightarrow> P (Suc m)) \<rbrakk> \<Longrightarrow> P m'"
  10.201 +apply (frule Suc_le_D)
  10.202 +apply fast
  10.203 +done
  10.204 +
  10.205 +lemma evaln_nonstrict [rule_format (no_asm), elim]: 
  10.206 +  "\<And>ws. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> ws \<Longrightarrow> \<forall>m. n\<le>m \<longrightarrow> G\<turnstile>s \<midarrow>t\<succ>\<midarrow>m\<rightarrow> ws"
  10.207 +apply (simp (no_asm_simp) only: split_tupled_all)
  10.208 +apply (erule evaln.induct)
  10.209 +apply (tactic {* ALLGOALS (EVERY'[strip_tac, TRY o etac (thm "Suc_le_D_lemma"),
  10.210 +  REPEAT o smp_tac 1, 
  10.211 +  resolve_tac (thms "evaln.intros") THEN_ALL_NEW TRY o atac]) *})
  10.212 +(* 3 subgoals *)
  10.213 +apply (auto split del: split_if)
  10.214 +done
  10.215 +
  10.216 +lemmas evaln_nonstrict_Suc = evaln_nonstrict [OF _ le_refl [THEN le_SucI]]
  10.217 +
  10.218 +lemma evaln_max2: "\<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>n1\<rightarrow> ws1; G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>n2\<rightarrow> ws2\<rbrakk> \<Longrightarrow> 
  10.219 +             G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max n1 n2\<rightarrow> ws1 \<and> G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max n1 n2\<rightarrow> ws2"
  10.220 +apply (fast intro: le_maxI1 le_maxI2)
  10.221 +done
  10.222 +
  10.223 +lemma evaln_max3: 
  10.224 +"\<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>n1\<rightarrow> ws1; G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>n2\<rightarrow> ws2; G\<turnstile>s3 \<midarrow>t3\<succ>\<midarrow>n3\<rightarrow> ws3\<rbrakk> \<Longrightarrow>
  10.225 + G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws1 \<and>
  10.226 + G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws2 \<and> 
  10.227 + G\<turnstile>s3 \<midarrow>t3\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws3"
  10.228 +apply (drule (1) evaln_max2, erule thin_rl)
  10.229 +apply (fast intro!: le_maxI1 le_maxI2)
  10.230 +done
  10.231 +
  10.232 +lemma eval_evaln: "\<And>ws. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws \<Longrightarrow> (\<exists>n. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> ws)"
  10.233 +apply (simp (no_asm_simp) only: split_tupled_all)
  10.234 +apply (erule eval.induct)
  10.235 +apply (tactic {* ALLGOALS 
  10.236 +         (asm_full_simp_tac (HOL_basic_ss addsplits [split_if_asm])) *})
  10.237 +apply (tactic {* ALLGOALS (EVERY'[
  10.238 +   REPEAT o eresolve_tac [exE, conjE], rtac exI,
  10.239 +                     TRY o datac (thm "evaln_max3") 2, REPEAT o etac conjE,
  10.240 +  resolve_tac (thms "evaln.intros") THEN_ALL_NEW 
  10.241 +  force_tac (HOL_cs, HOL_ss)]) *})
  10.242 +done
  10.243 +
  10.244 +declare split_if     [split del] split_if_asm     [split del]
  10.245 +        option.split [split del] option.split_asm [split del]
  10.246 +inductive_cases evaln_cases: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> vs'"
  10.247 +
  10.248 +inductive_cases evaln_elim_cases:
  10.249 +	"G\<turnstile>(Some xc, s) \<midarrow>t                        \<succ>\<midarrow>n\<rightarrow> vs'"
  10.250 +	"G\<turnstile>Norm s \<midarrow>In1r Skip                      \<succ>\<midarrow>n\<rightarrow> xs'"
  10.251 +        "G\<turnstile>Norm s \<midarrow>In1r (Do j)                    \<succ>\<midarrow>n\<rightarrow> xs'"
  10.252 +        "G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> c)                    \<succ>\<midarrow>n\<rightarrow> xs'"
  10.253 +	"G\<turnstile>Norm s \<midarrow>In3  ([])                      \<succ>\<midarrow>n\<rightarrow> vs'"
  10.254 +	"G\<turnstile>Norm s \<midarrow>In3  (e#es)                    \<succ>\<midarrow>n\<rightarrow> vs'"
  10.255 +	"G\<turnstile>Norm s \<midarrow>In1l (Lit w)                   \<succ>\<midarrow>n\<rightarrow> vs'"
  10.256 +	"G\<turnstile>Norm s \<midarrow>In2  (LVar vn)                 \<succ>\<midarrow>n\<rightarrow> vs'"
  10.257 +	"G\<turnstile>Norm s \<midarrow>In1l (Cast T e)                \<succ>\<midarrow>n\<rightarrow> vs'"
  10.258 +	"G\<turnstile>Norm s \<midarrow>In1l (e InstOf T)              \<succ>\<midarrow>n\<rightarrow> vs'"
  10.259 +	"G\<turnstile>Norm s \<midarrow>In1l (Super)                   \<succ>\<midarrow>n\<rightarrow> vs'"
  10.260 +	"G\<turnstile>Norm s \<midarrow>In1l (Acc va)                  \<succ>\<midarrow>n\<rightarrow> vs'"
  10.261 +	"G\<turnstile>Norm s \<midarrow>In1r (Expr e)                  \<succ>\<midarrow>n\<rightarrow> xs'"
  10.262 +	"G\<turnstile>Norm s \<midarrow>In1r (c1;; c2)                 \<succ>\<midarrow>n\<rightarrow> xs'"
  10.263 +	"G\<turnstile>Norm s \<midarrow>In1l (Methd C sig)             \<succ>\<midarrow>n\<rightarrow> xs'"
  10.264 +	"G\<turnstile>Norm s \<midarrow>In1l (Body D c)                \<succ>\<midarrow>n\<rightarrow> xs'"
  10.265 +	"G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2)            \<succ>\<midarrow>n\<rightarrow> vs'"
  10.266 +	"G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2)        \<succ>\<midarrow>n\<rightarrow> xs'"
  10.267 +	"G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> While(e) c)           \<succ>\<midarrow>n\<rightarrow> xs'"
  10.268 +	"G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2)           \<succ>\<midarrow>n\<rightarrow> xs'"
  10.269 +	"G\<turnstile>Norm s \<midarrow>In1r (Throw e)                 \<succ>\<midarrow>n\<rightarrow> xs'"
  10.270 +	"G\<turnstile>Norm s \<midarrow>In1l (NewC C)                  \<succ>\<midarrow>n\<rightarrow> vs'"
  10.271 +	"G\<turnstile>Norm s \<midarrow>In1l (New T[e])                \<succ>\<midarrow>n\<rightarrow> vs'"
  10.272 +	"G\<turnstile>Norm s \<midarrow>In1l (Ass va e)                \<succ>\<midarrow>n\<rightarrow> vs'"
  10.273 +	"G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2)  \<succ>\<midarrow>n\<rightarrow> xs'"
  10.274 +	"G\<turnstile>Norm s \<midarrow>In2  ({C,stat}e..fn)           \<succ>\<midarrow>n\<rightarrow> vs'"
  10.275 +	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                 \<succ>\<midarrow>n\<rightarrow> vs'"
  10.276 +	"G\<turnstile>Norm s \<midarrow>In1l ({statT,mode}e\<cdot>mn({pT}p)) \<succ>\<midarrow>n\<rightarrow> vs'"
  10.277 +	"G\<turnstile>Norm s \<midarrow>In1r (Init C)                  \<succ>\<midarrow>n\<rightarrow> xs'"
  10.278 +declare split_if     [split] split_if_asm     [split] 
  10.279 +        option.split [split] option.split_asm [split]
  10.280 +
  10.281 +lemma evaln_Inj_elim: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (w,s') \<Longrightarrow> case t of In1 ec \<Rightarrow>  
  10.282 +  (case ec of Inl e \<Rightarrow> (\<exists>v. w = In1 v) | Inr c \<Rightarrow> w = \<diamondsuit>)  
  10.283 +  | In2 e \<Rightarrow> (\<exists>v. w = In2 v) | In3 e \<Rightarrow> (\<exists>v. w = In3 v)"
  10.284 +apply (erule evaln_cases , auto)
  10.285 +apply (induct_tac "t")
  10.286 +apply   (induct_tac "a")
  10.287 +apply auto
  10.288 +done
  10.289 +
  10.290 +ML_setup {*
  10.291 +fun enf nam inj rhs =
  10.292 +let
  10.293 +  val name = "evaln_" ^ nam ^ "_eq"
  10.294 +  val lhs = "G\<turnstile>s \<midarrow>" ^ inj ^ " t\<succ>\<midarrow>n\<rightarrow> (w, s')"
  10.295 +  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
  10.296 +	(K [Auto_tac, ALLGOALS (ftac (thm "evaln_Inj_elim")) THEN Auto_tac])
  10.297 +  fun is_Inj (Const (inj,_) $ _) = true
  10.298 +    | is_Inj _                   = false
  10.299 +  fun pred (_ $ (Const ("Pair",_) $ _ $ (Const ("Pair", _) $ _ $ 
  10.300 +    (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ x $ _ )))) $ _ ) = is_Inj x
  10.301 +in
  10.302 +  make_simproc name lhs pred (thm name)
  10.303 +end;
  10.304 +
  10.305 +val evaln_expr_proc = enf "expr" "In1l" "\<exists>v.  w=In1 v  \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<midarrow>n\<rightarrow> s'";
  10.306 +val evaln_var_proc  = enf "var"  "In2"  "\<exists>vf. w=In2 vf \<and> G\<turnstile>s \<midarrow>t=\<succ>vf\<midarrow>n\<rightarrow> s'";
  10.307 +val evaln_exprs_proc= enf "exprs""In3"  "\<exists>vs. w=In3 vs \<and> G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s'";
  10.308 +val evaln_stmt_proc = enf "stmt" "In1r" "     w=\<diamondsuit>      \<and> G\<turnstile>s \<midarrow>t     \<midarrow>n\<rightarrow> s'";
  10.309 +Addsimprocs [evaln_expr_proc,evaln_var_proc,evaln_exprs_proc,evaln_stmt_proc];
  10.310 +
  10.311 +bind_thms ("evaln_AbruptIs", sum3_instantiate (thm "evaln.Abrupt"))
  10.312 +*}
  10.313 +declare evaln_AbruptIs [intro!]
  10.314 +
  10.315 +lemma evaln_abrupt_lemma: "G\<turnstile>s \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (v,s') \<Longrightarrow> 
  10.316 + fst s = Some xc \<longrightarrow> s' = s \<and> v = arbitrary3 e"
  10.317 +apply (erule evaln_cases , auto)
  10.318 +done
  10.319 +
  10.320 +lemma evaln_abrupt: 
  10.321 + "\<And>s'. G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (w,s') = (s' = (Some xc,s) \<and>  
  10.322 +  w=arbitrary3 e \<and> G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (arbitrary3 e,(Some xc,s)))"
  10.323 +apply auto
  10.324 +apply (frule evaln_abrupt_lemma, auto)+
  10.325 +done
  10.326 +
  10.327 +ML {*
  10.328 +local
  10.329 +  fun is_Some (Const ("Pair",_) $ (Const ("Option.option.Some",_) $ _)$ _) =true
  10.330 +    | is_Some _ = false
  10.331 +  fun pred (_ $ (Const ("Pair",_) $
  10.332 +     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
  10.333 +       (Const ("Pair", _) $ _ $ x)))) $ _ ) = is_Some x
  10.334 +in
  10.335 +  val evaln_abrupt_proc = 
  10.336 + make_simproc "evaln_abrupt" "G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (w,s')" pred (thm "evaln_abrupt")
  10.337 +end;
  10.338 +Addsimprocs [evaln_abrupt_proc]
  10.339 +*}
  10.340 +
  10.341 +lemma evaln_LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else arbitrary)\<midarrow>n\<rightarrow> s"
  10.342 +apply (case_tac "s", case_tac "a = None")
  10.343 +by (auto intro!: evaln.Lit)
  10.344 +
  10.345 +lemma CondI: 
  10.346 + "\<And>s1. \<lbrakk>G\<turnstile>s \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow> 
  10.347 +  G\<turnstile>s \<midarrow>e ? e1 : e2-\<succ>(if normal s1 then v else arbitrary)\<midarrow>n\<rightarrow> s2"
  10.348 +apply (case_tac "s", case_tac "a = None")
  10.349 +by (auto intro!: evaln.Cond)
  10.350 +
  10.351 +lemma evaln_SkipI [intro!]: "G\<turnstile>s \<midarrow>Skip\<midarrow>n\<rightarrow> s"
  10.352 +apply (case_tac "s", case_tac "a = None")
  10.353 +by (auto intro!: evaln.Skip)
  10.354 +
  10.355 +lemma evaln_ExprI: "G\<turnstile>s \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>Expr e\<midarrow>n\<rightarrow> s'"
  10.356 +apply (case_tac "s", case_tac "a = None")
  10.357 +by (auto intro!: evaln.Expr)
  10.358 +
  10.359 +lemma evaln_CompI: "\<lbrakk>G\<turnstile>s \<midarrow>c1\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c2\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow> G\<turnstile>s \<midarrow>c1;; c2\<midarrow>n\<rightarrow> s2"
  10.360 +apply (case_tac "s", case_tac "a = None")
  10.361 +by (auto intro!: evaln.Comp)
  10.362 +
  10.363 +lemma evaln_IfI: 
  10.364 + "\<lbrakk>G\<turnstile>s \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>(if the_Bool v then c1 else c2)\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow> 
  10.365 +  G\<turnstile>s \<midarrow>If(e) c1 Else c2\<midarrow>n\<rightarrow> s2"
  10.366 +apply (case_tac "s", case_tac "a = None")
  10.367 +by (auto intro!: evaln.If)
  10.368 +
  10.369 +lemma evaln_SkipD [dest!]: "G\<turnstile>s \<midarrow>Skip\<midarrow>n\<rightarrow> s' \<Longrightarrow> s' = s" 
  10.370 +by (erule evaln_cases, auto)
  10.371 +
  10.372 +lemma evaln_Skip_eq [simp]: "G\<turnstile>s \<midarrow>Skip\<midarrow>n\<rightarrow> s' = (s = s')"
  10.373 +apply auto
  10.374 +done
  10.375 +
  10.376 +end
    11.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    11.2 +++ b/src/HOL/Bali/Example.thy	Mon Jan 28 17:00:19 2002 +0100
    11.3 @@ -0,0 +1,1248 @@
    11.4 +(*  Title:      isabelle/Bali/Example.thy
    11.5 +    ID:         $Id$
    11.6 +    Author:     David von Oheimb
    11.7 +    Copyright   1997 Technische Universitaet Muenchen
    11.8 +*)
    11.9 +header {* Example Bali program *}
   11.10 +
   11.11 +theory Example = Eval + WellForm:
   11.12 +
   11.13 +
   11.14 +text {*
   11.15 +The following example Bali program includes:
   11.16 +\begin{itemize}
   11.17 +\item class and interface declarations with inheritance, hiding of fields,
   11.18 +    overriding of methods (with refined result type), array type,
   11.19 +\item method call (with dynamic binding), parameter access, return expressions,
   11.20 +\item expression statements, sequential composition, literal values, 
   11.21 +    local assignment, local access, field assignment, type cast,
   11.22 +\item exception generation and propagation, try and catch statement, throw 
   11.23 +      statement
   11.24 +\item instance creation and (default) static initialization
   11.25 +\end{itemize}
   11.26 +\begin{verbatim}
   11.27 +package java_lang
   11.28 +
   11.29 +public interface HasFoo {
   11.30 +  public Base foo(Base z);
   11.31 +}
   11.32 +
   11.33 +public class Base implements HasFoo {
   11.34 +  static boolean arr[] = new boolean[2];
   11.35 +  public HasFoo vee;
   11.36 +  public Base foo(Base z) {
   11.37 +    return z;
   11.38 +  }
   11.39 +}
   11.40 +
   11.41 +public class Ext extends Base {
   11.42 +  public int vee;
   11.43 +  public Ext foo(Base z) {
   11.44 +    ((Ext)z).vee = 1;
   11.45 +    return null;
   11.46 +  }
   11.47 +}
   11.48 +
   11.49 +public class Example {
   11.50 +  public static void main(String args[]) throws Throwable {
   11.51 +    Base e = new Ext();
   11.52 +    try {e.foo(null); }
   11.53 +    catch(NullPointerException z) { 
   11.54 +      while(Ext.arr[2]) ;
   11.55 +    }
   11.56 +  }
   11.57 +}
   11.58 +\end{verbatim}
   11.59 +*}
   11.60 +
   11.61 +declare widen.null [intro]
   11.62 +
   11.63 +lemma wf_fdecl_def2: "\<And>fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))"
   11.64 +apply (unfold wf_fdecl_def)
   11.65 +apply (simp (no_asm))
   11.66 +done
   11.67 +
   11.68 +declare wf_fdecl_def2 [iff]
   11.69 +
   11.70 +
   11.71 +section "type and expression names"
   11.72 +
   11.73 +(** unfortunately cannot simply instantiate tnam **)
   11.74 +datatype tnam_  = HasFoo_ | Base_ | Ext_
   11.75 +datatype vnam_  = arr_ | vee_ | z_ | e_
   11.76 +datatype label_ = lab1_
   11.77 +
   11.78 +consts
   11.79 +
   11.80 +  tnam_ :: "tnam_  \<Rightarrow> tnam"
   11.81 +  vnam_ :: "vnam_  \<Rightarrow> vname"
   11.82 +  label_:: "label_ \<Rightarrow> label"
   11.83 +axioms  (** tnam_, vnam_ and label are intended to be isomorphic 
   11.84 +            to tnam, vname and label **)
   11.85 +
   11.86 +  inj_tnam_  [simp]: "(tnam_  x = tnam_  y) = (x = y)"
   11.87 +  inj_vnam_  [simp]: "(vnam_  x = vnam_  y) = (x = y)"
   11.88 +  inj_label_ [simp]: "(label_ x = label_ y) = (x = y)"
   11.89 +  
   11.90 +  
   11.91 +  surj_tnam_:  "\<exists>m. n = tnam_  m"
   11.92 +  surj_vnam_:  "\<exists>m. n = vnam_  m"
   11.93 +  surj_label_:" \<exists>m. n = label_ m"
   11.94 +
   11.95 +syntax
   11.96 +
   11.97 +  HasFoo :: qtname
   11.98 +  Base   :: qtname
   11.99 +  Ext    :: qtname
  11.100 +  arr :: ename
  11.101 +  vee :: ename
  11.102 +  z   :: ename
  11.103 +  e   :: ename
  11.104 +  lab1:: label
  11.105 +translations
  11.106 +
  11.107 +  "HasFoo" == "\<lparr>pid=java_lang,tid=TName (tnam_ HasFoo_)\<rparr>"
  11.108 +  "Base"   == "\<lparr>pid=java_lang,tid=TName (tnam_ Base_)\<rparr>"
  11.109 +  "Ext"    == "\<lparr>pid=java_lang,tid=TName (tnam_ Ext_)\<rparr>"
  11.110 +  "arr"    ==        "(vnam_ arr_)"
  11.111 +  "vee"    ==        "(vnam_ vee_)"
  11.112 +  "z"      ==        "(vnam_ z_)"
  11.113 +  "e"      ==        "(vnam_ e_)"
  11.114 +  "lab1"   ==        "label_ lab1_"
  11.115 +
  11.116 +
  11.117 +lemma neq_Base_Object [simp]: "Base\<noteq>Object"
  11.118 +by (simp add: Object_def)
  11.119 +
  11.120 +lemma neq_Ext_Object [simp]: "Ext\<noteq>Object"
  11.121 +by (simp add: Object_def)
  11.122 +
  11.123 +lemma neq_Base_SXcpt [simp]: "Base\<noteq>SXcpt xn"
  11.124 +by (simp add: SXcpt_def)
  11.125 +
  11.126 +lemma neq_Ext_SXcpt [simp]: "Ext\<noteq>SXcpt xn"
  11.127 +by (simp add: SXcpt_def)
  11.128 +
  11.129 +section "classes and interfaces"
  11.130 +
  11.131 +defs
  11.132 +
  11.133 +  Object_mdecls_def: "Object_mdecls \<equiv> []"
  11.134 +  SXcpt_mdecls_def:  "SXcpt_mdecls  \<equiv> []"
  11.135 +
  11.136 +consts
  11.137 +  
  11.138 +  foo    :: mname
  11.139 +
  11.140 +constdefs
  11.141 +  
  11.142 +  foo_sig   :: sig
  11.143 + "foo_sig   \<equiv> \<lparr>name=foo,parTs=[Class Base]\<rparr>"
  11.144 +  
  11.145 +  foo_mhead :: mhead
  11.146 + "foo_mhead \<equiv> \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>"
  11.147 +
  11.148 +constdefs
  11.149 +  
  11.150 +  Base_foo :: mdecl
  11.151 + "Base_foo \<equiv> (foo_sig, \<lparr>access=Public,static=False,pars=[z],resT=Class Base,
  11.152 +                        mbody=\<lparr>lcls=[],stmt=Return (!!z)\<rparr>\<rparr>)"
  11.153 +  
  11.154 +  Ext_foo  :: mdecl
  11.155 + "Ext_foo  \<equiv> (foo_sig, 
  11.156 +              \<lparr>access=Public,static=False,pars=[z],resT=Class Ext,
  11.157 +	       mbody=\<lparr>lcls=[]
  11.158 +                     ,stmt=Expr({Ext,False}Cast (Class Ext) (!!z)..vee := 
  11.159 +       	                                                     Lit (Intg 1))\<rparr>
  11.160 +	      \<rparr>)"
  11.161 +
  11.162 +constdefs
  11.163 +  
  11.164 +arr_viewed_from :: "qtname \<Rightarrow> var"
  11.165 +"arr_viewed_from C \<equiv> {Base,True}StatRef (ClassT C)..arr"
  11.166 +
  11.167 +BaseCl :: class
  11.168 +"BaseCl \<equiv> \<lparr>access=Public,
  11.169 +           cfields=[(arr, \<lparr>access=Public,static=True ,type=PrimT Boolean.[]\<rparr>),
  11.170 +	            (vee, \<lparr>access=Public,static=False,type=Iface HasFoo    \<rparr>)],
  11.171 +           methods=[Base_foo],
  11.172 +           init=Expr(arr_viewed_from Base :=New (PrimT Boolean)[Lit (Intg 2)]),
  11.173 +           super=Object,
  11.174 +           superIfs=[HasFoo]\<rparr>"
  11.175 +  
  11.176 +ExtCl  :: class
  11.177 +"ExtCl  \<equiv> \<lparr>access=Public,
  11.178 +           cfields=[(vee, \<lparr>access=Public,static=False,type= PrimT Integer\<rparr>)], 
  11.179 +           methods=[Ext_foo],
  11.180 +           init=Skip,
  11.181 +           super=Base,
  11.182 +           superIfs=[]\<rparr>"
  11.183 +
  11.184 +constdefs
  11.185 +  
  11.186 +  HasFooInt :: iface
  11.187 + "HasFooInt \<equiv> \<lparr>access=Public,imethods=[(foo_sig, foo_mhead)],isuperIfs=[]\<rparr>"
  11.188 +
  11.189 +  Ifaces ::"idecl list"
  11.190 + "Ifaces \<equiv> [(HasFoo,HasFooInt)]"
  11.191 +
  11.192 + "Classes" ::"cdecl list"
  11.193 + "Classes \<equiv> [(Base,BaseCl),(Ext,ExtCl)]@standard_classes"
  11.194 +
  11.195 +lemmas table_classes_defs = 
  11.196 +     Classes_def standard_classes_def ObjectC_def SXcptC_def
  11.197 +
  11.198 +lemma table_ifaces [simp]: "table_of Ifaces = empty(HasFoo\<mapsto>HasFooInt)"
  11.199 +apply (unfold Ifaces_def)
  11.200 +apply (simp (no_asm))
  11.201 +done
  11.202 +
  11.203 +lemma table_classes_Object [simp]: 
  11.204 + "table_of Classes Object = Some \<lparr>access=Public,cfields=[]
  11.205 +                                 ,methods=Object_mdecls
  11.206 +                                 ,init=Skip,super=arbitrary,superIfs=[]\<rparr>"
  11.207 +apply (unfold table_classes_defs)
  11.208 +apply (simp (no_asm) add:Object_def)
  11.209 +done
  11.210 +
  11.211 +lemma table_classes_SXcpt [simp]: 
  11.212 +  "table_of Classes (SXcpt xn) 
  11.213 +    = Some \<lparr>access=Public,cfields=[],methods=SXcpt_mdecls,
  11.214 +            init=Skip,
  11.215 +            super=if xn = Throwable then Object else SXcpt Throwable,
  11.216 +            superIfs=[]\<rparr>"
  11.217 +apply (unfold table_classes_defs)
  11.218 +apply (induct_tac xn)
  11.219 +apply (simp add: Object_def SXcpt_def)+
  11.220 +done
  11.221 +
  11.222 +lemma table_classes_HasFoo [simp]: "table_of Classes HasFoo = None"
  11.223 +apply (unfold table_classes_defs)
  11.224 +apply (simp (no_asm) add: Object_def SXcpt_def)
  11.225 +done
  11.226 +
  11.227 +lemma table_classes_Base [simp]: "table_of Classes Base = Some BaseCl"
  11.228 +apply (unfold table_classes_defs )
  11.229 +apply (simp (no_asm) add: Object_def SXcpt_def)
  11.230 +done
  11.231 +
  11.232 +lemma table_classes_Ext [simp]: "table_of Classes Ext = Some ExtCl"
  11.233 +apply (unfold table_classes_defs )
  11.234 +apply (simp (no_asm) add: Object_def SXcpt_def)
  11.235 +done
  11.236 +
  11.237 +
  11.238 +section "program"
  11.239 +
  11.240 +syntax
  11.241 +  tprg :: prog
  11.242 +
  11.243 +translations
  11.244 +  "tprg" == "\<lparr>ifaces=Ifaces,classes=Classes\<rparr>"
  11.245 +
  11.246 +constdefs
  11.247 +  test    :: "(ty)list \<Rightarrow> stmt"
  11.248 + "test pTs \<equiv> e:==NewC Ext;; 
  11.249 +           \<spacespace> Try Expr({ClassT Base,IntVir}!!e\<cdot>foo({pTs}[Lit Null]))
  11.250 +           \<spacespace> Catch((SXcpt NullPointer) z)
  11.251 +           (lab1\<bullet> While(Acc (Acc (arr_viewed_from Ext).[Lit (Intg 2)])) Skip)"
  11.252 +
  11.253 +
  11.254 +section "well-structuredness"
  11.255 +
  11.256 +lemma not_Object_subcls_any [elim!]: "(Object, C) \<in> (subcls1 tprg)^+ \<Longrightarrow> R"
  11.257 +apply (auto dest!: tranclD subcls1D)
  11.258 +done
  11.259 +
  11.260 +lemma not_Throwable_subcls_SXcpt [elim!]: 
  11.261 +  "(SXcpt Throwable, SXcpt xn) \<in> (subcls1 tprg)^+ \<Longrightarrow> R"
  11.262 +apply (auto dest!: tranclD subcls1D) 
  11.263 +apply (simp add: Object_def SXcpt_def)
  11.264 +done
  11.265 +
  11.266 +lemma not_SXcpt_n_subcls_SXcpt_n [elim!]: 
  11.267 +  "(SXcpt xn, SXcpt xn)  \<in> (subcls1 tprg)^+ \<Longrightarrow> R"
  11.268 +apply (auto dest!: tranclD subcls1D)
  11.269 +apply (drule rtranclD)
  11.270 +apply auto
  11.271 +done
  11.272 +
  11.273 +lemma not_Base_subcls_Ext [elim!]: "(Base, Ext) \<in> (subcls1 tprg)^+  \<Longrightarrow> R"
  11.274 +apply (auto dest!: tranclD subcls1D simp add: BaseCl_def)
  11.275 +done
  11.276 +
  11.277 +lemma not_TName_n_subcls_TName_n [rule_format (no_asm), elim!]: 
  11.278 +  "(\<lparr>pid=java_lang,tid=TName tn\<rparr>, \<lparr>pid=java_lang,tid=TName tn\<rparr>) 
  11.279 +   \<in> (subcls1 tprg)^+ \<longrightarrow> R"
  11.280 +apply (rule_tac n1 = "tn" in surj_tnam_ [THEN exE])
  11.281 +apply (erule ssubst)
  11.282 +apply (rule tnam_.induct)
  11.283 +apply  safe
  11.284 +apply (auto dest!: tranclD subcls1D simp add: BaseCl_def ExtCl_def)
  11.285 +apply (drule rtranclD)
  11.286 +apply auto
  11.287 +done
  11.288 +
  11.289 +
  11.290 +lemma ws_idecl_HasFoo: "ws_idecl tprg HasFoo []"
  11.291 +apply (unfold ws_idecl_def)
  11.292 +apply (simp (no_asm))
  11.293 +done
  11.294 +
  11.295 +lemma ws_cdecl_Object: "ws_cdecl tprg Object any"
  11.296 +apply (unfold ws_cdecl_def)
  11.297 +apply auto
  11.298 +done
  11.299 +
  11.300 +lemma ws_cdecl_Throwable: "ws_cdecl tprg (SXcpt Throwable) Object"
  11.301 +apply (unfold ws_cdecl_def)
  11.302 +apply auto
  11.303 +done
  11.304 +
  11.305 +lemma ws_cdecl_SXcpt: "ws_cdecl tprg (SXcpt xn) (SXcpt Throwable)"
  11.306 +apply (unfold ws_cdecl_def)
  11.307 +apply auto
  11.308 +done
  11.309 +
  11.310 +lemma ws_cdecl_Base: "ws_cdecl tprg Base Object"
  11.311 +apply (unfold ws_cdecl_def)
  11.312 +apply auto
  11.313 +done
  11.314 +
  11.315 +lemma ws_cdecl_Ext: "ws_cdecl tprg Ext Base"
  11.316 +apply (unfold ws_cdecl_def)
  11.317 +apply auto
  11.318 +done
  11.319 +
  11.320 +lemmas ws_cdecls = ws_cdecl_SXcpt ws_cdecl_Object ws_cdecl_Throwable
  11.321 +                   ws_cdecl_Base ws_cdecl_Ext
  11.322 +
  11.323 +declare not_Object_subcls_any [rule del]
  11.324 +          not_Throwable_subcls_SXcpt [rule del] 
  11.325 +          not_SXcpt_n_subcls_SXcpt_n [rule del] 
  11.326 +          not_Base_subcls_Ext [rule del] not_TName_n_subcls_TName_n [rule del]
  11.327 +
  11.328 +lemma ws_idecl_all: 
  11.329 + "G=tprg \<Longrightarrow> (\<forall>(I,i)\<in>set Ifaces. ws_idecl G I (isuperIfs i))"
  11.330 +apply (simp (no_asm) add: Ifaces_def HasFooInt_def)
  11.331 +apply (auto intro!: ws_idecl_HasFoo)
  11.332 +done
  11.333 +
  11.334 +lemma ws_cdecl_all: "G=tprg \<Longrightarrow> (\<forall>(C,c)\<in>set Classes. ws_cdecl G C (super c))"
  11.335 +apply (simp (no_asm) add: Classes_def BaseCl_def ExtCl_def)
  11.336 +apply (auto intro!: ws_cdecls simp add: standard_classes_def ObjectC_def 
  11.337 +        SXcptC_def)
  11.338 +done
  11.339 +
  11.340 +lemma ws_tprg: "ws_prog tprg"
  11.341 +apply (unfold ws_prog_def)
  11.342 +apply (auto intro!: ws_idecl_all ws_cdecl_all)
  11.343 +done
  11.344 +
  11.345 +
  11.346 +section "misc program properties (independent of well-structuredness)"
  11.347 +
  11.348 +lemma single_iface [simp]: "is_iface tprg I = (I = HasFoo)"
  11.349 +apply (unfold Ifaces_def)
  11.350 +apply (simp (no_asm))
  11.351 +done
  11.352 +
  11.353 +lemma empty_subint1 [simp]: "subint1 tprg = {}"
  11.354 +apply (unfold subint1_def Ifaces_def HasFooInt_def)
  11.355 +apply auto
  11.356 +done
  11.357 +
  11.358 +lemma unique_ifaces: "unique Ifaces"
  11.359 +apply (unfold Ifaces_def)
  11.360 +apply (simp (no_asm))
  11.361 +done
  11.362 +
  11.363 +lemma unique_classes: "unique Classes"
  11.364 +apply (unfold table_classes_defs )
  11.365 +apply (simp )
  11.366 +done
  11.367 +
  11.368 +lemma SXcpt_subcls_Throwable [simp]: "tprg\<turnstile>SXcpt xn\<preceq>\<^sub>C SXcpt Throwable"
  11.369 +apply (rule SXcpt_subcls_Throwable_lemma)
  11.370 +apply auto
  11.371 +done
  11.372 +
  11.373 +lemma Ext_subclseq_Base [simp]: "tprg\<turnstile>Ext \<preceq>\<^sub>C Base"
  11.374 +apply (rule subcls_direct1)
  11.375 +apply  (simp (no_asm) add: ExtCl_def)
  11.376 +apply  (simp add: Object_def)
  11.377 +apply (simp (no_asm))
  11.378 +done
  11.379 +
  11.380 +lemma Ext_subcls_Base [simp]: "tprg\<turnstile>Ext \<prec>\<^sub>C Base"
  11.381 +apply (rule subcls_direct2)
  11.382 +apply  (simp (no_asm) add: ExtCl_def)
  11.383 +apply  (simp add: Object_def)
  11.384 +apply (simp (no_asm))
  11.385 +done
  11.386 +
  11.387 +
  11.388 +
  11.389 +section "fields and method lookup"
  11.390 +
  11.391 +lemma fields_tprg_Object [simp]: "DeclConcepts.fields tprg Object = []"
  11.392 +by (rule ws_tprg [THEN fields_emptyI], force+)
  11.393 +
  11.394 +lemma fields_tprg_Throwable [simp]: 
  11.395 +  "DeclConcepts.fields tprg (SXcpt Throwable) = []"
  11.396 +by (rule ws_tprg [THEN fields_emptyI], force+)
  11.397 +
  11.398 +lemma fields_tprg_SXcpt [simp]: "DeclConcepts.fields tprg (SXcpt xn) = []"
  11.399 +apply (case_tac "xn = Throwable")
  11.400 +apply  (simp (no_asm_simp))
  11.401 +by (rule ws_tprg [THEN fields_emptyI], force+)
  11.402 +
  11.403 +lemmas fields_rec_ = fields_rec [OF _ ws_tprg]
  11.404 +
  11.405 +lemma fields_Base [simp]: 
  11.406 +"DeclConcepts.fields tprg Base 
  11.407 +  = [((arr,Base), \<lparr>access=Public,static=True ,type=PrimT Boolean.[]\<rparr>),
  11.408 +     ((vee,Base), \<lparr>access=Public,static=False,type=Iface HasFoo    \<rparr>)]"
  11.409 +apply (subst fields_rec_)
  11.410 +apply   (auto simp add: BaseCl_def)
  11.411 +done
  11.412 +
  11.413 +lemma fields_Ext [simp]: 
  11.414 + "DeclConcepts.fields tprg Ext  
  11.415 +  = [((vee,Ext), \<lparr>access=Public,static=False,type= PrimT Integer\<rparr>)] 
  11.416 +    @ DeclConcepts.fields tprg Base"
  11.417 +apply (rule trans)
  11.418 +apply (rule fields_rec_)
  11.419 +apply   (auto simp add: ExtCl_def Object_def)
  11.420 +done
  11.421 +
  11.422 +lemmas imethds_rec_ = imethds_rec [OF _ ws_tprg]
  11.423 +lemmas methd_rec_  = methd_rec  [OF _ ws_tprg]
  11.424 +
  11.425 +lemma imethds_HasFoo [simp]: 
  11.426 +  "imethds tprg HasFoo = o2s \<circ> empty(foo_sig\<mapsto>(HasFoo, foo_mhead))"
  11.427 +apply (rule trans)
  11.428 +apply (rule imethds_rec_)
  11.429 +apply  (auto simp add: HasFooInt_def)
  11.430 +done
  11.431 +
  11.432 +lemma methd_tprg_Object [simp]: "methd tprg Object = empty"
  11.433 +apply (subst methd_rec_)
  11.434 +apply (auto simp add: Object_mdecls_def)
  11.435 +done
  11.436 +
  11.437 +lemma methd_Base [simp]: 
  11.438 +  "methd tprg Base = table_of [(\<lambda>(s,m). (s, Base, m)) Base_foo]"
  11.439 +apply (rule trans)
  11.440 +apply (rule methd_rec_)
  11.441 +apply   (auto simp add: BaseCl_def)
  11.442 +done
  11.443 +
  11.444 +(* ### To Table *)
  11.445 +lemma filter_tab_all_False: 
  11.446 + "\<forall> k y. t k = Some y \<longrightarrow> \<not> p k y \<Longrightarrow>filter_tab p t = empty"
  11.447 +by (auto simp add: filter_tab_def expand_fun_eq)
  11.448 +
  11.449 +
  11.450 +lemma memberid_Base_foo_simp [simp]:
  11.451 + "memberid (mdecl Base_foo) = mid foo_sig"
  11.452 +by (simp add: Base_foo_def)
  11.453 +
  11.454 +lemma memberid_Ext_foo_simp [simp]:
  11.455 + "memberid (mdecl Ext_foo) = mid foo_sig"
  11.456 +by (simp add: Ext_foo_def)
  11.457 +
  11.458 +lemma Base_declares_foo:
  11.459 +  "tprg\<turnstile>mdecl Base_foo declared_in Base"
  11.460 +by (auto simp add: declared_in_def cdeclaredmethd_def BaseCl_def Base_foo_def)
  11.461 +
  11.462 +lemma foo_sig_not_undeclared_in_Base:
  11.463 +  "\<not> tprg\<turnstile>mid foo_sig undeclared_in Base"
  11.464 +proof -
  11.465 +  from Base_declares_foo
  11.466 +  show ?thesis
  11.467 +    by (auto dest!: declared_not_undeclared )
  11.468 +qed
  11.469 +
  11.470 +lemma Ext_declares_foo:
  11.471 +  "tprg\<turnstile>mdecl Ext_foo declared_in Ext"
  11.472 +by (auto simp add: declared_in_def cdeclaredmethd_def ExtCl_def Ext_foo_def)
  11.473 +
  11.474 +lemma foo_sig_not_undeclared_in_Ext:
  11.475 +  "\<not> tprg\<turnstile>mid foo_sig undeclared_in Ext"
  11.476 +proof -
  11.477 +  from Ext_declares_foo
  11.478 +  show ?thesis
  11.479 +    by (auto dest!: declared_not_undeclared )
  11.480 +qed
  11.481 +
  11.482 +lemma Base_foo_not_inherited_in_Ext:
  11.483 + "\<not> tprg \<turnstile> Ext inherits (Base,mdecl Base_foo)"
  11.484 +by (auto simp add: inherits_def foo_sig_not_undeclared_in_Ext)
  11.485 +
  11.486 +lemma Ext_method_inheritance:
  11.487 + "filter_tab (\<lambda>sig m. tprg \<turnstile> Ext inherits method sig m)
  11.488 +     (empty(fst ((\<lambda>(s, m). (s, Base, m)) Base_foo)\<mapsto>
  11.489 +      snd ((\<lambda>(s, m). (s, Base, m)) Base_foo)))
  11.490 +  = empty"
  11.491 +proof -
  11.492 +  from Base_foo_not_inherited_in_Ext
  11.493 +  show ?thesis
  11.494 +    by (auto intro: filter_tab_all_False simp add: Base_foo_def)
  11.495 +qed
  11.496 +
  11.497 +
  11.498 +lemma methd_Ext [simp]: "methd tprg Ext =   
  11.499 +  table_of [(\<lambda>(s,m). (s, Ext, m)) Ext_foo]"
  11.500 +apply (rule trans)
  11.501 +apply (rule methd_rec_)
  11.502 +apply   (auto simp add: ExtCl_def Object_def Ext_method_inheritance)
  11.503 +done
  11.504 +
  11.505 +section "accessibility"
  11.506 +
  11.507 +lemma classesDefined: 
  11.508 + "\<lbrakk>class tprg C = Some c; C\<noteq>Object\<rbrakk> \<Longrightarrow> \<exists> sc. class tprg (super c) = Some sc"
  11.509 +apply (auto simp add: Classes_def standard_classes_def 
  11.510 +                      BaseCl_def ExtCl_def
  11.511 +                      SXcptC_def ObjectC_def) 
  11.512 +done
  11.513 +
  11.514 +lemma superclassesBase [simp]: "superclasses tprg Base={Object}"
  11.515 +proof -
  11.516 +  have ws: "ws_prog tprg" by (rule ws_tprg)
  11.517 +  then show ?thesis
  11.518 +    by (auto simp add: superclasses_rec  BaseCl_def)
  11.519 +qed
  11.520 +
  11.521 +lemma superclassesExt [simp]: "superclasses tprg Ext={Base,Object}"
  11.522 +proof -
  11.523 +  have ws: "ws_prog tprg" by (rule ws_tprg)
  11.524 +  then show ?thesis
  11.525 +    by (auto simp add: superclasses_rec  ExtCl_def BaseCl_def)
  11.526 +qed
  11.527 +
  11.528 +lemma HasFoo_accessible[simp]:"tprg\<turnstile>(Iface HasFoo) accessible_in P" 
  11.529 +by (simp add: accessible_in_RefT_simp is_public_def HasFooInt_def)
  11.530 +
  11.531 +lemma HasFoo_is_acc_iface[simp]: "is_acc_iface tprg P HasFoo"
  11.532 +by (simp add: is_acc_iface_def)
  11.533 +
  11.534 +lemma HasFoo_is_acc_type[simp]: "is_acc_type tprg P (Iface HasFoo)"
  11.535 +by (simp add: is_acc_type_def)
  11.536 +
  11.537 +lemma Base_accessible[simp]:"tprg\<turnstile>(Class Base) accessible_in P" 
  11.538 +by (simp add: accessible_in_RefT_simp is_public_def BaseCl_def)
  11.539 +
  11.540 +lemma Base_is_acc_class[simp]: "is_acc_class tprg P Base"
  11.541 +by (simp add: is_acc_class_def)
  11.542 +
  11.543 +lemma Base_is_acc_type[simp]: "is_acc_type tprg P (Class Base)"
  11.544 +by (simp add: is_acc_type_def)
  11.545 +
  11.546 +lemma Ext_accessible[simp]:"tprg\<turnstile>(Class Ext) accessible_in P" 
  11.547 +by (simp add: accessible_in_RefT_simp is_public_def ExtCl_def)
  11.548 +
  11.549 +lemma Ext_is_acc_class[simp]: "is_acc_class tprg P Ext"
  11.550 +by (simp add: is_acc_class_def)
  11.551 +
  11.552 +lemma Ext_is_acc_type[simp]: "is_acc_type tprg P (Class Ext)"
  11.553 +by (simp add: is_acc_type_def)
  11.554 +
  11.555 +lemma accmethd_tprg_Object [simp]: "accmethd tprg S Object = empty"
  11.556 +apply (unfold accmethd_def)
  11.557 +apply (simp)
  11.558 +done
  11.559 +
  11.560 +lemma  snd_special_simp: "snd ((\<lambda>(s, m). (s, a, m)) x) = (a,snd x)"
  11.561 +by (cases x) (auto)
  11.562 +
  11.563 +lemma  fst_special_simp: "fst ((\<lambda>(s, m). (s, a, m)) x) = fst x"
  11.564 +by (cases x) (auto)
  11.565 +
  11.566 +lemma foo_sig_undeclared_in_Object:
  11.567 +  "tprg\<turnstile>mid foo_sig undeclared_in Object"
  11.568 +by (auto simp add: undeclared_in_def cdeclaredmethd_def Object_mdecls_def)
  11.569 +
  11.570 +(* ### To DeclConcepts *)
  11.571 +lemma undeclared_not_declared:
  11.572 + "G\<turnstile> memberid m undeclared_in C \<Longrightarrow> \<not> G\<turnstile> m declared_in C" 
  11.573 +by (cases m) (auto simp add: declared_in_def undeclared_in_def)
  11.574 +
  11.575 +
  11.576 +lemma unique_sig_Base_foo:
  11.577 + "tprg\<turnstile> mdecl (sig, snd Base_foo) declared_in Base \<Longrightarrow> sig=foo_sig"
  11.578 +by (auto simp add: declared_in_def cdeclaredmethd_def 
  11.579 +                   Base_foo_def BaseCl_def)
  11.580 +
  11.581 +lemma Base_foo_no_override:
  11.582 + "tprg,sig\<turnstile>(Base,(snd Base_foo)) overrides old \<Longrightarrow> P"
  11.583 +apply (drule overrides_commonD)
  11.584 +apply (clarsimp)
  11.585 +apply (frule subclsEval)
  11.586 +apply    (rule ws_tprg)
  11.587 +apply    (simp)
  11.588 +apply    (rule classesDefined) 
  11.589 +apply    assumption+
  11.590 +apply (frule unique_sig_Base_foo) 
  11.591 +apply (auto dest!: declared_not_undeclared intro: foo_sig_undeclared_in_Object
  11.592 +            dest: unique_sig_Base_foo)
  11.593 +done
  11.594 +
  11.595 +lemma Base_foo_no_stat_override:
  11.596 + "tprg,sig\<turnstile>(Base,(snd Base_foo)) overrides\<^sub>S old \<Longrightarrow> P"
  11.597 +apply (drule stat_overrides_commonD)
  11.598 +apply (clarsimp)
  11.599 +apply (frule subclsEval)
  11.600 +apply    (rule ws_tprg)
  11.601 +apply    (simp)
  11.602 +apply    (rule classesDefined) 
  11.603 +apply    assumption+
  11.604 +apply (frule unique_sig_Base_foo) 
  11.605 +apply (auto dest!: declared_not_undeclared intro: foo_sig_undeclared_in_Object
  11.606 +            dest: unique_sig_Base_foo)
  11.607 +done
  11.608 +
  11.609 +
  11.610 +lemma Base_foo_no_hide:
  11.611 + "tprg,sig\<turnstile>(Base,(snd Base_foo)) hides old \<Longrightarrow> P"
  11.612 +by (auto dest: hidesD simp add: Base_foo_def member_is_static_simp)
  11.613 +
  11.614 +lemma Ext_foo_no_hide:
  11.615 + "tprg,sig\<turnstile>(Ext,(snd Ext_foo)) hides old \<Longrightarrow> P"
  11.616 +by (auto dest: hidesD simp add: Ext_foo_def member_is_static_simp)
  11.617 +
  11.618 +lemma unique_sig_Ext_foo:
  11.619 + "tprg\<turnstile> mdecl (sig, snd Ext_foo) declared_in Ext \<Longrightarrow> sig=foo_sig"
  11.620 +by (auto simp add: declared_in_def cdeclaredmethd_def 
  11.621 +                   Ext_foo_def ExtCl_def)
  11.622 +
  11.623 +(* ### To DeclConcepts *)
  11.624 +lemma unique_declaration: 
  11.625 + "\<lbrakk>G\<turnstile>m declared_in C;  G\<turnstile>n declared_in C; memberid m = memberid n \<rbrakk> 
  11.626 +  \<Longrightarrow> m = n"
  11.627 +apply (cases m)
  11.628 +apply (cases n,
  11.629 +        auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)+
  11.630 +done
  11.631 +
  11.632 +
  11.633 +lemma Ext_foo_override:
  11.634 + "tprg,sig\<turnstile>(Ext,(snd Ext_foo)) overrides old 
  11.635 +  \<Longrightarrow> old = (Base,(snd Base_foo))"
  11.636 +apply (drule overrides_commonD)
  11.637 +apply (clarsimp)
  11.638 +apply (frule subclsEval)
  11.639 +apply    (rule ws_tprg)
  11.640 +apply    (simp)
  11.641 +apply    (rule classesDefined) 
  11.642 +apply    assumption+
  11.643 +apply (frule unique_sig_Ext_foo) 
  11.644 +apply (case_tac "old")
  11.645 +apply (insert Base_declares_foo foo_sig_undeclared_in_Object) 
  11.646 +apply (auto simp add: ExtCl_def Ext_foo_def
  11.647 +                      BaseCl_def Base_foo_def Object_mdecls_def
  11.648 +                      split_paired_all
  11.649 +                      member_is_static_simp
  11.650 +           dest: declared_not_undeclared unique_declaration)
  11.651 +done
  11.652 +
  11.653 +lemma Ext_foo_stat_override:
  11.654 + "tprg,sig\<turnstile>(Ext,(snd Ext_foo)) overrides\<^sub>S old 
  11.655 +  \<Longrightarrow> old = (Base,(snd Base_foo))"
  11.656 +apply (drule stat_overrides_commonD)
  11.657 +apply (clarsimp)
  11.658 +apply (frule subclsEval)
  11.659 +apply    (rule ws_tprg)
  11.660 +apply    (simp)
  11.661 +apply    (rule classesDefined) 
  11.662 +apply    assumption+
  11.663 +apply (frule unique_sig_Ext_foo) 
  11.664 +apply (case_tac "old")
  11.665 +apply (insert Base_declares_foo foo_sig_undeclared_in_Object) 
  11.666 +apply (auto simp add: ExtCl_def Ext_foo_def
  11.667 +                      BaseCl_def Base_foo_def Object_mdecls_def
  11.668 +                      split_paired_all
  11.669 +                      member_is_static_simp
  11.670 +           dest: declared_not_undeclared unique_declaration)
  11.671 +done
  11.672 +
  11.673 +(*### weiter hoch *)
  11.674 +lemma Base_foo_member_of_Base: 
  11.675 +  "tprg\<turnstile>(Base,mdecl Base_foo) member_of Base"
  11.676 +by (auto intro!: members.Immediate Base_declares_foo)
  11.677 +
  11.678 +(*### weiter hoch *)
  11.679 +lemma Ext_foo_member_of_Ext: 
  11.680 +  "tprg\<turnstile>(Ext,mdecl Ext_foo) member_of Ext"
  11.681 +by (auto intro!: members.Immediate Ext_declares_foo)
  11.682 +
  11.683 +lemma Base_foo_permits_acc:
  11.684 + "tprg \<turnstile> (Base, mdecl Base_foo) in Base permits_acc_to S"
  11.685 +by ( simp add: permits_acc_def Base_foo_def)
  11.686 +
  11.687 +lemma Base_foo_accessible [simp]:
  11.688 + "tprg\<turnstile>(Base,mdecl Base_foo) of Base accessible_from S"
  11.689 +by (auto intro: accessible_fromR.immediate 
  11.690 +                Base_foo_member_of_Base Base_foo_permits_acc)
  11.691 +
  11.692 +lemma accmethd_Base [simp]: 
  11.693 +  "accmethd tprg S Base = methd tprg Base"
  11.694 +apply (simp add: accmethd_def)
  11.695 +apply (rule filter_tab_all_True)
  11.696 +apply (simp add: snd_special_simp fst_special_simp)
  11.697 +done
  11.698 +
  11.699 +lemma Ext_foo_permits_acc:
  11.700 + "tprg \<turnstile> (Ext, mdecl Ext_foo) in Ext permits_acc_to S"
  11.701 +by ( simp add: permits_acc_def Ext_foo_def)
  11.702 +
  11.703 +lemma Ext_foo_accessible [simp]:
  11.704 + "tprg\<turnstile>(Ext,mdecl Ext_foo) of Ext accessible_from S"
  11.705 +by (auto intro: accessible_fromR.immediate 
  11.706 +                Ext_foo_member_of_Ext Ext_foo_permits_acc)
  11.707 +
  11.708 +(*
  11.709 +lemma Base_foo_accessible_through_inheritance_in_Ext [simp]:
  11.710 + "tprg\<turnstile>(Base,snd Base_foo) accessible_through_inheritance_in Ext"
  11.711 +apply (rule accessible_through_inheritance.Direct)
  11.712 +apply   simp
  11.713 +apply   (simp add: accessible_for_inheritance_in_def Base_foo_def)
  11.714 +done
  11.715 +*)
  11.716 +
  11.717 +lemma Ext_foo_overrides_Base_foo:
  11.718 + "tprg\<turnstile>(Ext,Ext_foo) overrides (Base,Base_foo)"
  11.719 +proof (rule overridesR.Direct, simp_all)
  11.720 +  show "\<not> is_static Ext_foo" 
  11.721 +    by (simp add: member_is_static_simp Ext_foo_def)
  11.722 +  show "\<not> is_static Base_foo"
  11.723 +    by (simp add: member_is_static_simp Base_foo_def)
  11.724 +  show "accmodi Ext_foo \<noteq> Private"
  11.725 +    by (simp add: Ext_foo_def)
  11.726 +  show "msig (Ext, Ext_foo) = msig (Base, Base_foo)"
  11.727 +    by (simp add: Ext_foo_def Base_foo_def)
  11.728 +  show "tprg\<turnstile>mdecl Ext_foo declared_in Ext"
  11.729 +    by (auto intro: Ext_declares_foo)
  11.730 +  show "tprg\<turnstile>mdecl Base_foo declared_in Base"
  11.731 +    by (auto intro: Base_declares_foo)
  11.732 +  show "tprg \<turnstile>(Base, mdecl Base_foo) inheritable_in java_lang"
  11.733 +    by (simp add: inheritable_in_def Base_foo_def)
  11.734 +  show "tprg\<turnstile>resTy Ext_foo\<preceq>resTy Base_foo"
  11.735 +    by (simp add: Ext_foo_def Base_foo_def mhead_resTy_simp)
  11.736 +qed
  11.737 +
  11.738 +(*
  11.739 +lemma Base_foo_of_Ext_accessible[simp]:
  11.740 + "tprg\<turnstile>(Base, mdecl Base_foo) of Ext accessible_from S"
  11.741 +apply (auto intro: accessible_fromR.immediate 
  11.742 +                Base_foo_member_of_Base Base_foo_permits_acc)
  11.743 +apply (rule accessible_fromR.immediate)
  11.744 +apply (rule_tac "old"="(Base,Base_foo)" and  sup="Base" 
  11.745 +       in accessible_fromR.overriding)
  11.746 +apply (auto intro!: Ext_foo_overrides_Base_foo)
  11.747 +apply (auto 
  11.748 +apply (insert Ext_foo_overrides_Base_foo)
  11.749 +apply (rule accessible_fromR.overriding, simp_all)
  11.750 +apply (auto intro!: Ext_foo_overrides_Base_foo)
  11.751 +apply (auto intro!: accessible_fromR.overriding
  11.752 +             intro:   Ext_foo_overrides_Base_foo)
  11.753 +by
  11.754 +                Ext_foo_member_of_Ext Ext_foo_permits_acc)
  11.755 +apply (auto intro!: accessible 
  11.756 +apply (auto simp add: method_accessible_from_def accessible_from_def) 
  11.757 +apply (simp add: Base_foo_def)
  11.758 +done 
  11.759 +*)
  11.760 +
  11.761 +lemma accmethd_Ext [simp]: 
  11.762 +  "accmethd tprg S Ext = methd tprg Ext"
  11.763 +apply (simp add: accmethd_def)
  11.764 +apply (rule filter_tab_all_True)
  11.765 +apply (auto simp add: snd_special_simp fst_special_simp)
  11.766 +done
  11.767 +
  11.768 +(* ### Weiter hoch *)
  11.769 +lemma cls_Ext: "class tprg Ext = Some ExtCl"
  11.770 +by simp
  11.771 +lemma dynmethd_Ext_foo:
  11.772 + "dynmethd tprg Base Ext \<lparr>name = foo, parTs = [Class Base]\<rparr> 
  11.773 +  = Some (Ext,snd Ext_foo)"
  11.774 +proof -
  11.775 +  have "methd tprg Base \<lparr>name = foo, parTs = [Class Base]\<rparr> 
  11.776 +          = Some (Base,snd Base_foo)" and
  11.777 +       "methd tprg Ext \<lparr>name = foo, parTs = [Class Base]\<rparr> 
  11.778 +          = Some (Ext,snd Ext_foo)" 
  11.779 +    by (auto simp add: Ext_foo_def Base_foo_def foo_sig_def)
  11.780 +  with cls_Ext ws_tprg Ext_foo_overrides_Base_foo
  11.781 +  show ?thesis
  11.782 +    by (auto simp add: dynmethd_rec simp add: Ext_foo_def Base_foo_def)
  11.783 +qed
  11.784 +
  11.785 +lemma Base_fields_accessible[simp]:
  11.786 + "accfield tprg S Base 
  11.787 +  = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (DeclConcepts.fields tprg Base))"
  11.788 +apply (auto simp add: accfield_def expand_fun_eq Let_def 
  11.789 +                      accessible_in_RefT_simp
  11.790 +                      is_public_def
  11.791 +                      BaseCl_def
  11.792 +                      permits_acc_def
  11.793 +                      declared_in_def 
  11.794 +                      cdeclaredfield_def
  11.795 +               intro!: filter_tab_all_True_Some filter_tab_None
  11.796 +                       accessible_fromR.immediate
  11.797 +               intro: members.Immediate)
  11.798 +done
  11.799 +
  11.800 +
  11.801 +lemma arr_member_of_Base:
  11.802 +  "tprg\<turnstile>(Base, fdecl (arr, 
  11.803 +                 \<lparr>access = Public, static = True, type = PrimT Boolean.[]\<rparr>))
  11.804 +          member_of Base"
  11.805 +by (auto intro: members.Immediate 
  11.806 +       simp add: declared_in_def cdeclaredfield_def BaseCl_def)
  11.807 + 
  11.808 +lemma arr_member_of_Ext:
  11.809 +  "tprg\<turnstile>(Base, fdecl (arr, 
  11.810 +                    \<lparr>access = Public, static = True, type = PrimT Boolean.[]\<rparr>))
  11.811 +             member_of Ext"
  11.812 +apply (rule members.Inherited)
  11.813 +apply   (simp add: inheritable_in_def)
  11.814 +apply   (simp add: undeclared_in_def cdeclaredfield_def ExtCl_def)
  11.815 +apply   (auto intro: arr_member_of_Base simp add: subcls1_def ExtCl_def)
  11.816 +done
  11.817 +
  11.818 +lemma Ext_fields_accessible[simp]:
  11.819 +"accfield tprg S Ext 
  11.820 +  = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (DeclConcepts.fields tprg Ext))"
  11.821 +apply (auto simp add: accfield_def expand_fun_eq Let_def 
  11.822 +                      accessible_in_RefT_simp
  11.823 +                      is_public_def
  11.824 +                      BaseCl_def
  11.825 +                      ExtCl_def
  11.826 +                      permits_acc_def
  11.827 +               intro!: filter_tab_all_True_Some filter_tab_None
  11.828 +                       accessible_fromR.immediate)
  11.829 +apply (auto intro: members.Immediate arr_member_of_Ext
  11.830 +            simp add: declared_in_def cdeclaredfield_def ExtCl_def)
  11.831 +done
  11.832 +
  11.833 +lemma array_of_PrimT_acc [simp]:
  11.834 + "is_acc_type tprg java_lang (PrimT t.[])"
  11.835 +apply (simp add: is_acc_type_def accessible_in_RefT_simp)
  11.836 +done
  11.837 +
  11.838 +lemma PrimT_acc [simp]: 
  11.839 + "is_acc_type tprg java_lang (PrimT t)"
  11.840 +apply (simp add: is_acc_type_def accessible_in_RefT_simp)
  11.841 +done
  11.842 +
  11.843 +lemma Object_acc [simp]:
  11.844 + "is_acc_class tprg java_lang Object"
  11.845 +apply (auto simp add: is_acc_class_def accessible_in_RefT_simp is_public_def)
  11.846 +done
  11.847 +
  11.848 +
  11.849 +section "well-formedness"
  11.850 +
  11.851 +
  11.852 +lemma wf_HasFoo: "wf_idecl tprg (HasFoo, HasFooInt)"
  11.853 +apply (unfold wf_idecl_def HasFooInt_def)
  11.854 +apply (auto intro!: wf_mheadI ws_idecl_HasFoo 
  11.855 +            simp add: foo_sig_def foo_mhead_def mhead_resTy_simp
  11.856 +                      member_is_static_simp )
  11.857 +done
  11.858 +
  11.859 +declare wt.Skip [rule del] wt.Init [rule del]
  11.860 +lemmas Base_foo_defs = Base_foo_def foo_sig_def foo_mhead_def
  11.861 +lemmas Ext_foo_defs  = Ext_foo_def  foo_sig_def
  11.862 +ML {* bind_thms ("wt_intros",map (rewrite_rule [id_def]) (thms "wt.intros")) *}
  11.863 +lemmas wtIs = wt_Call wt_Super wt_FVar wt_StatRef wt_intros
  11.864 +
  11.865 +lemma wf_Base_foo: "wf_mdecl tprg Base Base_foo"
  11.866 +apply (unfold Base_foo_defs )
  11.867 +apply (auto intro!: wf_mdeclI wf_mheadI intro!: wtIs
  11.868 +            simp add: mhead_resTy_simp)
  11.869 +done
  11.870 +
  11.871 +lemma wf_Ext_foo: "wf_mdecl tprg Ext Ext_foo"
  11.872 +apply (unfold Ext_foo_defs )
  11.873 +apply (auto intro!: wf_mdeclI wf_mheadI intro!: wtIs 
  11.874 +            simp add: mhead_resTy_simp )
  11.875 +apply  (rule wt.Cast)
  11.876 +prefer 2
  11.877 +apply    simp
  11.878 +apply   (rule_tac [2] narrow.subcls [THEN cast.narrow])
  11.879 +apply   (auto intro!: wtIs)
  11.880 +done
  11.881 +
  11.882 +declare mhead_resTy_simp [simp add]
  11.883 +declare member_is_static_simp [simp add]
  11.884 +
  11.885 +lemma wf_BaseC: "wf_cdecl tprg (Base,BaseCl)"
  11.886 +apply (unfold wf_cdecl_def BaseCl_def arr_viewed_from_def)
  11.887 +apply (auto intro!: wf_Base_foo)
  11.888 +apply       (auto intro!: ws_cdecl_Base simp add: Base_foo_def foo_mhead_def)
  11.889 +apply (auto intro!: wtIs)
  11.890 +apply (auto simp add: Base_foo_defs entails_def Let_def)
  11.891 +apply   (insert Base_foo_no_stat_override, simp add: Base_foo_def,blast)+
  11.892 +apply   (insert Base_foo_no_hide         , simp add: Base_foo_def,blast)
  11.893 +done
  11.894 +
  11.895 +
  11.896 +lemma wf_ExtC: "wf_cdecl tprg (Ext,ExtCl)"
  11.897 +apply (unfold wf_cdecl_def ExtCl_def)
  11.898 +apply (auto intro!: wf_Ext_foo ws_cdecl_Ext)
  11.899 +apply (auto simp add: entails_def snd_special_simp)
  11.900 +apply (insert Ext_foo_stat_override)
  11.901 +apply  (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
  11.902 +apply  (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
  11.903 +apply  (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
  11.904 +apply (insert Ext_foo_no_hide)
  11.905 +apply  (simp_all add: qmdecl_def)
  11.906 +apply  blast+
  11.907 +done
  11.908 +
  11.909 +lemma wf_idecl_all: "p=tprg \<Longrightarrow> Ball (set Ifaces) (wf_idecl p)"
  11.910 +apply (simp (no_asm) add: Ifaces_def)
  11.911 +apply (simp (no_asm_simp))
  11.912 +apply (rule wf_HasFoo)
  11.913 +done
  11.914 +
  11.915 +lemma wf_cdecl_all_standard_classes: 
  11.916 +  "Ball (set standard_classes) (wf_cdecl tprg)"
  11.917 +apply (unfold standard_classes_def Let_def 
  11.918 +       ObjectC_def SXcptC_def Object_mdecls_def SXcpt_mdecls_def)
  11.919 +apply (simp (no_asm) add: wf_cdecl_def ws_cdecls)
  11.920 +apply (auto simp add:is_acc_class_def accessible_in_RefT_simp SXcpt_def)
  11.921 +apply (auto simp add: Object_def Classes_def standard_classes_def 
  11.922 +                      SXcptC_def SXcpt_def)
  11.923 +done
  11.924 +
  11.925 +lemma wf_cdecl_all: "p=tprg \<Longrightarrow> Ball (set Classes) (wf_cdecl p)"
  11.926 +apply (simp (no_asm) add: Classes_def)
  11.927 +apply (simp (no_asm_simp))
  11.928 +apply   (rule wf_BaseC [THEN conjI])
  11.929 +apply  (rule wf_ExtC [THEN conjI])
  11.930 +apply (rule wf_cdecl_all_standard_classes)
  11.931 +done
  11.932 +
  11.933 +theorem wf_tprg: "wf_prog tprg"
  11.934 +apply (unfold wf_prog_def Let_def)
  11.935 +apply (simp (no_asm) add: unique_ifaces unique_classes)
  11.936 +apply (rule conjI)
  11.937 +apply  ((simp (no_asm) add: Classes_def standard_classes_def))
  11.938 +apply (rule conjI)
  11.939 +apply (simp add: Object_mdecls_def)
  11.940 +apply safe
  11.941 +apply  (cut_tac xn_cases_old)   (* FIXME (insert xn_cases) *)
  11.942 +apply  (simp (no_asm_simp) add: Classes_def standard_classes_def)
  11.943 +apply  (insert wf_idecl_all)
  11.944 +apply  (insert wf_cdecl_all)
  11.945 +apply  auto
  11.946 +done
  11.947 +
  11.948 +
  11.949 +section "max spec"
  11.950 +
  11.951 +lemma appl_methds_Base_foo: 
  11.952 +"appl_methds tprg S (ClassT Base) \<lparr>name=foo, parTs=[NT]\<rparr> =  
  11.953 +  {((ClassT Base, \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>)
  11.954 +    ,[Class Base])}"
  11.955 +apply (unfold appl_methds_def)
  11.956 +apply (simp (no_asm))
  11.957 +apply (subgoal_tac "tprg\<turnstile>NT\<preceq> Class Base")
  11.958 +apply  (auto simp add: cmheads_def Base_foo_defs)
  11.959 +done
  11.960 +
  11.961 +lemma max_spec_Base_foo: "max_spec tprg S (ClassT Base) \<lparr>name=foo,parTs=[NT]\<rparr> = 
  11.962 +  {((ClassT Base, \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>)
  11.963 +   , [Class Base])}"
  11.964 +apply (unfold max_spec_def)
  11.965 +apply (simp (no_asm) add: appl_methds_Base_foo)
  11.966 +apply auto
  11.967 +done
  11.968 +
  11.969 +
  11.970 +section "well-typedness"
  11.971 +
  11.972 +lemma wt_test: "\<lparr>prg=tprg,cls=S,lcl=empty(VName e\<mapsto>Class Base)\<rparr>\<turnstile>test ?pTs\<Colon>\<surd>"
  11.973 +apply (unfold test_def arr_viewed_from_def)
  11.974 +(* ?pTs = [Class Base] *)
  11.975 +apply (rule wtIs (* ;; *))
  11.976 +apply  (rule wtIs (* Ass *))
  11.977 +apply  (rule wtIs (* NewC *))
  11.978 +apply     (rule wtIs (* LVar *))
  11.979 +apply      (simp)
  11.980 +apply     (simp)
  11.981 +apply    (simp)
  11.982 +apply   (rule wtIs (* NewC *))
  11.983 +apply   (simp)
  11.984 +apply  (simp)
  11.985 +apply (rule wtIs (* Try *))
  11.986 +prefer 4
  11.987 +apply    (simp)
  11.988 +defer
  11.989 +apply    (rule wtIs (* Expr *))
  11.990 +apply    (rule wtIs (* Call *))
  11.991 +apply       (rule wtIs (* Acc *))
  11.992 +apply       (rule wtIs (* LVar *))
  11.993 +apply        (simp)
  11.994 +apply       (simp)
  11.995 +apply      (rule wtIs (* Cons *))
  11.996 +apply       (rule wtIs (* Lit *))
  11.997 +apply       (simp)
  11.998 +apply      (rule wtIs (* Nil *))
  11.999 +apply     (simp)
 11.1000 +apply     (rule max_spec_Base_foo)
 11.1001 +apply    (simp)
 11.1002 +apply   (simp)
 11.1003 +apply  (simp)
 11.1004 +apply  (simp)
 11.1005 +apply (rule wtIs (* While *))
 11.1006 +apply  (rule wtIs (* Acc *))
 11.1007 +apply   (rule wtIs (* AVar *))
 11.1008 +apply   (rule wtIs (* Acc *))
 11.1009 +apply   (rule wtIs (* FVar *))
 11.1010 +apply    (rule wtIs (* StatRef *))
 11.1011 +apply    (simp)
 11.1012 +apply   (simp)
 11.1013 +apply   (simp )
 11.1014 +apply   (simp)
 11.1015 +apply  (rule wtIs (* LVar *))
 11.1016 +apply  (simp)
 11.1017 +apply (rule wtIs (* Skip *))
 11.1018 +done
 11.1019 +
 11.1020 +
 11.1021 +section "execution"
 11.1022 +
 11.1023 +lemma alloc_one: "\<And>a obj. \<lbrakk>the (new_Addr h) = a; atleast_free h (Suc n)\<rbrakk> \<Longrightarrow>  
 11.1024 +  new_Addr h = Some a \<and> atleast_free (h(a\<mapsto>obj)) n"
 11.1025 +apply (frule atleast_free_SucD)
 11.1026 +apply (drule atleast_free_Suc [THEN iffD1])
 11.1027 +apply clarsimp
 11.1028 +apply (frule new_Addr_SomeI)
 11.1029 +apply force
 11.1030 +done
 11.1031 +
 11.1032 +declare fvar_def2 [simp] avar_def2 [simp] init_lvars_def2 [simp]
 11.1033 +declare init_obj_def [simp] var_tys_def [simp] fields_table_def [simp]
 11.1034 +declare BaseCl_def [simp] ExtCl_def [simp] Ext_foo_def [simp]
 11.1035 +        Base_foo_defs  [simp]
 11.1036 +
 11.1037 +ML {* bind_thms ("eval_intros", map 
 11.1038 +        (simplify (simpset() delsimps [thm "Skip_eq"]
 11.1039 +                             addsimps [thm "lvar_def"]) o 
 11.1040 +         rewrite_rule [thm "assign_def",Let_def]) (thms "eval.intros")) *}
 11.1041 +lemmas eval_Is = eval_Init eval_StatRef AbruptIs eval_intros
 11.1042 +
 11.1043 +consts
 11.1044 +  a :: loc
 11.1045 +  b :: loc
 11.1046 +  c :: loc
 11.1047 +  
 11.1048 +syntax
 11.1049 +
 11.1050 +  tprg :: prog
 11.1051 +
 11.1052 +  obj_a :: obj
 11.1053 +  obj_b :: obj
 11.1054 +  obj_c :: obj
 11.1055 +  arr_N :: "(vn, val) table"
 11.1056 +  arr_a :: "(vn, val) table"
 11.1057 +  globs1 :: globs
 11.1058 +  globs2 :: globs
 11.1059 +  globs3 :: globs
 11.1060 +  globs8 :: globs
 11.1061 +  locs3 :: locals
 11.1062 +  locs4 :: locals
 11.1063 +  locs8 :: locals
 11.1064 +  s0  :: state
 11.1065 +  s0' :: state
 11.1066 +  s9' :: state
 11.1067 +  s1  :: state
 11.1068 +  s1' :: state
 11.1069 +  s2  :: state
 11.1070 +  s2' :: state
 11.1071 +  s3  :: state
 11.1072 +  s3' :: state
 11.1073 +  s4  :: state
 11.1074 +  s4' :: state
 11.1075 +  s6' :: state
 11.1076 +  s7' :: state
 11.1077 +  s8  :: state
 11.1078 +  s8' :: state
 11.1079 +
 11.1080 +translations
 11.1081 +
 11.1082 +  "tprg"    == "\<lparr>ifaces=Ifaces,classes=Classes\<rparr>"
 11.1083 +  
 11.1084 +  "obj_a"   <= "\<lparr>tag=Arr (PrimT Boolean) two
 11.1085 +                ,values=empty(Inr 0\<mapsto>Bool False)(Inr one\<mapsto>Bool False)\<rparr>"
 11.1086 +  "obj_b"   <= "\<lparr>tag=CInst Ext
 11.1087 +                ,values=(empty(Inl (vee, Base)\<mapsto>Null   ) 
 11.1088 +			      (Inl (vee, Ext )\<mapsto>Intg 0))\<rparr>"
 11.1089 +  "obj_c"   == "\<lparr>tag=CInst (SXcpt NullPointer),values=empty\<rparr>"
 11.1090 +  "arr_N"   == "empty(Inl (arr, Base)\<mapsto>Null)"
 11.1091 +  "arr_a"   == "empty(Inl (arr, Base)\<mapsto>Addr a)"
 11.1092 +  "globs1"  == "empty(Inr Ext   \<mapsto>\<lparr>tag=arbitrary, values=empty\<rparr>)
 11.1093 +		     (Inr Base  \<mapsto>\<lparr>tag=arbitrary, values=arr_N\<rparr>)
 11.1094 +		     (Inr Object\<mapsto>\<lparr>tag=arbitrary, values=empty\<rparr>)"
 11.1095 +  "globs2"  == "empty(Inr Ext   \<mapsto>\<lparr>tag=arbitrary, values=empty\<rparr>)
 11.1096 +		     (Inr Object\<mapsto>\<lparr>tag=arbitrary, values=empty\<rparr>)
 11.1097 +		     (Inl a\<mapsto>obj_a)
 11.1098 +		     (Inr Base  \<mapsto>\<lparr>tag=arbitrary, values=arr_a\<rparr>)"
 11.1099 +  "globs3"  == "globs2(Inl b\<mapsto>obj_b)"
 11.1100 +  "globs8"  == "globs3(Inl c\<mapsto>obj_c)"
 11.1101 +  "locs3"   == "empty(VName e\<mapsto>Addr b)"
 11.1102 +  "locs4"   == "empty(VName z\<mapsto>Null)(Inr()\<mapsto>Addr b)"
 11.1103 +  "locs8"   == "locs3(VName z\<mapsto>Addr c)"
 11.1104 +  "s0"  == "       st empty  empty"
 11.1105 +  "s0'" == " Norm  s0"
 11.1106 +  "s1"  == "       st globs1 empty"
 11.1107 +  "s1'" == " Norm  s1"
 11.1108 +  "s2"  == "       st globs2 empty"
 11.1109 +  "s2'" == " Norm  s2"
 11.1110 +  "s3"  == "       st globs3 locs3 "
 11.1111 +  "s3'" == " Norm  s3"
 11.1112 +  "s4"  == "       st globs3 locs4"
 11.1113 +  "s4'" == " Norm  s4"
 11.1114 +  "s6'" == "(Some (Xcpt (Std NullPointer)), s4)"
 11.1115 +  "s7'" == "(Some (Xcpt (Std NullPointer)), s3)"
 11.1116 +  "s8"  == "       st globs8 locs8"
 11.1117 +  "s8'" == " Norm  s8"
 11.1118 +  "s9'" == "(Some (Xcpt (Std IndOutBound)), s8)"
 11.1119 +
 11.1120 +
 11.1121 +syntax "four"::nat
 11.1122 +       "tree"::nat
 11.1123 +       "two" ::nat
 11.1124 +       "one" ::nat
 11.1125 +translations 
 11.1126 +  "one"  == "Suc 0"
 11.1127 +  "two"  == "Suc one"
 11.1128 +  "tree" == "Suc two"
 11.1129 +  "four" == "Suc tree"
 11.1130 +
 11.1131 +declare Pair_eq [simp del]
 11.1132 +lemma exec_test: 
 11.1133 +"\<lbrakk>the (new_Addr (heap  s1)) = a;  
 11.1134 +  the (new_Addr (heap ?s2)) = b;  
 11.1135 +  the (new_Addr (heap ?s3)) = c\<rbrakk> \<Longrightarrow>  
 11.1136 +  atleast_free  (heap s0) four \<Longrightarrow>  
 11.1137 +  tprg\<turnstile>s0' \<midarrow>test [Class Base]\<rightarrow> ?s9'"
 11.1138 +apply (unfold test_def arr_viewed_from_def)
 11.1139 +(* ?s9' = s9' *)
 11.1140 +apply (simp (no_asm_use))
 11.1141 +apply (drule (1) alloc_one,clarsimp)
 11.1142 +apply (rule eval_Is (* ;; *))
 11.1143 +apply  (erule_tac V = "the (new_Addr ?h) = c" in thin_rl)
 11.1144 +apply  (erule_tac [2] V = "new_Addr ?h = Some a" in thin_rl)
 11.1145 +apply  (erule_tac [2] V = "atleast_free ?h four" in thin_rl)
 11.1146 +apply  (rule eval_Is (* Expr *))
 11.1147 +apply  (rule eval_Is (* Ass *))
 11.1148 +apply   (rule eval_Is (* LVar *))
 11.1149 +apply  (rule eval_Is (* NewC *))
 11.1150 +      (* begin init Ext *)
 11.1151 +apply   (erule_tac V = "the (new_Addr ?h) = b" in thin_rl)
 11.1152 +apply   (erule_tac V = "atleast_free ?h tree" in thin_rl)
 11.1153 +apply   (erule_tac [2] V = "atleast_free ?h four" in thin_rl)
 11.1154 +apply   (erule_tac [2] V = "new_Addr ?h = Some a" in thin_rl)
 11.1155 +apply   (rule eval_Is (* Init Ext *))
 11.1156 +apply   (simp)
 11.1157 +apply   (rule conjI)
 11.1158 +prefer 2 apply (rule conjI HOL.refl)+
 11.1159 +apply   (rule eval_Is (* Init Base *))
 11.1160 +apply   (simp add: arr_viewed_from_def)
 11.1161 +apply   (rule conjI)
 11.1162 +apply    (rule eval_Is (* Init Object *))
 11.1163 +apply    (simp)
 11.1164 +apply    (rule conjI, rule HOL.refl)+
 11.1165 +apply    (rule HOL.refl)
 11.1166 +apply   (simp)
 11.1167 +apply   (rule conjI, rule_tac [2] HOL.refl)
 11.1168 +apply   (rule eval_Is (* Expr *))
 11.1169 +apply   (rule eval_Is (* Ass *))
 11.1170 +apply    (rule eval_Is (* FVar *))
 11.1171 +apply      (rule init_done, simp)
 11.1172 +apply     (rule eval_Is (* StatRef *))
 11.1173 +apply    (simp)
 11.1174 +apply   (rule eval_Is (* NewA *))
 11.1175 +apply     (simp)
 11.1176 +apply    (rule eval_Is (* Lit *))
 11.1177 +apply   (simp)
 11.1178 +apply   (rule halloc.New)
 11.1179 +apply    (simp (no_asm_simp))
 11.1180 +apply   (drule atleast_free_weaken,rotate_tac -1,drule atleast_free_weaken)
 11.1181 +apply   (simp (no_asm_simp))
 11.1182 +apply  (simp add: upd_gobj_def)
 11.1183 +      (* end init Ext *)
 11.1184 +apply  (rule halloc.New)
 11.1185 +apply   (drule alloc_one)
 11.1186 +prefer 2 apply fast
 11.1187 +apply   (simp (no_asm_simp))
 11.1188 +apply  (drule atleast_free_weaken)
 11.1189 +apply  force
 11.1190 +apply (simp)
 11.1191 +apply (drule alloc_one)
 11.1192 +apply  (simp (no_asm_simp))
 11.1193 +apply clarsimp
 11.1194 +apply (erule_tac V = "atleast_free ?h tree" in thin_rl)
 11.1195 +apply (drule_tac x = "a" in new_AddrD2 [THEN spec])
 11.1196 +apply (simp (no_asm_use))
 11.1197 +apply (rule eval_Is (* Try *))
 11.1198 +apply   (rule eval_Is (* Expr *))
 11.1199 +      (* begin method call *)
 11.1200 +apply   (rule eval_Is (* Call *))
 11.1201 +apply      (rule eval_Is (* Acc *))
 11.1202 +apply      (rule eval_Is (* LVar *))
 11.1203 +apply     (rule eval_Is (* Cons *))
 11.1204 +apply      (rule eval_Is (* Lit *))
 11.1205 +apply     (rule eval_Is (* Nil *))
 11.1206 +apply    (simp)
 11.1207 +apply   (simp)
 11.1208 +apply   (rule eval_Is (* Methd *))
 11.1209 +apply   (simp add: body_def Let_def)
 11.1210 +apply   (rule eval_Is (* Body *))
 11.1211 +apply     (rule init_done, simp)
 11.1212 +apply       (simp add: invocation_declclass_def dynlookup_def dynmethd_Ext_foo)
 11.1213 +apply    (simp add: invocation_declclass_def dynlookup_def dynmethd_Ext_foo)
 11.1214 +apply    (rule eval_Is (* Expr *))
 11.1215 +apply    (rule eval_Is (* Ass *))
 11.1216 +apply     (rule eval_Is (* FVar *))
 11.1217 +apply       (rule init_done, simp)
 11.1218 +apply      (rule eval_Is (* Cast *))
 11.1219 +apply       (rule eval_Is (* Acc *))
 11.1220 +apply       (rule eval_Is (* LVar *))
 11.1221 +apply      (simp)
 11.1222 +apply     (simp split del: split_if)
 11.1223 +apply    (rule eval_Is (* XcptE *))
 11.1224 +apply   (simp)
 11.1225 +      (* end method call *)
 11.1226 +apply  (rule sxalloc.intros)
 11.1227 +apply  (rule halloc.New)
 11.1228 +apply   (erule alloc_one [THEN conjunct1])
 11.1229 +apply   (simp (no_asm_simp))
 11.1230 +apply  (simp (no_asm_simp))
 11.1231 +apply (simp add: gupd_def lupd_def obj_ty_def split del: split_if)
 11.1232 +apply (drule alloc_one [THEN conjunct1])
 11.1233 +apply  (simp (no_asm_simp))
 11.1234 +apply (erule_tac V = "atleast_free ?h two" in thin_rl)
 11.1235 +apply (drule_tac x = "a" in new_AddrD2 [THEN spec])
 11.1236 +apply simp
 11.1237 +apply (rule eval_Is (* While *))
 11.1238 +apply  (rule eval_Is (* Acc *))
 11.1239 +apply  (rule eval_Is (* AVar *))
 11.1240 +apply    (rule eval_Is (* Acc *))
 11.1241 +apply    (rule eval_Is (* FVar *))
 11.1242 +apply      (rule init_done, simp)
 11.1243 +apply     (rule eval_Is (* StatRef *))
 11.1244 +apply    (simp)
 11.1245 +apply   (rule eval_Is (* Lit *))
 11.1246 +apply  (simp (no_asm_simp))
 11.1247 +apply (auto simp add: in_bounds_def)
 11.1248 +done
 11.1249 +declare Pair_eq [simp]
 11.1250 +
 11.1251 +end
    12.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    12.2 +++ b/src/HOL/Bali/Name.thy	Mon Jan 28 17:00:19 2002 +0100
    12.3 @@ -0,0 +1,123 @@
    12.4 +(*  Title:      isabelle/Bali/Name.thy
    12.5 +    ID:         $Id$
    12.6 +    Author:     David von Oheimb
    12.7 +    Copyright   1997 Technische Universitaet Muenchen
    12.8 +*)
    12.9 +header {* Java names *}
   12.10 +
   12.11 +theory Name = Basis:
   12.12 +
   12.13 +(* cf. 6.5 *) 
   12.14 +typedecl tnam		(* ordinary type name, i.e. class or interface name *)
   12.15 +typedecl pname          (* package name *)
   12.16 +typedecl mname		(* method name *)
   12.17 +typedecl vname          (* variable or field name *)
   12.18 +typedecl label          (* label as destination of break or continue *)
   12.19 +
   12.20 +arities
   12.21 +  tnam  :: "type"
   12.22 +  pname :: "type"
   12.23 +  vname :: "type"
   12.24 +  mname :: "type"
   12.25 +  label :: "type"
   12.26 +
   12.27 +
   12.28 +datatype ename        (* expression name *) 
   12.29 +        = VNam vname 
   12.30 +        | Res         (* special name to model the return value of methods *)
   12.31 +
   12.32 +datatype lname        (* names for local variables and the This pointer *)
   12.33 +        = EName ename 
   12.34 +        | This
   12.35 +syntax   
   12.36 +  VName  :: "vname \<Rightarrow> lname"
   12.37 +  Result :: lname
   12.38 +
   12.39 +translations
   12.40 +  "VName n" == "EName (VNam n)"
   12.41 +  "Result"  == "EName Res"
   12.42 +
   12.43 +datatype xname		(* names of standard exceptions *)
   12.44 +	= Throwable
   12.45 +	| NullPointer | OutOfMemory | ClassCast   
   12.46 +	| NegArrSize  | IndOutBound | ArrStore
   12.47 +
   12.48 +lemma xn_cases: 
   12.49 +  "xn = Throwable   \<or> xn = NullPointer \<or>  
   12.50 +         xn = OutOfMemory \<or> xn = ClassCast \<or> 
   12.51 +         xn = NegArrSize  \<or> xn = IndOutBound \<or> xn = ArrStore"
   12.52 +apply (induct xn)
   12.53 +apply auto
   12.54 +done
   12.55 +
   12.56 +lemma xn_cases_old:   (* FIXME cf. Example.thy *)
   12.57 +  "ALL xn. xn = Throwable   \<or> xn = NullPointer \<or>  
   12.58 +         xn = OutOfMemory \<or> xn = ClassCast \<or> 
   12.59 +         xn = NegArrSize  \<or> xn = IndOutBound \<or> xn = ArrStore"
   12.60 +apply (rule allI)
   12.61 +apply (induct_tac xn)
   12.62 +apply auto
   12.63 +done
   12.64 +
   12.65 +datatype tname	(* type names for standard classes and other type names *)
   12.66 +	= Object_
   12.67 +	| SXcpt_   xname
   12.68 +	| TName   tnam
   12.69 +
   12.70 +record   qtname = (* qualified tname cf. 6.5.3, 6.5.4*)
   12.71 +          pid :: pname 
   12.72 +          tid :: tname
   12.73 +
   12.74 +axclass has_pname < "type"
   12.75 +consts pname::"'a::has_pname \<Rightarrow> pname"
   12.76 +
   12.77 +instance pname::has_pname
   12.78 +by (intro_classes)
   12.79 +
   12.80 +defs (overloaded)
   12.81 +pname_pname_def: "pname (p::pname) \<equiv> p"
   12.82 +
   12.83 +axclass has_tname < "type"
   12.84 +consts tname::"'a::has_tname \<Rightarrow> tname"
   12.85 +
   12.86 +instance tname::has_tname
   12.87 +by (intro_classes)
   12.88 +
   12.89 +defs (overloaded)
   12.90 +tname_tname_def: "tname (t::tname) \<equiv> t"
   12.91 +
   12.92 +axclass has_qtname < "type"
   12.93 +consts qtname:: "'a::has_qtname \<Rightarrow> qtname"
   12.94 +
   12.95 +(* Declare qtname as instance of has_qtname *)
   12.96 +instance pid_field_type::(has_pname,"type") has_qtname
   12.97 +by intro_classes 
   12.98 +
   12.99 +defs (overloaded)
  12.100 +qtname_qtname_def: "qtname (q::qtname) \<equiv> q"
  12.101 +
  12.102 +translations
  12.103 +  "mname"  <= "Name.mname"
  12.104 +  "xname"  <= "Name.xname"
  12.105 +  "tname"  <= "Name.tname"
  12.106 +  "ename"  <= "Name.ename"
  12.107 +  "qtname" <= (type) "\<lparr>pid::pname,tid::tname\<rparr>"
  12.108 +  (type) "'a qtname_scheme" <= (type) "\<lparr>pid::pname,tid::tname,\<dots>::'a\<rparr>"
  12.109 +
  12.110 +
  12.111 +consts java_lang::pname (* package java.lang *)
  12.112 +
  12.113 +consts 
  12.114 +  Object :: qtname
  12.115 +  SXcpt  :: "xname \<Rightarrow> qtname"
  12.116 +defs
  12.117 +  Object_def: "Object  \<equiv> \<lparr>pid = java_lang, tid = Object_\<rparr>"
  12.118 +  SXcpt_def:  "SXcpt   \<equiv> \<lambda>x.  \<lparr>pid = java_lang, tid = SXcpt_ x\<rparr>"
  12.119 +
  12.120 +lemma Object_neq_SXcpt [simp]: "Object \<noteq> SXcpt xn"
  12.121 +by (simp add: Object_def SXcpt_def)
  12.122 +
  12.123 +lemma SXcpt_inject [simp]: "(SXcpt xn = SXcpt xm) = (xn = xm)"
  12.124 +by (simp add: SXcpt_def)
  12.125 +end
  12.126 +
    13.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    13.2 +++ b/src/HOL/Bali/ROOT.ML	Mon Jan 28 17:00:19 2002 +0100
    13.3 @@ -0,0 +1,13 @@
    13.4 +
    13.5 +use_thy "WellForm";
    13.6 +
    13.7 +(*The dynamic part of Bali, including type-safety*)
    13.8 +use_thy "Evaln";
    13.9 +use_thy "Example"; 
   13.10 +use_thy "TypeSafe";
   13.11 +(*###use_thy "Trans";*)
   13.12 +
   13.13 +(*The Hoare logic for Bali*)
   13.14 +use_thy "AxExample";
   13.15 +use_thy "AxSound";
   13.16 +use_thy "AxCompl";
    14.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    14.2 +++ b/src/HOL/Bali/State.thy	Mon Jan 28 17:00:19 2002 +0100
    14.3 @@ -0,0 +1,739 @@
    14.4 +(*  Title:      isabelle/Bali/State.thy
    14.5 +    ID:         $Id$
    14.6 +    Author:     David von Oheimb
    14.7 +    Copyright   1997 Technische Universitaet Muenchen
    14.8 +*)
    14.9 +header {* State for evaluation of Java expressions and statements *}
   14.10 +
   14.11 +theory State = DeclConcepts:
   14.12 +
   14.13 +text {*
   14.14 +design issues:
   14.15 +\begin{itemize}
   14.16 +\item all kinds of objects (class instances, arrays, and class objects)
   14.17 +  are handeled via a general object abstraction
   14.18 +\item the heap and the map for class objects are combined into a single table
   14.19 +  @{text "(recall (loc, obj) table \<times> (qtname, obj) table  ~=  (loc + qtname, obj) table)"}
   14.20 +\end{itemize}
   14.21 +*}
   14.22 +
   14.23 +section "objects"
   14.24 +
   14.25 +datatype  obj_tag =     (* tag for generic object   *)
   14.26 +	  CInst qtname   (* class instance           *)
   14.27 +	| Arr  ty int   (* array with component type and length *)
   14.28 +     (* | CStat            the tag is irrelevant for a class object,
   14.29 +			   i.e. the static fields of a class *)
   14.30 +
   14.31 +types	vn   = "fspec + int"                    (* variable name      *)
   14.32 +record	obj  = 
   14.33 +          tag :: "obj_tag"                      (* generalized object *)
   14.34 +          values :: "(vn, val) table"      
   14.35 +
   14.36 +translations 
   14.37 +  "fspec" <= (type) "vname \<times> qtname" 
   14.38 +  "vn"    <= (type) "fspec + int"
   14.39 +  "obj"   <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option\<rparr>"
   14.40 +  "obj"   <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option,\<dots>::'a\<rparr>"
   14.41 +
   14.42 +constdefs
   14.43 +  
   14.44 +  the_Arr :: "obj option \<Rightarrow> ty \<times> int \<times> (vn, val) table"
   14.45 + "the_Arr obj \<equiv> \<epsilon>(T,k,t). obj = Some \<lparr>tag=Arr T k,values=t\<rparr>"
   14.46 +
   14.47 +lemma the_Arr_Arr [simp]: "the_Arr (Some \<lparr>tag=Arr T k,values=cs\<rparr>) = (T,k,cs)"
   14.48 +apply (auto simp: the_Arr_def)
   14.49 +done
   14.50 +
   14.51 +lemma the_Arr_Arr1 [simp,intro,dest]:
   14.52 + "\<lbrakk>tag obj = Arr T k\<rbrakk> \<Longrightarrow> the_Arr (Some obj) = (T,k,values obj)"
   14.53 +apply (auto simp add: the_Arr_def)
   14.54 +done
   14.55 +
   14.56 +constdefs
   14.57 +
   14.58 +  upd_obj       :: "vn \<Rightarrow> val \<Rightarrow> obj \<Rightarrow> obj" 
   14.59 + "upd_obj n v \<equiv> \<lambda> obj . obj \<lparr>values:=(values obj)(n\<mapsto>v)\<rparr>"
   14.60 +
   14.61 +lemma upd_obj_def2 [simp]: 
   14.62 +  "upd_obj n v obj = obj \<lparr>values:=(values obj)(n\<mapsto>v)\<rparr>" 
   14.63 +apply (auto simp: upd_obj_def)
   14.64 +done
   14.65 +
   14.66 +constdefs
   14.67 +  obj_ty        :: "obj \<Rightarrow> ty"
   14.68 + "obj_ty obj    \<equiv> case tag obj of 
   14.69 +                    CInst C \<Rightarrow> Class C 
   14.70 +                  | Arr T k \<Rightarrow> T.[]"
   14.71 +
   14.72 +lemma obj_ty_eq [intro!]: "obj_ty \<lparr>tag=oi,values=x\<rparr> = obj_ty \<lparr>tag=oi,values=y\<rparr>" 
   14.73 +by (simp add: obj_ty_def)
   14.74 +
   14.75 +
   14.76 +lemma obj_ty_eq1 [intro!,dest]: 
   14.77 +  "tag obj = tag obj' \<Longrightarrow> obj_ty obj = obj_ty obj'" 
   14.78 +by (simp add: obj_ty_def)
   14.79 +
   14.80 +lemma obj_ty_cong [simp]: 
   14.81 +  "obj_ty (obj \<lparr>values:=vs\<rparr>) = obj_ty obj" 
   14.82 +by auto
   14.83 +(*
   14.84 +lemma obj_ty_cong [simp]: 
   14.85 +  "obj_ty (obj \<lparr>values:=vs(n\<mapsto>v)\<rparr>) = obj_ty (obj \<lparr>values:=vs\<rparr>)" 
   14.86 +by auto
   14.87 +*)
   14.88 +lemma obj_ty_CInst [simp]: 
   14.89 + "obj_ty \<lparr>tag=CInst C,values=vs\<rparr> = Class C" 
   14.90 +by (simp add: obj_ty_def)
   14.91 +
   14.92 +lemma obj_ty_CInst1 [simp,intro!,dest]: 
   14.93 + "\<lbrakk>tag obj = CInst C\<rbrakk> \<Longrightarrow> obj_ty obj = Class C" 
   14.94 +by (simp add: obj_ty_def)
   14.95 +
   14.96 +lemma obj_ty_Arr [simp]: 
   14.97 + "obj_ty \<lparr>tag=Arr T i,values=vs\<rparr> = T.[]"
   14.98 +by (simp add: obj_ty_def)
   14.99 +
  14.100 +lemma obj_ty_Arr1 [simp,intro!,dest]: 
  14.101 + "\<lbrakk>tag obj = Arr T i\<rbrakk> \<Longrightarrow> obj_ty obj = T.[]"
  14.102 +by (simp add: obj_ty_def)
  14.103 +
  14.104 +lemma obj_ty_widenD: 
  14.105 + "G\<turnstile>obj_ty obj\<preceq>RefT t \<Longrightarrow> (\<exists>C. tag obj = CInst C) \<or> (\<exists>T k. tag obj = Arr T k)"
  14.106 +apply (unfold obj_ty_def)
  14.107 +apply (auto split add: obj_tag.split_asm)
  14.108 +done
  14.109 +
  14.110 +constdefs
  14.111 +
  14.112 +  obj_class :: "obj \<Rightarrow> qtname"
  14.113 + "obj_class obj \<equiv> case tag obj of 
  14.114 +                    CInst C \<Rightarrow> C 
  14.115 +                  | Arr T k \<Rightarrow> Object"
  14.116 +
  14.117 +
  14.118 +lemma obj_class_CInst [simp]: "obj_class \<lparr>tag=CInst C,values=vs\<rparr> = C" 
  14.119 +by (auto simp: obj_class_def)
  14.120 +
  14.121 +lemma obj_class_CInst1 [simp,intro!,dest]: 
  14.122 +  "tag obj = CInst C \<Longrightarrow> obj_class obj = C" 
  14.123 +by (auto simp: obj_class_def)
  14.124 +
  14.125 +lemma obj_class_Arr [simp]: "obj_class \<lparr>tag=Arr T k,values=vs\<rparr> = Object" 
  14.126 +by (auto simp: obj_class_def)
  14.127 +
  14.128 +lemma obj_class_Arr1 [simp,intro!,dest]: 
  14.129 + "tag obj = Arr T k \<Longrightarrow> obj_class obj = Object" 
  14.130 +by (auto simp: obj_class_def)
  14.131 +
  14.132 +lemma obj_ty_obj_class: "G\<turnstile>obj_ty obj\<preceq> Class statC = G\<turnstile>obj_class obj \<preceq>\<^sub>C statC"
  14.133 +apply (case_tac "tag obj")
  14.134 +apply (auto simp add: obj_ty_def obj_class_def)
  14.135 +apply (case_tac "statC = Object")
  14.136 +apply (auto dest: widen_Array_Class)
  14.137 +done
  14.138 +
  14.139 +section "object references"
  14.140 +
  14.141 +types oref = "loc + qtname"          (* generalized object reference *)
  14.142 +syntax
  14.143 +  Heap  :: "loc   \<Rightarrow> oref"
  14.144 +  Stat  :: "qtname \<Rightarrow> oref"
  14.145 +
  14.146 +translations
  14.147 +  "Heap" => "Inl"
  14.148 +  "Stat" => "Inr"
  14.149 +  "oref" <= (type) "loc + qtname"
  14.150 +
  14.151 +constdefs
  14.152 +  fields_table::
  14.153 +    "prog \<Rightarrow> qtname \<Rightarrow> (fspec \<Rightarrow> field \<Rightarrow> bool)  \<Rightarrow> (fspec, ty) table"
  14.154 + "fields_table G C P 
  14.155 +    \<equiv> option_map type \<circ> table_of (filter (split P) (DeclConcepts.fields G C))"
  14.156 +
  14.157 +lemma fields_table_SomeI: 
  14.158 +"\<lbrakk>table_of (DeclConcepts.fields G C) n = Some f; P n f\<rbrakk> 
  14.159 + \<Longrightarrow> fields_table G C P n = Some (type f)"
  14.160 +apply (unfold fields_table_def)
  14.161 +apply clarsimp
  14.162 +apply (rule exI)
  14.163 +apply (rule conjI)
  14.164 +apply (erule map_of_filter_in)
  14.165 +apply assumption
  14.166 +apply simp
  14.167 +done
  14.168 +
  14.169 +(* unused *)
  14.170 +lemma fields_table_SomeD': "fields_table G C P fn = Some T \<Longrightarrow>  
  14.171 +  \<exists>f. (fn,f)\<in>set(DeclConcepts.fields G C) \<and> type f = T"
  14.172 +apply (unfold fields_table_def)
  14.173 +apply clarsimp
  14.174 +apply (drule map_of_SomeD)
  14.175 +apply auto
  14.176 +done
  14.177 +
  14.178 +lemma fields_table_SomeD: 
  14.179 +"\<lbrakk>fields_table G C P fn = Some T; unique (DeclConcepts.fields G C)\<rbrakk> \<Longrightarrow>  
  14.180 +  \<exists>f. table_of (DeclConcepts.fields G C) fn = Some f \<and> type f = T"
  14.181 +apply (unfold fields_table_def)
  14.182 +apply clarsimp
  14.183 +apply (rule exI)
  14.184 +apply (rule conjI)
  14.185 +apply (erule table_of_filter_unique_SomeD)
  14.186 +apply assumption
  14.187 +apply simp
  14.188 +done
  14.189 +
  14.190 +constdefs
  14.191 +  in_bounds :: "int \<Rightarrow> int \<Rightarrow> bool"            ("(_/ in'_bounds _)" [50, 51] 50)
  14.192 + "i in_bounds k \<equiv> 0 \<le> i \<and> i < k"
  14.193 +
  14.194 +  arr_comps :: "'a \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a option"
  14.195 + "arr_comps T k \<equiv> \<lambda>i. if i in_bounds k then Some T else None"
  14.196 +  
  14.197 +  var_tys       :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> (vn, ty) table"
  14.198 +"var_tys G oi r 
  14.199 +  \<equiv> case r of 
  14.200 +      Heap a \<Rightarrow> (case oi of 
  14.201 +                   CInst C \<Rightarrow> fields_table G C (\<lambda>n f. \<not>static f) (+) empty
  14.202 +                 | Arr T k \<Rightarrow> empty (+) arr_comps T k)
  14.203 +    | Stat C \<Rightarrow> fields_table G C (\<lambda>fn f. declclassf fn = C \<and> static f) 
  14.204 +                (+) empty"
  14.205 +
  14.206 +lemma var_tys_Some_eq: 
  14.207 + "var_tys G oi r n = Some T 
  14.208 +  = (case r of 
  14.209 +       Inl a \<Rightarrow> (case oi of  
  14.210 +                   CInst C \<Rightarrow> (\<exists>nt. n = Inl nt \<and> fields_table G C (\<lambda>n f. 
  14.211 +                               \<not>static f) nt = Some T)  
  14.212 +                 | Arr t k \<Rightarrow> (\<exists> i. n = Inr i  \<and> i in_bounds k \<and> t = T))  
  14.213 +     | Inr C \<Rightarrow> (\<exists>nt. n = Inl nt \<and> 
  14.214 +                 fields_table G C (\<lambda>fn f. declclassf fn = C \<and> static f) nt 
  14.215 +                  = Some T))"
  14.216 +apply (unfold var_tys_def arr_comps_def)
  14.217 +apply (force split add: sum.split_asm sum.split obj_tag.split)
  14.218 +done
  14.219 +
  14.220 +
  14.221 +section "stores"
  14.222 +
  14.223 +types	globs                  (* global variables: heap and static variables *)
  14.224 +	= "(oref , obj) table"
  14.225 +	heap
  14.226 +	= "(loc  , obj) table"
  14.227 +	locals
  14.228 +	= "(lname, val) table" (* local variables *)
  14.229 +
  14.230 +translations
  14.231 + "globs"  <= (type) "(oref , obj) table"
  14.232 + "heap"   <= (type) "(loc  , obj) table"
  14.233 + "locals" <= (type) "(lname, val) table"
  14.234 +
  14.235 +datatype st = (* pure state, i.e. contents of all variables *)
  14.236 +	 st globs locals
  14.237 +
  14.238 +subsection "access"
  14.239 +
  14.240 +constdefs
  14.241 +
  14.242 +  globs  :: "st \<Rightarrow> globs"
  14.243 + "globs  \<equiv> st_case (\<lambda>g l. g)"
  14.244 +  
  14.245 +  locals :: "st \<Rightarrow> locals"
  14.246 + "locals \<equiv> st_case (\<lambda>g l. l)"
  14.247 +
  14.248 +  heap   :: "st \<Rightarrow> heap"
  14.249 + "heap s \<equiv> globs s \<circ> Heap"
  14.250 +
  14.251 +
  14.252 +lemma globs_def2 [simp]: " globs (st g l) = g"
  14.253 +by (simp add: globs_def)
  14.254 +
  14.255 +lemma locals_def2 [simp]: "locals (st g l) = l"
  14.256 +by (simp add: locals_def)
  14.257 +
  14.258 +lemma heap_def2 [simp]:  "heap s a=globs s (Heap a)"
  14.259 +by (simp add: heap_def)
  14.260 +
  14.261 +
  14.262 +syntax
  14.263 +  val_this     :: "st \<Rightarrow> val"
  14.264 +  lookup_obj   :: "st \<Rightarrow> val \<Rightarrow> obj"
  14.265 +
  14.266 +translations
  14.267 + "val_this s"       == "the (locals s This)" 
  14.268 + "lookup_obj s a'"  == "the (heap s (the_Addr a'))"
  14.269 +
  14.270 +subsection "memory allocation"
  14.271 +
  14.272 +constdefs
  14.273 +  new_Addr     :: "heap \<Rightarrow> loc option"
  14.274 + "new_Addr h   \<equiv> if (\<forall>a. h a \<noteq> None) then None else Some (\<epsilon>a. h a = None)"
  14.275 +
  14.276 +lemma new_AddrD: "new_Addr h = Some a \<Longrightarrow> h a = None"
  14.277 +apply (unfold new_Addr_def)
  14.278 +apply auto
  14.279 +apply (case_tac "h (SOME a\<Colon>loc. h a = None)")
  14.280 +apply simp
  14.281 +apply (fast intro: someI2)
  14.282 +done
  14.283 +
  14.284 +lemma new_AddrD2: "new_Addr h = Some a \<Longrightarrow> \<forall>b. h b \<noteq> None \<longrightarrow> b \<noteq> a"
  14.285 +apply (drule new_AddrD)
  14.286 +apply auto
  14.287 +done
  14.288 +
  14.289 +lemma new_Addr_SomeI: "h a = None \<Longrightarrow> \<exists>b. new_Addr h = Some b \<and> h b = None"
  14.290 +apply (unfold new_Addr_def)
  14.291 +apply (frule not_Some_eq [THEN iffD2])
  14.292 +apply auto
  14.293 +apply  (drule not_Some_eq [THEN iffD2])
  14.294 +apply  auto
  14.295 +apply (fast intro!: someI2)
  14.296 +done
  14.297 +
  14.298 +
  14.299 +subsection "initialization"
  14.300 +
  14.301 +syntax
  14.302 +
  14.303 +  init_vals     :: "('a, ty) table \<Rightarrow> ('a, val) table"
  14.304 +
  14.305 +translations
  14.306 + "init_vals vs"    == "option_map default_val \<circ> vs"
  14.307 +
  14.308 +lemma init_arr_comps_base [simp]: "init_vals (arr_comps T 0) = empty"
  14.309 +apply (unfold arr_comps_def in_bounds_def)
  14.310 +apply (rule ext)
  14.311 +apply auto
  14.312 +done
  14.313 +
  14.314 +lemma init_arr_comps_step [simp]: 
  14.315 +"0 < j \<Longrightarrow> init_vals (arr_comps T  j    ) =  
  14.316 +           init_vals (arr_comps T (j - 1))(j - 1\<mapsto>default_val T)"
  14.317 +apply (unfold arr_comps_def in_bounds_def)
  14.318 +apply (rule ext)
  14.319 +apply auto
  14.320 +done
  14.321 +
  14.322 +subsection "update"
  14.323 +
  14.324 +constdefs
  14.325 +  gupd       :: "oref  \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st"        ("gupd'(_\<mapsto>_')"[10,10]1000)
  14.326 + "gupd r obj  \<equiv> st_case (\<lambda>g l. st (g(r\<mapsto>obj)) l)"
  14.327 +
  14.328 +  lupd       :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st"        ("lupd'(_\<mapsto>_')"[10,10]1000)
  14.329 + "lupd vn v   \<equiv> st_case (\<lambda>g l. st g (l(vn\<mapsto>v)))"
  14.330 +
  14.331 +  upd_gobj   :: "oref \<Rightarrow> vn \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" 
  14.332 + "upd_gobj r n v \<equiv> st_case (\<lambda>g l. st (chg_map (upd_obj n v) r g) l)"
  14.333 +
  14.334 +  set_locals  :: "locals \<Rightarrow> st \<Rightarrow> st"
  14.335 + "set_locals l \<equiv> st_case (\<lambda>g l'. st g l)"
  14.336 +  
  14.337 +  init_obj    :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> st \<Rightarrow> st"
  14.338 + "init_obj G oi r \<equiv> gupd(r\<mapsto>\<lparr>tag=oi, values=init_vals (var_tys G oi r)\<rparr>)"
  14.339 +
  14.340 +syntax
  14.341 +  init_class_obj :: "prog \<Rightarrow> qtname \<Rightarrow> st \<Rightarrow> st"
  14.342 +
  14.343 +translations
  14.344 + "init_class_obj G C" == "init_obj G arbitrary (Inr C)"
  14.345 +
  14.346 +lemma gupd_def2 [simp]: "gupd(r\<mapsto>obj) (st g l) = st (g(r\<mapsto>obj)) l"
  14.347 +apply (unfold gupd_def)
  14.348 +apply (simp (no_asm))
  14.349 +done
  14.350 +
  14.351 +lemma lupd_def2 [simp]: "lupd(vn\<mapsto>v) (st g l) = st g (l(vn\<mapsto>v))"
  14.352 +apply (unfold lupd_def)
  14.353 +apply (simp (no_asm))
  14.354 +done
  14.355 +
  14.356 +lemma globs_gupd [simp]: "globs  (gupd(r\<mapsto>obj) s) = globs s(r\<mapsto>obj)"
  14.357 +apply (induct "s")
  14.358 +by (simp add: gupd_def)
  14.359 +
  14.360 +lemma globs_lupd [simp]: "globs  (lupd(vn\<mapsto>v ) s) = globs  s"
  14.361 +apply (induct "s")
  14.362 +by (simp add: lupd_def)
  14.363 +
  14.364 +lemma locals_gupd [simp]: "locals (gupd(r\<mapsto>obj) s) = locals s"
  14.365 +apply (induct "s")
  14.366 +by (simp add: gupd_def)
  14.367 +
  14.368 +lemma locals_lupd [simp]: "locals (lupd(vn\<mapsto>v ) s) = locals s(vn\<mapsto>v )"
  14.369 +apply (induct "s")
  14.370 +by (simp add: lupd_def)
  14.371 +
  14.372 +lemma globs_upd_gobj_new [rule_format (no_asm), simp]: 
  14.373 +  "globs s r = None \<longrightarrow> globs (upd_gobj r n v s) = globs s"
  14.374 +apply (unfold upd_gobj_def)
  14.375 +apply (induct "s")
  14.376 +apply auto
  14.377 +done
  14.378 +
  14.379 +lemma globs_upd_gobj_upd [rule_format (no_asm), simp]: 
  14.380 +"globs s r=Some obj\<longrightarrow> globs (upd_gobj r n v s) = globs s(r\<mapsto>upd_obj n v obj)"
  14.381 +apply (unfold upd_gobj_def)
  14.382 +apply (induct "s")
  14.383 +apply auto
  14.384 +done
  14.385 +
  14.386 +lemma locals_upd_gobj [simp]: "locals (upd_gobj r n v s) = locals s"
  14.387 +apply (induct "s")
  14.388 +by (simp add: upd_gobj_def) 
  14.389 +
  14.390 +
  14.391 +lemma globs_init_obj [simp]: "globs (init_obj G oi r s) t =  
  14.392 +  (if t=r then Some \<lparr>tag=oi,values=init_vals (var_tys G oi r)\<rparr> else globs s t)"
  14.393 +apply (unfold init_obj_def)
  14.394 +apply (simp (no_asm))
  14.395 +done
  14.396 +
  14.397 +lemma locals_init_obj [simp]: "locals (init_obj G oi r s) = locals s"
  14.398 +by (simp add: init_obj_def)
  14.399 +  
  14.400 +lemma surjective_st [simp]: "st (globs s) (locals s) = s"
  14.401 +apply (induct "s")
  14.402 +by auto
  14.403 +
  14.404 +lemma surjective_st_init_obj: 
  14.405 + "st (globs (init_obj G oi r s)) (locals s) = init_obj G oi r s"
  14.406 +apply (subst locals_init_obj [THEN sym])
  14.407 +apply (rule surjective_st)
  14.408 +done
  14.409 +
  14.410 +lemma heap_heap_upd [simp]: 
  14.411 +  "heap (st (g(Inl a\<mapsto>obj)) l) = heap (st g l)(a\<mapsto>obj)"
  14.412 +apply (rule ext)
  14.413 +apply (simp (no_asm))
  14.414 +done
  14.415 +lemma heap_stat_upd [simp]: "heap (st (g(Inr C\<mapsto>obj)) l) = heap (st g l)"
  14.416 +apply (rule ext)
  14.417 +apply (simp (no_asm))
  14.418 +done
  14.419 +lemma heap_local_upd [simp]: "heap (st g (l(vn\<mapsto>v))) = heap (st g l)"
  14.420 +apply (rule ext)
  14.421 +apply (simp (no_asm))
  14.422 +done
  14.423 +
  14.424 +lemma heap_gupd_Heap [simp]: "heap (gupd(Heap a\<mapsto>obj) s) = heap s(a\<mapsto>obj)"
  14.425 +apply (rule ext)
  14.426 +apply (simp (no_asm))
  14.427 +done
  14.428 +lemma heap_gupd_Stat [simp]: "heap (gupd(Stat C\<mapsto>obj) s) = heap s"
  14.429 +apply (rule ext)
  14.430 +apply (simp (no_asm))
  14.431 +done
  14.432 +lemma heap_lupd [simp]: "heap (lupd(vn\<mapsto>v) s) = heap s"
  14.433 +apply (rule ext)
  14.434 +apply (simp (no_asm))
  14.435 +done
  14.436 +
  14.437 +(*
  14.438 +lemma heap_upd_gobj_Heap: "!!a. heap (upd_gobj (Heap a) n v s) = ?X"
  14.439 +apply (rule ext)
  14.440 +apply (simp (no_asm))
  14.441 +apply (case_tac "globs s (Heap a)")
  14.442 +apply  auto
  14.443 +*)
  14.444 +
  14.445 +lemma heap_upd_gobj_Stat [simp]: "heap (upd_gobj (Stat C) n v s) = heap s"
  14.446 +apply (rule ext)
  14.447 +apply (simp (no_asm))
  14.448 +apply (case_tac "globs s (Stat C)")
  14.449 +apply  auto
  14.450 +done
  14.451 +
  14.452 +lemma set_locals_def2 [simp]: "set_locals l (st g l') = st g l"
  14.453 +apply (unfold set_locals_def)
  14.454 +apply (simp (no_asm))
  14.455 +done
  14.456 +
  14.457 +lemma set_locals_id [simp]: "set_locals (locals s) s = s"
  14.458 +apply (unfold set_locals_def)
  14.459 +apply (induct_tac "s")
  14.460 +apply (simp (no_asm))
  14.461 +done
  14.462 +
  14.463 +lemma set_set_locals [simp]: "set_locals l (set_locals l' s) = set_locals l s"
  14.464 +apply (unfold set_locals_def)
  14.465 +apply (induct_tac "s")
  14.466 +apply (simp (no_asm))
  14.467 +done
  14.468 +
  14.469 +lemma locals_set_locals [simp]: "locals (set_locals l s) = l"
  14.470 +apply (unfold set_locals_def)
  14.471 +apply (induct_tac "s")
  14.472 +apply (simp (no_asm))
  14.473 +done
  14.474 +
  14.475 +lemma globs_set_locals [simp]: "globs (set_locals l s) = globs s"
  14.476 +apply (unfold set_locals_def)
  14.477 +apply (induct_tac "s")
  14.478 +apply (simp (no_asm))
  14.479 +done
  14.480 +
  14.481 +lemma heap_set_locals [simp]: "heap (set_locals l s) = heap s"
  14.482 +apply (unfold heap_def)
  14.483 +apply (induct_tac "s")
  14.484 +apply (simp (no_asm))
  14.485 +done
  14.486 +
  14.487 +
  14.488 +section "abrupt completion"
  14.489 +
  14.490 +
  14.491 +datatype xcpt        (* exception *)
  14.492 +	= Loc loc    (* location of allocated execption object *)
  14.493 +	| Std xname  (* intermediate standard exception, see Eval.thy *)
  14.494 +
  14.495 +datatype abrupt      (* abrupt completion *) 
  14.496 +        = Xcpt xcpt  (* exception *)
  14.497 +        | Jump jump  (* break, continue, return *)
  14.498 +consts
  14.499 +
  14.500 +  the_Xcpt :: "abrupt \<Rightarrow> xcpt"
  14.501 +  the_Jump :: "abrupt => jump"
  14.502 +  the_Loc  :: "xcpt \<Rightarrow> loc"
  14.503 +  the_Std  :: "xcpt \<Rightarrow> xname"
  14.504 +
  14.505 +primrec "the_Xcpt (Xcpt x) = x"
  14.506 +primrec "the_Jump (Jump j) = j"
  14.507 +primrec "the_Loc (Loc a) = a"
  14.508 +primrec "the_Std (Std x) = x"
  14.509 +
  14.510 +types
  14.511 +  abopt  = "abrupt option"
  14.512 +	
  14.513 +
  14.514 +constdefs
  14.515 +  abrupt_if    :: "bool \<Rightarrow> abopt \<Rightarrow> abopt \<Rightarrow> abopt"
  14.516 + "abrupt_if c x' x \<equiv> if c \<and> (x = None) then x' else x"
  14.517 +
  14.518 +lemma abrupt_if_True_None [simp]: "abrupt_if True x None = x"
  14.519 +by (simp add: abrupt_if_def)
  14.520 +
  14.521 +lemma abrupt_if_True_not_None [simp]: "x \<noteq> None \<Longrightarrow> abrupt_if True x y \<noteq> None"
  14.522 +by (simp add: abrupt_if_def)
  14.523 +
  14.524 +lemma abrupt_if_False [simp]: "abrupt_if False x y = y"
  14.525 +by (simp add: abrupt_if_def)
  14.526 +
  14.527 +lemma abrupt_if_Some [simp]: "abrupt_if c x (Some y) = Some y"
  14.528 +by (simp add: abrupt_if_def)
  14.529 +
  14.530 +lemma abrupt_if_not_None [simp]: "y \<noteq> None \<Longrightarrow> abrupt_if c x y = y"
  14.531 +apply (simp add: abrupt_if_def)
  14.532 +by auto
  14.533 +
  14.534 +
  14.535 +lemma split_abrupt_if: 
  14.536 +"P (abrupt_if c x' x) = 
  14.537 +      ((c \<and> x = None \<longrightarrow> P x') \<and> (\<not> (c \<and> x = None) \<longrightarrow> P x))"
  14.538 +apply (unfold abrupt_if_def)
  14.539 +apply (split split_if)
  14.540 +apply auto
  14.541 +done
  14.542 +
  14.543 +syntax
  14.544 +
  14.545 +  raise_if :: "bool \<Rightarrow> xname \<Rightarrow> abopt \<Rightarrow> abopt"
  14.546 +  np       :: "val  \<spacespace>        \<Rightarrow> abopt \<Rightarrow> abopt"
  14.547 +  check_neg:: "val  \<spacespace>        \<Rightarrow> abopt \<Rightarrow> abopt"
  14.548 +  
  14.549 +translations
  14.550 +
  14.551 + "raise_if c xn" == "abrupt_if c (Some (Xcpt (Std xn)))"
  14.552 + "np v"          == "raise_if (v = Null)      NullPointer"
  14.553 + "check_neg i'"  == "raise_if (the_Intg i'<0) NegArrSize"
  14.554 +
  14.555 +lemma raise_if_None [simp]: "(raise_if c x y = None) = (\<not>c \<and> y = None)"
  14.556 +apply (simp add: abrupt_if_def)
  14.557 +by auto
  14.558 +declare raise_if_None [THEN iffD1, dest!]
  14.559 +
  14.560 +lemma if_raise_if_None [simp]: 
  14.561 +  "((if b then y else raise_if c x y) = None) = ((c \<longrightarrow> b) \<and> y = None)"
  14.562 +apply (simp add: abrupt_if_def)
  14.563 +apply auto
  14.564 +done
  14.565 +
  14.566 +lemma raise_if_SomeD [dest!]:
  14.567 +  "raise_if c x y = Some z \<Longrightarrow> c \<and> z=(Xcpt (Std x)) \<and> y=None \<or> (y=Some z)"
  14.568 +apply (case_tac y)
  14.569 +apply (case_tac c)
  14.570 +apply (simp add: abrupt_if_def)
  14.571 +apply (simp add: abrupt_if_def)
  14.572 +apply auto
  14.573 +done
  14.574 +
  14.575 +constdefs
  14.576 +   absorb :: "jump \<Rightarrow> abopt \<Rightarrow> abopt"
  14.577 +  "absorb j a \<equiv> if a=Some (Jump j) then None else a"
  14.578 +
  14.579 +lemma absorb_SomeD [dest!]: "absorb j a = Some x \<Longrightarrow> a = Some x"
  14.580 +by (auto simp add: absorb_def)
  14.581 +
  14.582 +lemma absorb_same [simp]: "absorb j (Some (Jump j)) = None"
  14.583 +by (auto simp add: absorb_def)
  14.584 +
  14.585 +lemma absorb_other [simp]: "a \<noteq> Some (Jump j) \<Longrightarrow> absorb j a = a"
  14.586 +by (auto simp add: absorb_def)
  14.587 +
  14.588 +
  14.589 +section "full program state"
  14.590 +
  14.591 +types
  14.592 +  state = "abopt \<times> st"          (* state including exception information *)
  14.593 +
  14.594 +syntax 
  14.595 +  Norm   :: "st \<Rightarrow> state"
  14.596 +  abrupt :: "state \<Rightarrow> abopt"
  14.597 +  store  :: "state \<Rightarrow> st"
  14.598 +
  14.599 +translations
  14.600 +   
  14.601 +  "Norm s"     == "(None,s)" 
  14.602 +  "abrupt"     => "fst"
  14.603 +  "store"      => "snd"
  14.604 +  "abopt"       <= (type) "State.abrupt option"
  14.605 +  "abopt"       <= (type) "abrupt option"
  14.606 +  "state"      <= (type) "abopt \<times> State.st"
  14.607 +  "state"      <= (type) "abopt \<times> st"
  14.608 +
  14.609 +
  14.610 +
  14.611 +lemma single_stateE: "\<forall>Z. Z = (s::state) \<Longrightarrow> False"
  14.612 +apply (erule_tac x = "(Some k,y)" in all_dupE)
  14.613 +apply (erule_tac x = "(None,y)" in allE)
  14.614 +apply clarify
  14.615 +done
  14.616 +
  14.617 +lemma state_not_single: "All (op = (x::state)) \<Longrightarrow> R"
  14.618 +apply (drule_tac x = "(if abrupt x = None then Some ?x else None,?y)" in spec)
  14.619 +apply clarsimp
  14.620 +done
  14.621 +
  14.622 +constdefs
  14.623 +
  14.624 +  normal     :: "state \<Rightarrow> bool"
  14.625 + "normal \<equiv> \<lambda>s. abrupt s = None"
  14.626 +
  14.627 +lemma normal_def2 [simp]: "normal s = (abrupt s = None)"
  14.628 +apply (unfold normal_def)
  14.629 +apply (simp (no_asm))
  14.630 +done
  14.631 +
  14.632 +constdefs
  14.633 +  heap_free :: "nat \<Rightarrow> state \<Rightarrow> bool"
  14.634 + "heap_free n \<equiv> \<lambda>s. atleast_free (heap (store s)) n"
  14.635 +
  14.636 +lemma heap_free_def2 [simp]: "heap_free n s = atleast_free (heap (store s)) n"
  14.637 +apply (unfold heap_free_def)
  14.638 +apply simp
  14.639 +done
  14.640 +
  14.641 +subsection "update"
  14.642 +
  14.643 +constdefs
  14.644 + 
  14.645 +  abupd     :: "(abopt \<Rightarrow> abopt) \<Rightarrow> state \<Rightarrow> state"
  14.646 + "abupd f \<equiv> prod_fun f id"
  14.647 +
  14.648 +  supd     :: "(st \<Rightarrow> st) \<Rightarrow> state \<Rightarrow> state" 
  14.649 + "supd \<equiv> prod_fun id"
  14.650 +  
  14.651 +lemma abupd_def2 [simp]: "abupd f (x,s) = (f x,s)"
  14.652 +by (simp add: abupd_def)
  14.653 +
  14.654 +lemma abupd_abrupt_if_False [simp]: "\<And> s. abupd (abrupt_if False xo) s = s"
  14.655 +by simp
  14.656 +
  14.657 +lemma supd_def2 [simp]: "supd f (x,s) = (x,f s)"
  14.658 +by (simp add: supd_def)
  14.659 +
  14.660 +lemma supd_lupd [simp]: 
  14.661 + "\<And> s. supd (lupd vn v ) s = (abrupt s,lupd vn v (store s))"
  14.662 +apply (simp (no_asm_simp) only: split_tupled_all)
  14.663 +apply (simp (no_asm))
  14.664 +done
  14.665 +
  14.666 +
  14.667 +lemma supd_gupd [simp]: 
  14.668 + "\<And> s. supd (gupd r obj) s = (abrupt s,gupd r obj (store s))"
  14.669 +apply (simp (no_asm_simp) only: split_tupled_all)
  14.670 +apply (simp (no_asm))
  14.671 +done
  14.672 +
  14.673 +lemma supd_init_obj [simp]: 
  14.674 + "supd (init_obj G oi r) s = (abrupt s,init_obj G oi r (store s))"
  14.675 +apply (unfold init_obj_def)
  14.676 +apply (simp (no_asm))
  14.677 +done
  14.678 +
  14.679 +syntax
  14.680 +
  14.681 +  set_lvars     :: "locals \<Rightarrow> state \<Rightarrow> state"
  14.682 +  restore_lvars :: "state  \<Rightarrow> state \<Rightarrow> state"
  14.683 +  
  14.684 +translations
  14.685 +
  14.686 + "set_lvars l" == "supd (set_locals l)"
  14.687 + "restore_lvars s' s" == "set_lvars (locals (store s')) s"
  14.688 +
  14.689 +lemma set_set_lvars [simp]: "\<And> s. set_lvars l (set_lvars l' s) = set_lvars l s"
  14.690 +apply (simp (no_asm_simp) only: split_tupled_all)
  14.691 +apply (simp (no_asm))
  14.692 +done
  14.693 +
  14.694 +lemma set_lvars_id [simp]: "\<And> s. set_lvars (locals (store s)) s = s"
  14.695 +apply (simp (no_asm_simp) only: split_tupled_all)
  14.696 +apply (simp (no_asm))
  14.697 +done
  14.698 +
  14.699 +section "initialisation test"
  14.700 +
  14.701 +constdefs
  14.702 +
  14.703 +  inited   :: "qtname \<Rightarrow> globs \<Rightarrow> bool"
  14.704 + "inited C g \<equiv> g (Stat C) \<noteq> None"
  14.705 +
  14.706 +  initd    :: "qtname \<Rightarrow> state \<Rightarrow> bool"
  14.707 + "initd C \<equiv> inited C \<circ> globs \<circ> store"
  14.708 +
  14.709 +lemma not_inited_empty [simp]: "\<not>inited C empty"
  14.710 +apply (unfold inited_def)
  14.711 +apply (simp (no_asm))
  14.712 +done
  14.713 +
  14.714 +lemma inited_gupdate [simp]: "inited C (g(r\<mapsto>obj)) = (inited C g \<or> r = Stat C)"
  14.715 +apply (unfold inited_def)
  14.716 +apply (auto split add: st.split)
  14.717 +done
  14.718 +
  14.719 +lemma inited_init_class_obj [intro!]: "inited C (globs (init_class_obj G C s))"
  14.720 +apply (unfold inited_def)
  14.721 +apply (simp (no_asm))
  14.722 +done
  14.723 +
  14.724 +lemma not_initedD: "\<not> inited C g \<Longrightarrow> g (Stat C) = None"
  14.725 +apply (unfold inited_def)
  14.726 +apply (erule notnotD)
  14.727 +done
  14.728 +
  14.729 +lemma initedD: "inited C g \<Longrightarrow> \<exists> obj. g (Stat C) = Some obj"
  14.730 +apply (unfold inited_def)
  14.731 +apply auto
  14.732 +done
  14.733 +
  14.734 +lemma initd_def2 [simp]: "initd C s = inited C (globs (store s))"
  14.735 +apply (unfold initd_def)
  14.736 +apply (simp (no_asm))
  14.737 +done
  14.738 +
  14.739 +
  14.740 +
  14.741 +end
  14.742 +
    15.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    15.2 +++ b/src/HOL/Bali/Table.thy	Mon Jan 28 17:00:19 2002 +0100
    15.3 @@ -0,0 +1,445 @@
    15.4 +(*  Title:      isabelle/Bali/Table.thy
    15.5 +    ID:         $Id$
    15.6 +    Author:     David von Oheimb
    15.7 +    Copyright   1997 Technische Universitaet Muenchen
    15.8 +*)
    15.9 +header {* Abstract tables and their implementation as lists *}
   15.10 +
   15.11 +theory Table = Basis:
   15.12 +
   15.13 +text {*
   15.14 +design issues:
   15.15 +\begin{itemize}
   15.16 +\item definition of table: infinite map vs. list vs. finite set
   15.17 +      list chosen, because:
   15.18 +  \begin{itemize} 
   15.19 +  \item[+]  a priori finite
   15.20 +  \item[+]  lookup is more operational than for finite set
   15.21 +  \item[-]  not very abstract, but function table converts it to abstract 
   15.22 +            mapping
   15.23 +  \end{itemize}
   15.24 +\item coding of lookup result: Some/None vs. value/arbitrary
   15.25 +   Some/None chosen, because:
   15.26 +  \begin{itemize}
   15.27 +  \item[++] makes definedness check possible (applies also to finite set),
   15.28 +     which is important for the type standard, hiding/overriding, etc.
   15.29 +     (though it may perhaps be possible at least for the operational semantics
   15.30 +      to treat programs as infinite, i.e. where classes, fields, methods etc.
   15.31 +      of any name are considered to be defined)
   15.32 +  \item[-]  sometimes awkward case distinctions, alleviated by operator 'the'
   15.33 +  \end{itemize}
   15.34 +\end{itemize}
   15.35 +*}
   15.36 +
   15.37 +
   15.38 +types ('a, 'b) table    (* table with key type 'a and contents type 'b *)
   15.39 +      = "'a \<leadsto> 'b"
   15.40 +      ('a, 'b) tables   (* non-unique table with key 'a and contents 'b *)
   15.41 +      = "'a \<Rightarrow> 'b set"
   15.42 +
   15.43 +
   15.44 +section "map of / table of"
   15.45 +
   15.46 +syntax
   15.47 +  table_of      :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) table"    (* concrete table *)
   15.48 +
   15.49 +translations
   15.50 +  "table_of" == "map_of"
   15.51 +  
   15.52 +  (type)"'a \<leadsto> 'b"       <= (type)"'a \<Rightarrow> 'b Option.option"
   15.53 +  (type)"('a, 'b) table" <= (type)"'a \<leadsto> 'b"
   15.54 +
   15.55 +(* ### To map *)
   15.56 +lemma override_find_left[simp]:
   15.57 +"n k = None \<Longrightarrow> (m ++ n) k = m k"
   15.58 +by (simp add: override_def)
   15.59 +
   15.60 +section {* Conditional Override *}
   15.61 +constdefs
   15.62 +cond_override:: 
   15.63 +  "('b \<Rightarrow>'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b) table"
   15.64 +
   15.65 +(* when merging tables old and new, only override an entry of table old when  
   15.66 +   the condition cond holds *)
   15.67 +"cond_override cond old new \<equiv>
   15.68 + \<lambda> k.
   15.69 +  (case new k of
   15.70 +     None         \<Rightarrow> old k                       
   15.71 +   | Some new_val \<Rightarrow> (case old k of
   15.72 +                        None         \<Rightarrow> Some new_val          
   15.73 +                      | Some old_val \<Rightarrow> (if cond new_val old_val
   15.74 +                                         then Some new_val     
   15.75 +                                         else Some old_val)))"
   15.76 +
   15.77 +lemma cond_override_empty1[simp]: "cond_override c empty t = t"
   15.78 +by (simp add: cond_override_def expand_fun_eq)
   15.79 +
   15.80 +lemma cond_override_empty2[simp]: "cond_override c t empty = t"
   15.81 +by (simp add: cond_override_def expand_fun_eq)
   15.82 +
   15.83 +lemma cond_override_None[simp]:
   15.84 + "old k = None \<Longrightarrow> (cond_override c old new) k = new k"
   15.85 +by (simp add: cond_override_def)
   15.86 +
   15.87 +lemma cond_override_override:
   15.88 + "\<lbrakk>old k = Some ov;new k = Some nv; C nv ov\<rbrakk> 
   15.89 +  \<Longrightarrow> (cond_override C old new) k = Some nv"
   15.90 +by (auto simp add: cond_override_def)
   15.91 +
   15.92 +lemma cond_override_noOverride:
   15.93 + "\<lbrakk>old k = Some ov;new k = Some nv; \<not> (C nv ov)\<rbrakk> 
   15.94 +  \<Longrightarrow> (cond_override C old new) k = Some ov"
   15.95 +by (auto simp add: cond_override_def)
   15.96 +
   15.97 +lemma dom_cond_override: "dom (cond_override C s t) \<subseteq> dom s \<union> dom t"
   15.98 +by (auto simp add: cond_override_def dom_def)
   15.99 +
  15.100 +lemma finite_dom_cond_override:
  15.101 + "\<lbrakk> finite (dom s); finite (dom t) \<rbrakk> \<Longrightarrow> finite (dom (cond_override C s t))"
  15.102 +apply (rule_tac B="dom s \<union> dom t" in finite_subset)
  15.103 +apply (rule dom_cond_override)
  15.104 +by (rule finite_UnI)
  15.105 +
  15.106 +section {* Filter on Tables *}
  15.107 +
  15.108 +constdefs
  15.109 +filter_tab:: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) table \<Rightarrow> ('a, 'b) table"
  15.110 +"filter_tab c t \<equiv> \<lambda> k. (case t k of 
  15.111 +                           None   \<Rightarrow> None
  15.112 +                         | Some x \<Rightarrow> if c k x then Some x else None)"
  15.113 +
  15.114 +lemma filter_tab_empty[simp]: "filter_tab c empty = empty"
  15.115 +by (simp add: filter_tab_def empty_def)
  15.116 +
  15.117 +lemma filter_tab_True[simp]: "filter_tab (\<lambda>x y. True) t = t"
  15.118 +by (simp add: expand_fun_eq filter_tab_def)
  15.119 +
  15.120 +lemma filter_tab_False[simp]: "filter_tab (\<lambda>x y. False) t = empty"
  15.121 +by (simp add: expand_fun_eq filter_tab_def empty_def)
  15.122 +
  15.123 +lemma filter_tab_ran_subset: "ran (filter_tab c t) \<subseteq> ran t"
  15.124 +by (auto simp add: filter_tab_def ran_def)
  15.125 +
  15.126 +lemma filter_tab_range_subset: "range (filter_tab c t) \<subseteq> range t \<union> {None}"
  15.127 +apply (auto simp add: filter_tab_def)
  15.128 +apply (drule sym, blast)
  15.129 +done
  15.130 +
  15.131 +lemma finite_range_filter_tab:
  15.132 +  "finite (range t) \<Longrightarrow> finite (range (filter_tab c t))"
  15.133 +apply (rule_tac B="range t \<union> {None} " in finite_subset)
  15.134 +apply (rule filter_tab_range_subset)
  15.135 +apply (auto intro: finite_UnI)
  15.136 +done
  15.137 +
  15.138 +lemma filter_tab_SomeD[dest!]: 
  15.139 +"filter_tab c t k = Some x \<Longrightarrow> (t k = Some x) \<and> c k x"
  15.140 +by (auto simp add: filter_tab_def)
  15.141 +
  15.142 +lemma filter_tab_SomeI: "\<lbrakk>t k = Some x;C k x\<rbrakk> \<Longrightarrow>filter_tab C t k = Some x"
  15.143 +by (simp add: filter_tab_def)
  15.144 +
  15.145 +lemma filter_tab_all_True: 
  15.146 + "\<forall> k y. t k = Some y \<longrightarrow> p k y \<Longrightarrow>filter_tab p t = t"
  15.147 +apply (auto simp add: filter_tab_def expand_fun_eq)
  15.148 +done
  15.149 +
  15.150 +lemma filter_tab_all_True_Some:
  15.151 + "\<lbrakk>\<forall> k y. t k = Some y \<longrightarrow> p k y; t k = Some v\<rbrakk> \<Longrightarrow> filter_tab p t k = Some v"
  15.152 +by (auto simp add: filter_tab_def expand_fun_eq)
  15.153 +
  15.154 +lemma filter_tab_None: "t k = None \<Longrightarrow> filter_tab p t k = None"
  15.155 +apply (simp add: filter_tab_def expand_fun_eq)
  15.156 +done
  15.157 +
  15.158 +lemma filter_tab_dom_subset: "dom (filter_tab C t) \<subseteq> dom t"
  15.159 +by (auto simp add: filter_tab_def dom_def)
  15.160 +
  15.161 +lemma filter_tab_eq: "\<lbrakk>a=b\<rbrakk> \<Longrightarrow> filter_tab C a = filter_tab C b"
  15.162 +by (auto simp add: expand_fun_eq filter_tab_def)
  15.163 +
  15.164 +lemma finite_dom_filter_tab:
  15.165 +"finite (dom t) \<Longrightarrow> finite (dom (filter_tab C t))"
  15.166 +apply (rule_tac B="dom t" in finite_subset)
  15.167 +by (rule filter_tab_dom_subset)
  15.168 +
  15.169 +
  15.170 +lemma filter_tab_weaken:
  15.171 +"\<lbrakk>\<forall> a \<in> t k: \<exists> b \<in> s k: P a b; 
  15.172 +  \<And> k x y. \<lbrakk>t k = Some x;s k = Some y\<rbrakk> \<Longrightarrow> cond k x \<longrightarrow> cond k y
  15.173 + \<rbrakk> \<Longrightarrow> \<forall> a \<in> filter_tab cond t k: \<exists> b \<in> filter_tab cond s k: P a b"
  15.174 +apply (auto simp add: filter_tab_def)
  15.175 +done
  15.176 +
  15.177 +lemma cond_override_filter: 
  15.178 +  "\<lbrakk>\<And> k old new. \<lbrakk>s k = Some new; t k = Some old\<rbrakk> 
  15.179 +    \<Longrightarrow> (\<not> overC new old \<longrightarrow>  \<not> filterC k new) \<and> 
  15.180 +        (overC new old \<longrightarrow> filterC k old \<longrightarrow> filterC k new)
  15.181 +   \<rbrakk> \<Longrightarrow>
  15.182 +    cond_override overC (filter_tab filterC t) (filter_tab filterC s) 
  15.183 +    = filter_tab filterC (cond_override overC t s)"
  15.184 +by (auto simp add: expand_fun_eq cond_override_def filter_tab_def )
  15.185 +
  15.186 +section {* Misc. *}
  15.187 +
  15.188 +lemma Ball_set_table: "(\<forall> (x,y)\<in> set l. P x y) \<Longrightarrow> \<forall> x. \<forall> y\<in> map_of l x: P x y"
  15.189 +apply (erule make_imp)
  15.190 +apply (induct l)
  15.191 +apply simp
  15.192 +apply (simp (no_asm))
  15.193 +apply auto
  15.194 +done
  15.195 +
  15.196 +lemma Ball_set_tableD: 
  15.197 +  "\<lbrakk>(\<forall> (x,y)\<in> set l. P x y); x \<in> o2s (table_of l xa)\<rbrakk> \<Longrightarrow> P xa x"
  15.198 +apply (frule Ball_set_table)
  15.199 +by auto
  15.200 +
  15.201 +declare map_of_SomeD [elim]
  15.202 +
  15.203 +lemma table_of_Some_in_set:
  15.204 +"table_of l k = Some x \<Longrightarrow> (k,x) \<in> set l"
  15.205 +by auto
  15.206 +
  15.207 +lemma set_get_eq: 
  15.208 +  "unique l \<Longrightarrow> (k, the (table_of l k)) \<in> set l = (table_of l k \<noteq> None)"
  15.209 +apply safe
  15.210 +apply (fast dest!: weak_map_of_SomeI)
  15.211 +apply auto
  15.212 +done
  15.213 +
  15.214 +lemma inj_Pair_const2: "inj (\<lambda>k. (k, C))"
  15.215 +apply (rule injI)
  15.216 +apply auto
  15.217 +done
  15.218 +
  15.219 +lemma table_of_mapconst_SomeI:
  15.220 +  "\<lbrakk>table_of t k = Some y'; snd y=y'; fst y=c\<rbrakk> \<Longrightarrow>
  15.221 +   table_of (map (\<lambda>(k,x). (k,c,x)) t) k = Some y"
  15.222 +apply (induct t)
  15.223 +apply auto
  15.224 +done
  15.225 +
  15.226 +lemma table_of_mapconst_NoneI:
  15.227 +  "\<lbrakk>table_of t k = None\<rbrakk> \<Longrightarrow>
  15.228 +   table_of (map (\<lambda>(k,x). (k,c,x)) t) k = None"
  15.229 +apply (induct t)
  15.230 +apply auto
  15.231 +done
  15.232 +
  15.233 +lemmas table_of_map2_SomeI = inj_Pair_const2 [THEN map_of_mapk_SomeI, standard]
  15.234 +
  15.235 +lemma table_of_map_SomeI [rule_format (no_asm)]: "table_of t k = Some x \<longrightarrow>
  15.236 +   table_of (map (\<lambda>(k,x). (k, f x)) t) k = Some (f x)"
  15.237 +apply (induct_tac "t")
  15.238 +apply auto
  15.239 +done
  15.240 +
  15.241 +lemma table_of_remap_SomeD [rule_format (no_asm)]: 
  15.242 +  "table_of (map (\<lambda>((k,k'),x). (k,(k',x))) t) k = Some (k',x) \<longrightarrow>
  15.243 +  table_of t (k, k') = Some x"
  15.244 +apply (induct_tac "t")
  15.245 +apply  auto
  15.246 +done
  15.247 +
  15.248 +lemma table_of_mapf_Some [rule_format (no_asm)]: "\<forall>x y. f x = f y \<longrightarrow> x = y \<Longrightarrow> 
  15.249 +  table_of (map (\<lambda>(k,x). (k,f x)) t) k = Some (f x) \<longrightarrow> table_of t k = Some x"
  15.250 +apply (induct_tac "t")
  15.251 +apply  auto
  15.252 +done
  15.253 +
  15.254 +lemma table_of_mapf_SomeD [rule_format (no_asm), dest!]: 
  15.255 +"table_of (map (\<lambda>(k,x). (k, f x)) t) k = Some z \<longrightarrow> (\<exists>y\<in>table_of t k: z=f y)"
  15.256 +apply (induct_tac "t")
  15.257 +apply  auto
  15.258 +done
  15.259 +
  15.260 +lemma table_of_mapf_NoneD [rule_format (no_asm), dest!]: 
  15.261 +"table_of (map (\<lambda>(k,x). (k, f x)) t) k = None \<longrightarrow> (table_of t k = None)"
  15.262 +apply (induct_tac "t")
  15.263 +apply auto
  15.264 +done
  15.265 +
  15.266 +lemma table_of_mapkey_SomeD [rule_format (no_asm), dest!]: 
  15.267 +  "table_of (map (\<lambda>(k,x). ((k,C),x)) t) (k,D) = Some x \<longrightarrow> C = D \<and> table_of t k = Some x"
  15.268 +apply (induct_tac "t")
  15.269 +apply  auto
  15.270 +done
  15.271 +lemma table_of_mapkey_SomeD2 [rule_format (no_asm), dest!]: 
  15.272 +  "table_of (map (\<lambda>(k,x). ((k,C),x)) t) ek = Some x 
  15.273 +   \<longrightarrow> C = snd ek \<and> table_of t (fst ek) = Some x"
  15.274 +apply (induct_tac "t")
  15.275 +apply  auto
  15.276 +done
  15.277 +
  15.278 +lemma table_append_Some_iff: "table_of (xs@ys) k = Some z = 
  15.279 + (table_of xs k = Some z \<or> (table_of xs k = None \<and> table_of ys k = Some z))"
  15.280 +apply (simp only: map_of_override [THEN sym])
  15.281 +apply (rule override_Some_iff)
  15.282 +done
  15.283 +
  15.284 +lemma table_of_filter_unique_SomeD [rule_format (no_asm)]:
  15.285 +  "table_of (filter P xs) k = Some z \<Longrightarrow> unique xs \<longrightarrow> table_of xs k = Some z"
  15.286 +apply (induct xs)
  15.287 +apply (auto del: map_of_SomeD intro!: map_of_SomeD)
  15.288 +done
  15.289 +
  15.290 +
  15.291 +consts
  15.292 +  Un_tables      :: "('a, 'b) tables set \<Rightarrow> ('a, 'b) tables"
  15.293 +  overrides_t    :: "('a, 'b) tables     \<Rightarrow> ('a, 'b) tables \<Rightarrow>
  15.294 +                     ('a, 'b) tables"             (infixl "\<oplus>\<oplus>" 100)
  15.295 +  hidings_entails:: "('a, 'b) tables \<Rightarrow> ('a, 'c) tables \<Rightarrow> 
  15.296 +                     ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"   ("_ hidings _ entails _" 20)
  15.297 +  (* variant for unique table: *)
  15.298 +  hiding_entails :: "('a, 'b) table  \<Rightarrow> ('a, 'c) table  \<Rightarrow> 
  15.299 +                     ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"   ("_ hiding _ entails _"  20)
  15.300 +  (* variant for a unique table and conditional overriding: *)
  15.301 +  cond_hiding_entails :: "('a, 'b) table  \<Rightarrow> ('a, 'c) table  
  15.302 +                          \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"  
  15.303 +                          ("_ hiding _ under _ entails _"  20)
  15.304 +
  15.305 +defs
  15.306 +  Un_tables_def:       "Un_tables ts\<spacespace>\<spacespace>\<equiv> \<lambda>k. \<Union>t\<in>ts. t k"
  15.307 +  overrides_t_def:     "s \<oplus>\<oplus> t        \<equiv> \<lambda>k. if t k = {} then s k else t k"
  15.308 +  hidings_entails_def: "t hidings s entails R \<equiv> \<forall>k. \<forall>x\<in>t k. \<forall>y\<in>s k. R x y"
  15.309 +  hiding_entails_def:  "t hiding  s entails R \<equiv> \<forall>k. \<forall>x\<in>t k: \<forall>y\<in>s k: R x y"
  15.310 +  cond_hiding_entails_def: "t hiding  s under C entails R
  15.311 +                            \<equiv> \<forall>k. \<forall>x\<in>t k: \<forall>y\<in>s k: C x y \<longrightarrow> R x y"
  15.312 +
  15.313 +section "Untables"
  15.314 +
  15.315 +lemma Un_tablesI [intro]:  "\<And>x. \<lbrakk>t \<in> ts; x \<in> t k\<rbrakk> \<Longrightarrow> x \<in> Un_tables ts k"
  15.316 +apply (simp add: Un_tables_def)
  15.317 +apply auto
  15.318 +done
  15.319 +
  15.320 +lemma Un_tablesD [dest!]: "\<And>x. x \<in> Un_tables ts k \<Longrightarrow> \<exists>t. t \<in> ts \<and> x \<in> t k"
  15.321 +apply (simp add: Un_tables_def)
  15.322 +apply auto
  15.323 +done
  15.324 +
  15.325 +lemma Un_tables_empty [simp]: "Un_tables {} = (\<lambda>k. {})"
  15.326 +apply (unfold Un_tables_def)
  15.327 +apply (simp (no_asm))
  15.328 +done
  15.329 +
  15.330 +
  15.331 +section "overrides"
  15.332 +
  15.333 +lemma empty_overrides_t [simp]: "(\<lambda>k. {}) \<oplus>\<oplus> m = m"
  15.334 +apply (unfold overrides_t_def)
  15.335 +apply (simp (no_asm))
  15.336 +done
  15.337 +lemma overrides_empty_t [simp]: "m \<oplus>\<oplus> (\<lambda>k. {}) = m"
  15.338 +apply (unfold overrides_t_def)
  15.339 +apply (simp (no_asm))
  15.340 +done
  15.341 +
  15.342 +lemma overrides_t_Some_iff: 
  15.343 + "(x \<in> (s \<oplus>\<oplus> t) k) = (x \<in> t k \<or> t k = {} \<and> x \<in> s k)"
  15.344 +by (simp add: overrides_t_def)
  15.345 +
  15.346 +lemmas overrides_t_SomeD = overrides_t_Some_iff [THEN iffD1, dest!]
  15.347 +
  15.348 +lemma overrides_t_right_empty [simp]: "n k = {} \<Longrightarrow> (m \<oplus>\<oplus> n) k = m k"  
  15.349 +by (simp add: overrides_t_def)
  15.350 +
  15.351 +lemma overrides_t_find_right [simp]: "n k \<noteq> {} \<Longrightarrow> (m \<oplus>\<oplus> n) k = n k"  
  15.352 +by (simp add: overrides_t_def)
  15.353 +
  15.354 +section "hiding entails"
  15.355 +
  15.356 +lemma hiding_entailsD: 
  15.357 +  "\<lbrakk>t hiding s entails R; t k = Some x; s k = Some y\<rbrakk> \<Longrightarrow> R x y"
  15.358 +by (simp add: hiding_entails_def)
  15.359 +
  15.360 +lemma empty_hiding_entails: "empty hiding s entails R"
  15.361 +by (simp add: hiding_entails_def)
  15.362 +
  15.363 +lemma hiding_empty_entails: "t hiding empty entails R"
  15.364 +by (simp add: hiding_entails_def)
  15.365 +declare empty_hiding_entails [simp] hiding_empty_entails [simp]
  15.366 +
  15.367 +section "cond hiding entails"
  15.368 +
  15.369 +lemma cond_hiding_entailsD: 
  15.370 +"\<lbrakk>t hiding s under C entails R; t k = Some x; s k = Some y; C x y\<rbrakk> \<Longrightarrow> R x y"
  15.371 +by (simp add: cond_hiding_entails_def)
  15.372 +
  15.373 +lemma empty_cond_hiding_entails[simp]: "empty hiding s under C entails R"
  15.374 +by (simp add: cond_hiding_entails_def)
  15.375 +
  15.376 +lemma cond_hiding_empty_entails[simp]: "t hiding empty under C entails R"
  15.377 +by (simp add: cond_hiding_entails_def)
  15.378 +
  15.379 +lemma hidings_entailsD: "\<lbrakk>t hidings s entails R; x \<in> t k; y \<in> s k\<rbrakk> \<Longrightarrow> R x y"
  15.380 +by (simp add: hidings_entails_def)
  15.381 +
  15.382 +lemma hidings_empty_entails: "t hidings (\<lambda>k. {}) entails R"
  15.383 +apply (unfold hidings_entails_def)
  15.384 +apply (simp (no_asm))
  15.385 +done
  15.386 +
  15.387 +lemma empty_hidings_entails: 
  15.388 +  "(\<lambda>k. {}) hidings s entails R"apply (unfold hidings_entails_def)
  15.389 +by (simp (no_asm))
  15.390 +declare empty_hidings_entails [intro!] hidings_empty_entails [intro!]
  15.391 +
  15.392 +
  15.393 +
  15.394 +(*###TO Map?*)
  15.395 +consts
  15.396 +  atleast_free :: "('a ~=> 'b) => nat => bool"
  15.397 +primrec
  15.398 + "atleast_free m 0       = True"
  15.399 + atleast_free_Suc: 
  15.400 + "atleast_free m (Suc n) = (? a. m a = None & (!b. atleast_free (m(a|->b)) n))"
  15.401 +
  15.402 +lemma atleast_free_weaken [rule_format (no_asm)]: 
  15.403 +  "!m. atleast_free m (Suc n) \<longrightarrow> atleast_free m n"
  15.404 +apply (induct_tac "n")
  15.405 +apply (simp (no_asm))
  15.406 +apply clarify
  15.407 +apply (simp (no_asm))
  15.408 +apply (drule atleast_free_Suc [THEN iffD1])
  15.409 +apply fast
  15.410 +done
  15.411 +
  15.412 +lemma atleast_free_SucI: 
  15.413 +"[| h a = None; !obj. atleast_free (h(a|->obj)) n |] ==> atleast_free h (Suc n)"
  15.414 +by force
  15.415 +
  15.416 +declare fun_upd_apply [simp del]
  15.417 +lemma atleast_free_SucD_lemma [rule_format (no_asm)]: 
  15.418 +" !m a. m a = None --> (!c. atleast_free (m(a|->c)) n) -->  
  15.419 +  (!b d. a ~= b --> atleast_free (m(b|->d)) n)"
  15.420 +apply (induct_tac "n")
  15.421 +apply  auto
  15.422 +apply (rule_tac x = "a" in exI)
  15.423 +apply  (rule conjI)
  15.424 +apply  (force simp add: fun_upd_apply)
  15.425 +apply (erule_tac V = "m a = None" in thin_rl)
  15.426 +apply clarify
  15.427 +apply (subst fun_upd_twist)
  15.428 +apply  (erule not_sym)
  15.429 +apply (rename_tac "ba")
  15.430 +apply (drule_tac x = "ba" in spec)
  15.431 +apply clarify
  15.432 +apply (tactic "smp_tac 2 1")
  15.433 +apply (erule (1) notE impE)
  15.434 +apply (case_tac "aa = b")
  15.435 +apply fast+
  15.436 +done
  15.437 +declare fun_upd_apply [simp]
  15.438 +
  15.439 +lemma atleast_free_SucD [rule_format (no_asm)]: "atleast_free h (Suc n) ==> atleast_free (h(a|->b)) n"
  15.440 +apply auto
  15.441 +apply (case_tac "aa = a")
  15.442 +apply auto
  15.443 +apply (erule atleast_free_SucD_lemma)
  15.444 +apply auto
  15.445 +done
  15.446 +
  15.447 +declare atleast_free_Suc [simp del]
  15.448 +end
    16.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    16.2 +++ b/src/HOL/Bali/Term.thy	Mon Jan 28 17:00:19 2002 +0100
    16.3 @@ -0,0 +1,231 @@
    16.4 +(*  Title:      isabelle/Bali/Term.thy
    16.5 +    ID:         $Id$
    16.6 +    Author:     David von Oheimb
    16.7 +    Copyright   1997 Technische Universitaet Muenchen
    16.8 +*)
    16.9 +
   16.10 +header {* Java expressions and statements *}
   16.11 +
   16.12 +theory Term = Value:
   16.13 +
   16.14 +text {*
   16.15 +design issues:
   16.16 +\begin{itemize}
   16.17 +\item invocation frames for local variables could be reduced to special static
   16.18 +  objects (one per method). This would reduce redundancy, but yield a rather
   16.19 +  non-standard execution model more difficult to understand.
   16.20 +\item method bodies separated from calls to handle assumptions in axiomat. 
   16.21 +  semantics
   16.22 +  NB: Body is intended to be in the environment of the called method.
   16.23 +\item class initialization is regarded as (auxiliary) statement 
   16.24 +      (required for AxSem)
   16.25 +\item result expression of method return is handled by a special result variable
   16.26 +  result variable is treated uniformly with local variables
   16.27 +  \begin{itemize}
   16.28 +  \item[+] welltypedness and existence of the result/return expression is 
   16.29 +           ensured without extra efford
   16.30 +  \end{itemize}
   16.31 +\end{itemize}
   16.32 +
   16.33 +simplifications:
   16.34 +\begin{itemize}
   16.35 +\item expression statement allowed for any expression
   16.36 +\item no unary, binary, etc, operators
   16.37 +\item This  is modeled as a special non-assignable local variable
   16.38 +\item Super is modeled as a general expression with the same value as This
   16.39 +\item access to field x in current class via This.x
   16.40 +\item NewA creates only one-dimensional arrays;
   16.41 +  initialization of further subarrays may be simulated with nested NewAs
   16.42 +\item The 'Lit' constructor is allowed to contain a reference value.
   16.43 +  But this is assumed to be prohibited in the input language, which is enforced
   16.44 +  by the type-checking rules.
   16.45 +\item a call of a static method via a type name may be simulated by a dummy 
   16.46 +      variable
   16.47 +\item no nested blocks with inner local variables
   16.48 +\item no synchronized statements
   16.49 +\item no secondary forms of if, while (e.g. no for) (may be easily simulated)
   16.50 +\item no switch (may be simulated with if)
   16.51 +\item the @{text try_catch_finally} statement is divided into the 
   16.52 +      @{text try_catch} statement 
   16.53 +      and a finally statement, which may be considered as try..finally with 
   16.54 +      empty catch
   16.55 +\item the @{text try_catch} statement has exactly one catch clause; 
   16.56 +      multiple ones can be
   16.57 +  simulated with instanceof
   16.58 +\item the compiler is supposed to add the annotations {@{text _}} during 
   16.59 +      type-checking. This
   16.60 +  transformation is left out as its result is checked by the type rules anyway
   16.61 +\end{itemize}
   16.62 +*}
   16.63 +
   16.64 +datatype inv_mode                  (* invocation mode for method calls *)
   16.65 +	= Static                   (* static *)
   16.66 +	| SuperM                   (* super  *)
   16.67 +	| IntVir                   (* interface or virtual *)
   16.68 +
   16.69 +record  sig =            (* signature of a method, cf. 8.4.2  *)
   16.70 +	  name ::"mname"      (* acutally belongs to Decl.thy *)
   16.71 +          parTs::"ty list"        
   16.72 +
   16.73 +translations
   16.74 +  "sig" <= (type) "\<lparr>name::mname,parTs::ty list\<rparr>"
   16.75 +  "sig" <= (type) "\<lparr>name::mname,parTs::ty list,\<dots>::'a\<rparr>"
   16.76 +
   16.77 +datatype jump
   16.78 +        = Break label (* break *)
   16.79 +        | Cont label  (* continue *)
   16.80 +        | Ret         (* return from method *)
   16.81 +
   16.82 +datatype var
   16.83 +	= LVar                  lname(* local variable (incl. parameters) *)
   16.84 +        | FVar qtname bool expr vname(*class field*)("{_,_}_.._"[10,10,85,99]90)
   16.85 +	| AVar        expr expr      (* array component *) ("_.[_]"[90,10   ]90)
   16.86 +
   16.87 +and expr
   16.88 +	= NewC qtname              (* class instance creation *)
   16.89 +	| NewA ty expr             (* array creation *) ("New _[_]"[99,10   ]85)
   16.90 +	| Cast ty expr             (* type cast  *)
   16.91 +	| Inst expr ref_ty         (* instanceof *)     ("_ InstOf _"[85,99] 85)
   16.92 +	| Lit  val                 (* literal value, references not allowed *)
   16.93 +	| Super                    (* special Super keyword *)
   16.94 +	| Acc  var                 (* variable access *)
   16.95 +	| Ass  var expr            (* variable assign *) ("_:=_"   [90,85   ]85)
   16.96 +	| Cond expr expr expr      (* conditional *)  ("_ ? _ : _" [85,85,80]80)
   16.97 +        | Call ref_ty inv_mode expr mname "(ty list)" (* method call *)
   16.98 +                  "(expr list)" ("{_,_}_\<cdot>_'( {_}_')"[10,10,85,99,10,10]85)
   16.99 +        | Methd qtname sig          (*   (folded) method (see below) *)
  16.100 +        | Body qtname stmt          (* (unfolded) method body *)
  16.101 +and  stmt
  16.102 +	= Skip                     (* empty      statement *)
  16.103 +	| Expr  expr               (* expression statement *)
  16.104 +        | Lab   label stmt         ("_\<bullet> _"(* labeled statement*)[      99,66]66)
  16.105 +                                   (* handles break *)
  16.106 +	| Comp  stmt stmt          ("_;; _"                     [      66,65]65)
  16.107 +	| If_   expr stmt stmt     ("If'(_') _ Else _"          [   80,79,79]70)
  16.108 +	| Loop  label expr stmt    ("_\<bullet> While'(_') _"           [   99,80,79]70)
  16.109 +        | Do jump                  (* break, continue, return *)
  16.110 +	| Throw expr
  16.111 +        | TryC  stmt
  16.112 +	        qtname vname stmt   ("Try _ Catch'(_ _') _"     [79,99,80,79]70)
  16.113 +	| Fin   stmt stmt          ("_ Finally _"               [      79,79]70)
  16.114 +	| Init  qtname              (* class initialization *)
  16.115 +
  16.116 +
  16.117 +text {*
  16.118 +The expressions Methd and Body are artificial program constructs, in the
  16.119 +sense that they are not used to define a concrete Bali program. In the 
  16.120 +evaluation semantic definition they are "generated on the fly" to decompose 
  16.121 +the task to define the behaviour of the Call expression. They are crucial 
  16.122 +for the axiomatic semantics to give a syntactic hook to insert 
  16.123 +some assertions (cf. AxSem.thy, Eval.thy).
  16.124 +Also the Init statement (to initialize a class on its first use) is inserted 
  16.125 +in various places by the evaluation semantics.   
  16.126 +*}
  16.127 + 
  16.128 +types "term" = "(expr+stmt, var, expr list) sum3"
  16.129 +translations
  16.130 +  "sig"   <= (type) "mname \<times> ty list"
  16.131 +  "var"   <= (type) "Term.var"
  16.132 +  "expr"  <= (type) "Term.expr"
  16.133 +  "stmt"  <= (type) "Term.stmt"
  16.134 +  "term"  <= (type) "(expr+stmt, var, expr list) sum3"
  16.135 +
  16.136 +syntax
  16.137 +  
  16.138 +  this    :: expr
  16.139 +  LAcc    :: "vname \<Rightarrow>         expr" ("!!")
  16.140 +  LAss    :: "vname \<Rightarrow> expr \<Rightarrow> stmt" ("_:==_" [90,85] 85)
  16.141 +  Return  :: "expr \<Rightarrow> stmt"
  16.142 +  StatRef :: "ref_ty \<Rightarrow> expr"
  16.143 +
  16.144 +translations
  16.145 +  
  16.146 + "this"       == "Acc (LVar This)"
  16.147 + "!!v"        == "Acc (LVar (EName (VNam v)))"
  16.148 + "v:==e"      == "Expr (Ass (LVar (EName (VNam  v))) e)"
  16.149 + "Return e"   == "Expr (Ass (LVar (EName Res)) e);; Do Ret" 
  16.150 +                                                   (* Res := e;; Do Ret *)
  16.151 + "StatRef rt" == "Cast (RefT rt) (Lit Null)"
  16.152 +  
  16.153 +constdefs
  16.154 +
  16.155 +  is_stmt :: "term \<Rightarrow> bool"
  16.156 + "is_stmt t \<equiv> \<exists>c. t=In1r c"
  16.157 +
  16.158 +ML {*
  16.159 +bind_thms ("is_stmt_rews", sum3_instantiate (thm "is_stmt_def"));
  16.160 +*}
  16.161 +
  16.162 +declare is_stmt_rews [simp]
  16.163 +
  16.164 +
  16.165 +(* ############# Just testing syntax *)
  16.166 +(** unfortunately cannot simply instantiate tnam **)
  16.167 +(*
  16.168 +datatype tnam_  = HasFoo_ | Base_ | Ext_
  16.169 +datatype vnam_  = arr_ | vee_ | z_ | e_
  16.170 +datatype label_ = lab1_
  16.171 +
  16.172 +consts
  16.173 +
  16.174 +  tnam_ :: "tnam_  \<Rightarrow> tnam"
  16.175 +  vnam_ :: "vnam_  \<Rightarrow> vname"
  16.176 +  label_:: "label_ \<Rightarrow> label"
  16.177 +axioms  
  16.178 +
  16.179 +  inj_tnam_  [simp]: "(tnam_  x = tnam_  y) = (x = y)"
  16.180 +  inj_vnam_  [simp]: "(vnam_  x = vnam_  y) = (x = y)"
  16.181 +  inj_label_ [simp]: "(label_ x = label_ y) = (x = y)"
  16.182 +  
  16.183 +  
  16.184 +  surj_tnam_:  "\<exists>m. n = tnam_  m"
  16.185 +  surj_vnam_:  "\<exists>m. n = vnam_  m"
  16.186 +  surj_label_:" \<exists>m. n = label_ m"
  16.187 +
  16.188 +syntax
  16.189 +
  16.190 +  HasFoo :: qtname
  16.191 +  Base   :: qtname
  16.192 +  Ext    :: qtname
  16.193 +  arr :: ename
  16.194 +  vee :: ename
  16.195 +  z   :: ename
  16.196 +  e   :: ename
  16.197 +  lab1:: label
  16.198 +
  16.199 +consts
  16.200 +  
  16.201 +  foo    :: mname
  16.202 +translations
  16.203 +
  16.204 +  "HasFoo" == "\<lparr>pid=java_lang,tid=TName (tnam_ HasFoo_)\<rparr>"
  16.205 +  "Base"   == "\<lparr>pid=java_lang,tid=TName (tnam_ Base_)\<rparr>"
  16.206 +  "Ext"    == "\<lparr>pid=java_lang,tid=TName (tnam_ Ext_)\<rparr>"
  16.207 +  "arr"    ==        "(vnam_ arr_)"
  16.208 +  "vee"    ==        "(vnam_ vee_)"
  16.209 +  "z"      ==        "(vnam_ z_)"
  16.210 +  "e"      ==        "(vnam_ e_)"
  16.211 +  "lab1"   ==        "label_ lab1_"
  16.212 +
  16.213 +constdefs test::stmt
  16.214 +"test \<equiv>
  16.215 +(lab1@ While(Acc  
  16.216 +      (Acc ({Base,True}StatRef (ClassT Object).arr).[Lit (Intg #2)])) Skip)"
  16.217 +
  16.218 +consts
  16.219 + pTs :: "ty list"
  16.220 +   
  16.221 +constdefs 
  16.222 +
  16.223 +test1::stmt
  16.224 +"test1 \<equiv> 
  16.225 +  Expr({ClassT Base,IntVir}!!e\<cdot>foo({pTs}[Lit Null]))"
  16.226 +
  16.227 +
  16.228 +
  16.229 +constdefs test::stmt
  16.230 +"test \<equiv>
  16.231 +(lab1\<cdot> While(Acc 
  16.232 +      (Acc ({Base,True}StatRef (ClassT Object)..arr).[Lit (Intg #2)])) Skip)"
  16.233 +*)
  16.234 +end
  16.235 \ No newline at end of file
    17.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    17.2 +++ b/src/HOL/Bali/Trans.thy	Mon Jan 28 17:00:19 2002 +0100
    17.3 @@ -0,0 +1,144 @@
    17.4 +(*  Title:      isabelle/Bali/Trans.thy
    17.5 +    ID:         $Id$
    17.6 +    Author:     David von Oheimb
    17.7 +    Copyright   1997 Technische Universitaet Muenchen
    17.8 +
    17.9 +Operational transition (small-step) semantics of the 
   17.10 +execution of Java expressions and statements
   17.11 +
   17.12 +PRELIMINARY!!!!!!!!
   17.13 +
   17.14 +*)
   17.15 +
   17.16 +Trans = Eval +
   17.17 +
   17.18 +consts
   17.19 +  texpr_tstmt	:: "prog \<Rightarrow> (((expr \<times> state) \<times> (expr \<times> state)) +
   17.20 +		            ((stmt \<times> state) \<times> (stmt \<times> state))) set"
   17.21 +
   17.22 +syntax (symbols)
   17.23 +  texpr :: "[prog, expr \<times> state, expr \<times> state] \<Rightarrow> bool "("_\<turnstile>_ \<rightarrow>1 _"[61,82,82] 81)
   17.24 +  tstmt :: "[prog, stmt \<times> state, stmt \<times> state] \<Rightarrow> bool "("_\<turnstile>_ \<mapsto>1 _"[61,82,82] 81)
   17.25 +  Ref   :: "loc \<Rightarrow> expr"
   17.26 +
   17.27 +translations
   17.28 +
   17.29 +  "G\<turnstile>e_s \<rightarrow>1 ex_s'" == "Inl (e_s, ex_s') \<in> texpr_tstmt G"
   17.30 +  "G\<turnstile>s_s \<mapsto>1 s'_s'" == "Inr (s_s, s'_s') \<in> texpr_tstmt G"
   17.31 +  "Ref a" == "Lit (Addr a)"
   17.32 +
   17.33 +inductive "texpr_tstmt G" intrs 
   17.34 +
   17.35 +(* evaluation of expression *)
   17.36 +  (* cf. 15.5 *)
   17.37 +  XcptE	"\<lbrakk>\<forall>v. e \<noteq> Lit v\<rbrakk> \\<Longrightarrow>
   17.38 +				  G\<turnstile>(e,Some xc,s) \<rightarrow>1 (Lit arbitrary,Some xc,s)"
   17.39 +
   17.40 +  (* cf. 15.8.1 *)
   17.41 +  NewC	"\<lbrakk>new_Addr (heap s) = Some (a,x);
   17.42 +	  s' = assign (hupd[a\<mapsto>init_Obj G C]s) (x,s)\<rbrakk> \\<Longrightarrow>
   17.43 +				G\<turnstile>(NewC C,None,s) \<rightarrow>1 (Ref a,s')"
   17.44 +
   17.45 +  (* cf. 15.9.1 *)
   17.46 +  NewA1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
   17.47 +			      G\<turnstile>(New T[e],None,s) \<rightarrow>1 (New T[e'],s')"
   17.48 +  NewA	"\<lbrakk>i = the_Intg i'; new_Addr (heap s) = Some (a, x);
   17.49 +	  s' = assign (hupd[a\<mapsto>init_Arr T i]s)(raise_if (i<#0) NegArrSize x,s)\<rbrakk> \\<Longrightarrow>
   17.50 +			G\<turnstile>(New T[Lit i'],None,s) \<rightarrow>1 (Ref a,s')"
   17.51 +  (* cf. 15.15 *)
   17.52 +  Cast1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
   17.53 +			      G\<turnstile>(Cast T e,None,s) \<rightarrow>1 (Cast T e',s')"
   17.54 +  Cast	"\<lbrakk>x'= raise_if (\<questiondown>G,s\<turnstile>v fits T) ClassCast None\<rbrakk> \\<Longrightarrow>
   17.55 +		        G\<turnstile>(Cast T (Lit v),None,s) \<rightarrow>1 (Lit v,x',s)"
   17.56 +
   17.57 +  (* cf. 15.7.1 *)
   17.58 +(*Lit				"G\<turnstile>(Lit v,None,s) \<rightarrow>1 (Lit v,None,s)"*)
   17.59 +
   17.60 +  (* cf. 15.13.1, 15.2 *)
   17.61 +  LAcc	"\<lbrakk>v = the (locals s vn)\<rbrakk> \\<Longrightarrow>
   17.62 +			       G\<turnstile>(LAcc vn,None,s) \<rightarrow>1 (Lit v,None,s)"
   17.63 +
   17.64 +  (* cf. 15.25.1 *)
   17.65 +  LAss1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
   17.66 +				 G\<turnstile>(f vn:=e,None,s) \<rightarrow>1 (vn:=e',s')"
   17.67 +  LAss			    "G\<turnstile>(f vn:=Lit v,None,s) \<rightarrow>1 (Lit v,None,lupd[vn\<mapsto>v]s)"
   17.68 +
   17.69 +  (* cf. 15.10.1, 15.2 *)
   17.70 +  FAcc1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
   17.71 +			       G\<turnstile>({T}e..fn,None,s) \<rightarrow>1 ({T}e'..fn,s')"
   17.72 +  FAcc	"\<lbrakk>v = the (snd (the_Obj (heap s (the_Addr a'))) (fn,T))\<rbrakk> \\<Longrightarrow>
   17.73 +			  G\<turnstile>({T}Lit a'..fn,None,s) \<rightarrow>1 (Lit v,np a' None,s)"
   17.74 +
   17.75 +  (* cf. 15.25.1 *)
   17.76 +  FAss1	"\<lbrakk>G\<turnstile>(e1,None,s) \<rightarrow>1 (e1',s')\<rbrakk> \\<Longrightarrow>
   17.77 +			  G\<turnstile>(f ({T}e1..fn):=e2,None,s) \<rightarrow>1 (f({T}e1'..fn):=e2,s')"
   17.78 +  FAss2	"\<lbrakk>G\<turnstile>(e2,np a' None,s) \<rightarrow>1 (e2',s')\<rbrakk> \\<Longrightarrow>
   17.79 +		      G\<turnstile>(f({T}Lit a'..fn):=e2,None,s) \<rightarrow>1 (f({T}Lit a'..fn):=e2',s')"
   17.80 +  FAss	"\<lbrakk>a = the_Addr a'; (c,fs) = the_Obj (heap s a);
   17.81 +	  s'= assign (hupd[a\<mapsto>Obj c (fs[(fn,T)\<mapsto>v])]s) (None,s)\<rbrakk> \\<Longrightarrow>
   17.82 +		   G\<turnstile>(f({T}Lit a'..fn):=Lit v,None,s) \<rightarrow>1 (Lit v,s')"
   17.83 +
   17.84 +
   17.85 +
   17.86 +
   17.87 +
   17.88 +  (* cf. 15.12.1 *)
   17.89 +  AAcc1	"\<lbrakk>G\<turnstile>(e1,None,s) \<rightarrow>1 (e1',s')\<rbrakk> \\<Longrightarrow>
   17.90 +				G\<turnstile>(e1[e2],None,s) \<rightarrow>1 (e1'[e2],s')"
   17.91 +  AAcc2	"\<lbrakk>G\<turnstile>(e2,None,s) \<rightarrow>1 (e2',s')\<rbrakk> \\<Longrightarrow>
   17.92 +			    G\<turnstile>(Lit a'[e2],None,s) \<rightarrow>1 (Lit a'[e2'],s')"
   17.93 +  AAcc	"\<lbrakk>vo = snd (the_Arr (heap s (the_Addr a'))) (the_Intg i');
   17.94 +	  x' = raise_if (vo = None) IndOutBound (np a' None)\<rbrakk> \\<Longrightarrow>
   17.95 +			G\<turnstile>(Lit a'[Lit i'],None,s) \<rightarrow>1 (Lit (the vo),x',s)"
   17.96 +
   17.97 +
   17.98 +  (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *)
   17.99 +  Call1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
  17.100 +			  G\<turnstile>(e..mn({pT}p),None,s) \<rightarrow>1 (e'..mn({pT}p),s')"
  17.101 +  Call2	"\<lbrakk>G\<turnstile>(p,None,s) \<rightarrow>1 (p',s')\<rbrakk> \\<Longrightarrow>
  17.102 +		     G\<turnstile>(Lit a'..mn({pT}p),None,s) \<rightarrow>1 (Lit a'..mn({pT}p'),s')"
  17.103 +  Call	"\<lbrakk>a = the_Addr a'; (md,(pn,rT),lvars,blk,res) = 
  17.104 + 			   the (cmethd G (fst (the_Obj (h a))) (mn,pT))\<rbrakk> \\<Longrightarrow>
  17.105 +	    G\<turnstile>(Lit a'..mn({pT}Lit pv),None,(h,l)) \<rightarrow>1 
  17.106 +      (Body blk res l,np a' x,(h,init_vals lvars[This\<mapsto>a'][Super\<mapsto>a'][pn\<mapsto>pv]))"
  17.107 +  Body1	"\<lbrakk>G\<turnstile>(s0,None,s) \<mapsto>1 (s0',s')\<rbrakk> \\<Longrightarrow>
  17.108 +		   G\<turnstile>(Body s0    e      l,None,s) \<rightarrow>1 (Body s0'  e  l,s')"
  17.109 +  Body2	"\<lbrakk>G\<turnstile>(e ,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
  17.110 +		   G\<turnstile>(Body Skip  e      l,None,s) \<rightarrow>1 (Body Skip e' l,s')"
  17.111 +  Body		  "G\<turnstile>(Body Skip (Lit v) l,None,s) \<rightarrow>1 (Lit v,None,(heap s,l))"
  17.112 +
  17.113 +(* execution of statements *)
  17.114 +
  17.115 +  (* cf. 14.1 *)
  17.116 +  XcptS	"\<lbrakk>s0 \<noteq> Skip\<rbrakk> \\<Longrightarrow>
  17.117 +				 G\<turnstile>(s0,Some xc,s) \<mapsto>1 (Skip,Some xc,s)"
  17.118 +
  17.119 +  (* cf. 14.5 *)
  17.120 +(*Skip	 			 "G\<turnstile>(Skip,None,s) \<mapsto>1 (Skip,None,s)"*)
  17.121 +
  17.122 +  (* cf. 14.2 *)
  17.123 +  Comp1	"\<lbrakk>G\<turnstile>(s1,None,s) \<mapsto>1 (s1',s')\<rbrakk> \\<Longrightarrow>
  17.124 +			       G\<turnstile>(s1;; s2,None,s) \<mapsto>1 (s1';; s2,s')"
  17.125 +  Comp			    "G\<turnstile>(Skip;; s2,None,s) \<mapsto>1 (s2,None,s)"
  17.126 +
  17.127 +
  17.128 +
  17.129 +
  17.130 +
  17.131 +
  17.132 +  (* cf. 14.7 *)
  17.133 +  Expr1	"\<lbrakk>G\<turnstile>(e ,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
  17.134 +				G\<turnstile>(Expr e,None,s) \<mapsto>1 (Expr e',s')"
  17.135 +  Expr			 "G\<turnstile>(Expr (Lit v),None,s) \<mapsto>1 (Skip,None,s)"
  17.136 +
  17.137 +  (* cf. 14.8.2 *)
  17.138 +  If1	"\<lbrakk>G\<turnstile>(e ,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
  17.139 +		      G\<turnstile>(If(e) s1 Else s2,None,s) \<mapsto>1 (If(e') s1 Else s2,s')"
  17.140 +  If		 "G\<turnstile>(If(Lit v) s1 Else s2,None,s) \<mapsto>1 
  17.141 +		    (if the_Bool v then s1 else s2,None,s)"
  17.142 +
  17.143 +  (* cf. 14.10, 14.10.1 *)
  17.144 +  Loop			  "G\<turnstile>(While(e) s0,None,s) \<mapsto>1 
  17.145 +			     (If(e) (s0;; While(e) s0) Else Skip,None,s)"
  17.146 +
  17.147 +end
    18.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    18.2 +++ b/src/HOL/Bali/Type.thy	Mon Jan 28 17:00:19 2002 +0100
    18.3 @@ -0,0 +1,59 @@
    18.4 +(*  Title:      isabelle/Bali/Type.thy
    18.5 +    ID:         $Id$
    18.6 +    Author:     David von Oheimb
    18.7 +    Copyright   1997 Technische Universitaet Muenchen
    18.8 +*)
    18.9 +
   18.10 +header {* Java types *}
   18.11 +
   18.12 +theory Type = Name:
   18.13 +
   18.14 +text {*
   18.15 +simplifications:
   18.16 +\begin{itemize}
   18.17 +\item only the most important primitive types
   18.18 +\item the null type is regarded as reference type
   18.19 +\end{itemize}
   18.20 +*}
   18.21 +
   18.22 +datatype prim_ty       	(* primitive type, cf. 4.2 *)
   18.23 +	= Void    	(* 'result type' of void methods *)
   18.24 +	| Boolean
   18.25 +	| Integer
   18.26 +
   18.27 +
   18.28 +datatype ref_ty		(* reference type, cf. 4.3 *)
   18.29 +	= NullT		(* null type, cf. 4.1 *)
   18.30 +	| IfaceT qtname	(* interface type *)
   18.31 +	| ClassT qtname	(* class type *)
   18.32 +	| ArrayT ty	(* array type *)
   18.33 +
   18.34 +and ty	    		(* any type, cf. 4.1 *)
   18.35 +	= PrimT prim_ty	(* primitive type *)
   18.36 +	| RefT  ref_ty	(* reference type *)
   18.37 +
   18.38 +translations
   18.39 +  "prim_ty" <= (type) "Type.prim_ty"
   18.40 +  "ref_ty"  <= (type) "Type.ref_ty"
   18.41 +  "ty"      <= (type) "Type.ty"
   18.42 +
   18.43 +syntax
   18.44 +	 NT	:: "       \<spacespace> ty"
   18.45 +	 Iface	:: "qtname  \<Rightarrow> ty"
   18.46 +	 Class	:: "qtname  \<Rightarrow> ty"
   18.47 +	 Array	:: "ty     \<Rightarrow> ty"	("_.[]" [90] 90)
   18.48 +
   18.49 +translations
   18.50 +	"NT"      == "RefT   NullT"
   18.51 +	"Iface I" == "RefT (IfaceT I)"
   18.52 +	"Class C" == "RefT (ClassT C)"
   18.53 +	"T.[]"    == "RefT (ArrayT T)"
   18.54 +
   18.55 +constdefs
   18.56 +  the_Class :: "ty \<Rightarrow> qtname"
   18.57 + "the_Class T \<equiv> \<epsilon>C. T = Class C" (** primrec not possible here **)
   18.58 + 
   18.59 +lemma the_Class_eq [simp]: "the_Class (Class C)= C"
   18.60 +by (auto simp add: the_Class_def)
   18.61 +
   18.62 +end
    19.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    19.2 +++ b/src/HOL/Bali/TypeRel.thy	Mon Jan 28 17:00:19 2002 +0100
    19.3 @@ -0,0 +1,667 @@
    19.4 +(*  Title:      isabelle/Bali/TypeRel.thy
    19.5 +    ID:         $Id$
    19.6 +    Author:     David von Oheimb
    19.7 +    Copyright   1997 Technische Universitaet Muenchen
    19.8 +*)
    19.9 +header {* The relations between Java types *}
   19.10 +
   19.11 +theory TypeRel = Decl:
   19.12 +
   19.13 +text {*
   19.14 +simplifications:
   19.15 +\begin{itemize}
   19.16 +\item subinterface, subclass and widening relation includes identity
   19.17 +\end{itemize}
   19.18 +improvements over Java Specification 1.0:
   19.19 +\begin{itemize}
   19.20 +\item narrowing reference conversion also in cases where the return types of a 
   19.21 +      pair of methods common to both types are in widening (rather identity) 
   19.22 +      relation
   19.23 +\item one could add similar constraints also for other cases
   19.24 +\end{itemize}
   19.25 +design issues:
   19.26 +\begin{itemize}
   19.27 +\item the type relations do not require @{text is_type} for their arguments
   19.28 +\item the subint1 and subcls1 relations imply @{text is_iface}/@{text is_class}
   19.29 +      for their first arguments, which is required for their finiteness
   19.30 +\end{itemize}
   19.31 +*}
   19.32 +
   19.33 +consts
   19.34 +
   19.35 +(*subint1, in Decl.thy*)                     (* direct subinterface       *)
   19.36 +(*subint , by translation*)                  (* subinterface (+ identity) *)
   19.37 +(*subcls1, in Decl.thy*)                     (* direct subclass           *)
   19.38 +(*subcls , by translation*)                  (*        subclass           *)
   19.39 +(*subclseq, by translation*)                 (* subclass + identity       *)
   19.40 +  implmt1   :: "prog \<Rightarrow> (qtname \<times> qtname) set" (* direct implementation *)
   19.41 +  implmt    :: "prog \<Rightarrow> (qtname \<times> qtname) set" (*        implementation *)
   19.42 +  widen     :: "prog \<Rightarrow> (ty    \<times> ty   ) set" (*        widening       *)
   19.43 +  narrow    :: "prog \<Rightarrow> (ty    \<times> ty   ) set" (*        narrowing      *)
   19.44 +  cast     :: "prog \<Rightarrow> (ty    \<times> ty   ) set"  (*        casting        *)
   19.45 +
   19.46 +syntax
   19.47 +
   19.48 + "@subint1" :: "prog => [qtname, qtname] => bool" ("_|-_<:I1_" [71,71,71] 70)
   19.49 + "@subint"  :: "prog => [qtname, qtname] => bool" ("_|-_<=:I _"[71,71,71] 70)
   19.50 + (* Defined in Decl.thy:
   19.51 + "@subcls1" :: "prog => [qtname, qtname] => bool" ("_|-_<:C1_" [71,71,71] 70)
   19.52 + "@subclseq":: "prog => [qtname, qtname] => bool" ("_|-_<=:C _"[71,71,71] 70)
   19.53 + "@subcls"  :: "prog => [qtname, qtname] => bool" ("_|-_<:C _"[71,71,71] 70)
   19.54 + *)
   19.55 + "@implmt1" :: "prog => [qtname, qtname] => bool" ("_|-_~>1_"  [71,71,71] 70)
   19.56 + "@implmt"  :: "prog => [qtname, qtname] => bool" ("_|-_~>_"   [71,71,71] 70)
   19.57 + "@widen"   :: "prog => [ty   , ty   ] => bool" ("_|-_<=:_"  [71,71,71] 70)
   19.58 + "@narrow"  :: "prog => [ty   , ty   ] => bool" ("_|-_:>_"   [71,71,71] 70)
   19.59 + "@cast"    :: "prog => [ty   , ty   ] => bool" ("_|-_<=:? _"[71,71,71] 70)
   19.60 +
   19.61 +syntax (symbols)
   19.62 +
   19.63 +  "@subint1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>I1_"  [71,71,71] 70)
   19.64 +  "@subint"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>I _"  [71,71,71] 70)
   19.65 +  (* Defined in Decl.thy:
   19.66 +\  "@subcls1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>\<^sub>C\<^sub>1_"  [71,71,71] 70)
   19.67 +  "@subclseq":: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>\<^sub>C _"  [71,71,71] 70)
   19.68 +  "@subcls"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>\<^sub>C _"  [71,71,71] 70)
   19.69 +  *)
   19.70 +  "@implmt1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<leadsto>1_"  [71,71,71] 70)
   19.71 +  "@implmt"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<leadsto>_"   [71,71,71] 70)
   19.72 +  "@widen"   :: "prog \<Rightarrow> [ty   , ty   ] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>_"    [71,71,71] 70)
   19.73 +  "@narrow"  :: "prog \<Rightarrow> [ty   , ty   ] \<Rightarrow> bool" ("_\<turnstile>_\<succ>_"    [71,71,71] 70)
   19.74 +  "@cast"    :: "prog \<Rightarrow> [ty   , ty   ] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>? _"  [71,71,71] 70)
   19.75 +
   19.76 +translations
   19.77 +
   19.78 +	"G\<turnstile>I \<prec>I1 J" == "(I,J) \<in> subint1 G"
   19.79 +	"G\<turnstile>I \<preceq>I  J" == "(I,J) \<in>(subint1 G)^*" (* cf. 9.1.3 *)
   19.80 +  	(* Defined in Decl.thy:
   19.81 +        "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" == "(C,D) \<in> subcls1 G"
   19.82 +	"G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" 
   19.83 +	*)
   19.84 +        "G\<turnstile>C \<leadsto>1 I" == "(C,I) \<in> implmt1 G"
   19.85 +	"G\<turnstile>C \<leadsto>  I" == "(C,I) \<in> implmt  G"
   19.86 +	"G\<turnstile>S \<preceq>   T" == "(S,T) \<in> widen   G"
   19.87 +	"G\<turnstile>S \<succ>   T" == "(S,T) \<in> narrow  G"
   19.88 +	"G\<turnstile>S \<preceq>?  T" == "(S,T) \<in> cast    G"
   19.89 +
   19.90 +
   19.91 +section "subclass and subinterface relations"
   19.92 +
   19.93 +  (* direct subinterface in Decl.thy, cf. 9.1.3 *)
   19.94 +  (* direct subclass     in Decl.thy, cf. 8.1.3 *)
   19.95 +
   19.96 +lemmas subcls_direct = subcls1I [THEN r_into_rtrancl, standard]
   19.97 +
   19.98 +lemma subcls_direct1:
   19.99 + "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<preceq>\<^sub>C D"
  19.100 +apply (auto dest: subcls_direct)
  19.101 +done
  19.102 +
  19.103 +lemma subcls1I1:
  19.104 + "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>\<^sub>C\<^sub>1 D"
  19.105 +apply (auto dest: subcls1I)
  19.106 +done
  19.107 +
  19.108 +lemma subcls_direct2:
  19.109 + "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>\<^sub>C D"
  19.110 +apply (auto dest: subcls1I1)
  19.111 +done
  19.112 +
  19.113 +lemma subclseq_trans: "\<lbrakk>G\<turnstile>A \<preceq>\<^sub>C B; G\<turnstile>B \<preceq>\<^sub>C C\<rbrakk> \<Longrightarrow> G\<turnstile>A \<preceq>\<^sub>C C"
  19.114 +by (blast intro: rtrancl_trans)
  19.115 +
  19.116 +lemma subcls_trans: "\<lbrakk>G\<turnstile>A \<prec>\<^sub>C B; G\<turnstile>B \<prec>\<^sub>C C\<rbrakk> \<Longrightarrow> G\<turnstile>A \<prec>\<^sub>C C"
  19.117 +by (blast intro: trancl_trans)
  19.118 +
  19.119 +lemma SXcpt_subcls_Throwable_lemma: 
  19.120 +"\<lbrakk>class G (SXcpt xn) = Some xc;
  19.121 +  super xc = (if xn = Throwable then Object else  SXcpt Throwable)\<rbrakk> 
  19.122 +\<Longrightarrow> G\<turnstile>SXcpt xn\<preceq>\<^sub>C SXcpt Throwable"
  19.123 +apply (case_tac "xn = Throwable")
  19.124 +apply  simp_all
  19.125 +apply (drule subcls_direct)
  19.126 +apply (auto dest: sym)
  19.127 +done
  19.128 +
  19.129 +lemma subcls_ObjectI: "\<lbrakk>is_class G C; ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>C\<preceq>\<^sub>C Object"
  19.130 +apply (erule ws_subcls1_induct)
  19.131 +apply clarsimp
  19.132 +apply (case_tac "C = Object")
  19.133 +apply  (fast intro: r_into_rtrancl [THEN rtrancl_trans])+
  19.134 +done
  19.135 +
  19.136 +lemma subclseq_ObjectD [dest!]: "G\<turnstile>Object\<preceq>\<^sub>C C \<Longrightarrow> C = Object"
  19.137 +apply (erule rtrancl_induct)
  19.138 +apply  (auto dest: subcls1D)
  19.139 +done
  19.140 +
  19.141 +lemma subcls_ObjectD [dest!]: "G\<turnstile>Object\<prec>\<^sub>C C \<Longrightarrow> False"
  19.142 +apply (erule trancl_induct)
  19.143 +apply  (auto dest: subcls1D)
  19.144 +done
  19.145 +
  19.146 +lemma subcls_ObjectI1 [intro!]: 
  19.147 + "\<lbrakk>C \<noteq> Object;is_class G C;ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>C \<prec>\<^sub>C Object" 
  19.148 +apply (drule (1) subcls_ObjectI)
  19.149 +apply (auto intro: rtrancl_into_trancl3)
  19.150 +done
  19.151 +
  19.152 +lemma subcls_is_class: "(C,D) \<in> (subcls1 G)^+ \<Longrightarrow> is_class G C"
  19.153 +apply (erule trancl_trans_induct)
  19.154 +apply (auto dest!: subcls1D)
  19.155 +done
  19.156 +
  19.157 +lemma subcls_is_class2 [rule_format (no_asm)]: 
  19.158 + "G\<turnstile>C\<preceq>\<^sub>C D \<Longrightarrow> is_class G D \<longrightarrow> is_class G C"
  19.159 +apply (erule rtrancl_induct)
  19.160 +apply (drule_tac [2] subcls1D)
  19.161 +apply  auto
  19.162 +done
  19.163 +
  19.164 +lemma single_inheritance: 
  19.165 +"\<lbrakk>G\<turnstile>A \<prec>\<^sub>C\<^sub>1 B; G\<turnstile>A \<prec>\<^sub>C\<^sub>1 C\<rbrakk> \<Longrightarrow> B = C"
  19.166 +by (auto simp add: subcls1_def)
  19.167 +  
  19.168 +lemma subcls_compareable:
  19.169 +"\<lbrakk>G\<turnstile>A \<preceq>\<^sub>C X; G\<turnstile>A \<preceq>\<^sub>C Y 
  19.170 + \<rbrakk> \<Longrightarrow> G\<turnstile>X \<preceq>\<^sub>C Y \<or> G\<turnstile>Y \<preceq>\<^sub>C X"
  19.171 +by (rule triangle_lemma)  (auto intro: single_inheritance) 
  19.172 +
  19.173 +lemma subcls1_irrefl: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D; ws_prog G \<rbrakk>
  19.174 + \<Longrightarrow> C \<noteq> D"
  19.175 +proof 
  19.176 +  assume ws: "ws_prog G" and
  19.177 +    subcls1: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" and
  19.178 +     eq_C_D: "C=D"
  19.179 +  from subcls1 obtain c 
  19.180 +    where
  19.181 +      neq_C_Object: "C\<noteq>Object" and
  19.182 +              clsC: "class G C = Some c" and
  19.183 +           super_c: "super c = D"
  19.184 +    by (auto simp add: subcls1_def)
  19.185 +  with super_c subcls1 eq_C_D 
  19.186 +  have subcls_super_c_C: "G\<turnstile>super c \<prec>\<^sub>C C"
  19.187 +    by auto
  19.188 +  from ws clsC neq_C_Object 
  19.189 +  have "\<not> G\<turnstile>super c \<prec>\<^sub>C C"
  19.190 +    by (auto dest: ws_prog_cdeclD)
  19.191 +  from this subcls_super_c_C 
  19.192 +  show "False"
  19.193 +    by (rule notE)
  19.194 +qed
  19.195 +
  19.196 +lemma no_subcls_Object: "G\<turnstile>C \<prec>\<^sub>C D \<Longrightarrow> C \<noteq> Object"
  19.197 +by (erule converse_trancl_induct) (auto dest: subcls1D)
  19.198 +
  19.199 +lemma subcls_acyclic: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D; ws_prog G\<rbrakk> \<Longrightarrow> \<not> G\<turnstile>D \<prec>\<^sub>C C"
  19.200 +proof -
  19.201 +  assume         ws: "ws_prog G"
  19.202 +  assume subcls_C_D: "G\<turnstile>C \<prec>\<^sub>C D"
  19.203 +  then show ?thesis
  19.204 +  proof (induct rule: converse_trancl_induct)
  19.205 +    fix C
  19.206 +    assume subcls1_C_D: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D"
  19.207 +    then obtain c  where
  19.208 +        "C\<noteq>Object" and
  19.209 +        "class G C = Some c" and
  19.210 +        "super c = D"
  19.211 +      by (auto simp add: subcls1_def)
  19.212 +    with ws 
  19.213 +    show "\<not> G\<turnstile>D \<prec>\<^sub>C C"
  19.214 +      by (auto dest: ws_prog_cdeclD)
  19.215 +  next
  19.216 +    fix C Z
  19.217 +    assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 Z" and
  19.218 +            subcls_Z_D: "G\<turnstile>Z \<prec>\<^sub>C D" and
  19.219 +           nsubcls_D_Z: "\<not> G\<turnstile>D \<prec>\<^sub>C Z"
  19.220 +    show "\<not> G\<turnstile>D \<prec>\<^sub>C C"
  19.221 +    proof
  19.222 +      assume subcls_D_C: "G\<turnstile>D \<prec>\<^sub>C C"
  19.223 +      show "False"
  19.224 +      proof -
  19.225 +	from subcls_D_C subcls1_C_Z
  19.226 +	have "G\<turnstile>D \<prec>\<^sub>C Z"
  19.227 +	  by (auto dest: r_into_trancl trancl_trans)
  19.228 +	with nsubcls_D_Z
  19.229 +	show ?thesis
  19.230 +	  by (rule notE)
  19.231 +      qed
  19.232 +    qed
  19.233 +  qed  
  19.234 +qed
  19.235 +
  19.236 +lemma subclseq_cases [consumes 1, case_names Eq Subcls]:
  19.237 + "\<lbrakk>G\<turnstile>C \<preceq>\<^sub>C D; C = D \<Longrightarrow> P; G\<turnstile>C \<prec>\<^sub>C D \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
  19.238 +by (blast intro: rtrancl_cases)
  19.239 +
  19.240 +lemma subclseq_acyclic: 
  19.241 + "\<lbrakk>G\<turnstile>C \<preceq>\<^sub>C D; G\<turnstile>D \<preceq>\<^sub>C C; ws_prog G\<rbrakk> \<Longrightarrow> C=D"
  19.242 +by (auto elim: subclseq_cases dest: subcls_acyclic)
  19.243 +
  19.244 +lemma subcls_irrefl: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D; ws_prog G\<rbrakk>
  19.245 + \<Longrightarrow> C \<noteq> D"
  19.246 +proof -
  19.247 +  assume     ws: "ws_prog G"
  19.248 +  assume subcls: "G\<turnstile>C \<prec>\<^sub>C D"
  19.249 +  then show ?thesis
  19.250 +  proof (induct rule: converse_trancl_induct)
  19.251 +    fix C
  19.252 +    assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D"
  19.253 +    with ws 
  19.254 +    show "C\<noteq>D" 
  19.255 +      by (blast dest: subcls1_irrefl)
  19.256 +  next
  19.257 +    fix C Z
  19.258 +    assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 Z" and
  19.259 +            subcls_Z_D: "G\<turnstile>Z \<prec>\<^sub>C D" and
  19.260 +               neq_Z_D: "Z \<noteq> D"
  19.261 +    show "C\<noteq>D"
  19.262 +    proof
  19.263 +      assume eq_C_D: "C=D"
  19.264 +      show "False"
  19.265 +      proof -
  19.266 +	from subcls1_C_Z eq_C_D 
  19.267 +	have "G\<turnstile>D \<prec>\<^sub>C Z"
  19.268 +	  by (auto)
  19.269 +	also
  19.270 +	from subcls_Z_D ws   
  19.271 +	have "\<not> G\<turnstile>D \<prec>\<^sub>C Z"
  19.272 +	  by (rule subcls_acyclic)
  19.273 +	ultimately 
  19.274 +	show ?thesis 
  19.275 +	  by - (rule notE)
  19.276 +      qed
  19.277 +    qed
  19.278 +  qed
  19.279 +qed
  19.280 +
  19.281 +lemma invert_subclseq:
  19.282 +"\<lbrakk>G\<turnstile>C \<preceq>\<^sub>C D; ws_prog G\<rbrakk>
  19.283 + \<Longrightarrow> \<not> G\<turnstile>D \<prec>\<^sub>C C"
  19.284 +proof -
  19.285 +  assume       ws: "ws_prog G" and
  19.286 +     subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D"
  19.287 +  show ?thesis
  19.288 +  proof (cases "D=C")
  19.289 +    case True
  19.290 +    with ws 
  19.291 +    show ?thesis 
  19.292 +      by (auto dest: subcls_irrefl)
  19.293 +  next
  19.294 +    case False
  19.295 +    with subclseq_C_D 
  19.296 +    have "G\<turnstile>C \<prec>\<^sub>C D"
  19.297 +      by (blast intro: rtrancl_into_trancl3) 
  19.298 +    with ws 
  19.299 +    show ?thesis 
  19.300 +      by (blast dest: subcls_acyclic)
  19.301 +  qed
  19.302 +qed
  19.303 +
  19.304 +lemma invert_subcls:
  19.305 +"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D; ws_prog G\<rbrakk>
  19.306 + \<Longrightarrow> \<not> G\<turnstile>D \<preceq>\<^sub>C C"
  19.307 +proof -
  19.308 +  assume        ws: "ws_prog G" and
  19.309 +        subcls_C_D: "G\<turnstile>C \<prec>\<^sub>C D"
  19.310 +  then 
  19.311 +  have nsubcls_D_C: "\<not> G\<turnstile>D \<prec>\<^sub>C C"
  19.312 +    by (blast dest: subcls_acyclic)
  19.313 +  show ?thesis
  19.314 +  proof
  19.315 +    assume "G\<turnstile>D \<preceq>\<^sub>C C"
  19.316 +    then show "False"
  19.317 +    proof (cases rule: subclseq_cases)
  19.318 +      case Eq
  19.319 +      with ws subcls_C_D
  19.320 +      show ?thesis 
  19.321 +	by (auto dest: subcls_irrefl)
  19.322 +    next
  19.323 +      case Subcls
  19.324 +      with nsubcls_D_C
  19.325 +      show ?thesis
  19.326 +	by blast
  19.327 +    qed
  19.328 +  qed
  19.329 +qed
  19.330 +
  19.331 +lemma subcls_superD:
  19.332 + "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D; class G C = Some c\<rbrakk> \<Longrightarrow> G\<turnstile>(super c) \<preceq>\<^sub>C D"
  19.333 +proof -
  19.334 +  assume       clsC: "class G C = Some c"
  19.335 +  assume subcls_C_C: "G\<turnstile>C \<prec>\<^sub>C D"
  19.336 +  then obtain S where 
  19.337 +                  "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" and
  19.338 +    subclseq_S_D: "G\<turnstile>S \<preceq>\<^sub>C D"
  19.339 +    by (blast dest: tranclD)
  19.340 +  with clsC 
  19.341 +  have "S=super c" 
  19.342 +    by (auto dest: subcls1D)
  19.343 +  with subclseq_S_D show ?thesis by simp
  19.344 +qed
  19.345 + 
  19.346 +
  19.347 +lemma subclseq_superD:
  19.348 + "\<lbrakk>G\<turnstile>C \<preceq>\<^sub>C D; C\<noteq>D;class G C = Some c\<rbrakk> \<Longrightarrow> G\<turnstile>(super c) \<preceq>\<^sub>C D"
  19.349 +proof -
  19.350 +  assume neq_C_D: "C\<noteq>D"
  19.351 +  assume    clsC: "class G C = Some c"
  19.352 +  assume subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" 
  19.353 +  then show ?thesis
  19.354 +  proof (cases rule: subclseq_cases)
  19.355 +    case Eq with neq_C_D show ?thesis by contradiction
  19.356 +  next
  19.357 +    case Subcls
  19.358 +    with clsC show ?thesis by (blast dest: subcls_superD)
  19.359 +  qed
  19.360 +qed
  19.361 +
  19.362 +section "implementation relation"
  19.363 +
  19.364 +defs
  19.365 +  (* direct implementation, cf. 8.1.3 *)
  19.366 +implmt1_def:"implmt1 G\<equiv>{(C,I). C\<noteq>Object \<and> (\<exists>c\<in>class G C: I\<in>set (superIfs c))}"
  19.367 +
  19.368 +lemma implmt1D: "G\<turnstile>C\<leadsto>1I \<Longrightarrow> C\<noteq>Object \<and> (\<exists>c\<in>class G C: I\<in>set (superIfs c))"
  19.369 +apply (unfold implmt1_def)
  19.370 +apply auto
  19.371 +done
  19.372 +
  19.373 +
  19.374 +inductive "implmt G" intros                    (* cf. 8.1.4 *)
  19.375 +
  19.376 +  direct:         "G\<turnstile>C\<leadsto>1J     \<spacespace>\<spacespace>     \<Longrightarrow> G\<turnstile>C\<leadsto>J"
  19.377 +  subint:        "\<lbrakk>G\<turnstile>C\<leadsto>1I; G\<turnstile>I\<preceq>I J\<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
  19.378 +  subcls1:       "\<lbrakk>G\<turnstile>C\<prec>\<^sub>C\<^sub>1D; G\<turnstile>D\<leadsto>J \<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
  19.379 +
  19.380 +lemma implmtD: "G\<turnstile>C\<leadsto>J \<Longrightarrow> (\<exists>I. G\<turnstile>C\<leadsto>1I \<and> G\<turnstile>I\<preceq>I J) \<or> (\<exists>D. G\<turnstile>C\<prec>\<^sub>C\<^sub>1D \<and> G\<turnstile>D\<leadsto>J)" 
  19.381 +apply (erule implmt.induct)
  19.382 +apply fast+
  19.383 +done
  19.384 +
  19.385 +lemma implmt_ObjectE [elim!]: "G\<turnstile>Object\<leadsto>I \<Longrightarrow> R"
  19.386 +by (auto dest!: implmtD implmt1D subcls1D)
  19.387 +
  19.388 +lemma subcls_implmt [rule_format (no_asm)]: "G\<turnstile>A\<preceq>\<^sub>C B \<Longrightarrow> G\<turnstile>B\<leadsto>K \<longrightarrow> G\<turnstile>A\<leadsto>K"
  19.389 +apply (erule rtrancl_induct)
  19.390 +apply  (auto intro: implmt.subcls1)
  19.391 +done
  19.392 +
  19.393 +lemma implmt_subint2: "\<lbrakk> G\<turnstile>A\<leadsto>J; G\<turnstile>J\<preceq>I K\<rbrakk> \<Longrightarrow> G\<turnstile>A\<leadsto>K"
  19.394 +apply (erule make_imp, erule implmt.induct)
  19.395 +apply (auto dest: implmt.subint rtrancl_trans implmt.subcls1)
  19.396 +done
  19.397 +
  19.398 +lemma implmt_is_class: "G\<turnstile>C\<leadsto>I \<Longrightarrow> is_class G C"
  19.399 +apply (erule implmt.induct)
  19.400 +apply (blast dest: implmt1D subcls1D)+
  19.401 +done
  19.402 +
  19.403 +
  19.404 +section "widening relation"
  19.405 +
  19.406 +inductive "widen G" intros(*widening, viz. method invocation conversion, cf. 5.3
  19.407 +			    i.e. kind of syntactic subtyping *)
  19.408 +  refl:    "G\<turnstile>T\<preceq>T"(*identity conversion, cf. 5.1.1 *)
  19.409 +  subint:  "G\<turnstile>I\<preceq>I J  \<Longrightarrow> G\<turnstile>Iface I\<preceq> Iface J"(*wid.ref.conv.,cf. 5.1.4 *)
  19.410 +  int_obj: "G\<turnstile>Iface I\<preceq> Class Object"
  19.411 +  subcls:  "G\<turnstile>C\<preceq>\<^sub>C D  \<Longrightarrow> G\<turnstile>Class C\<preceq> Class D"
  19.412 +  implmt:  "G\<turnstile>C\<leadsto>I   \<Longrightarrow> G\<turnstile>Class C\<preceq> Iface I"
  19.413 +  null:    "G\<turnstile>NT\<preceq> RefT R"
  19.414 +  arr_obj: "G\<turnstile>T.[]\<preceq> Class Object"
  19.415 +  array:   "G\<turnstile>RefT S\<preceq>RefT T \<Longrightarrow> G\<turnstile>RefT S.[]\<preceq> RefT T.[]"
  19.416 +
  19.417 +declare widen.refl [intro!]
  19.418 +declare widen.intros [simp]
  19.419 +
  19.420 +(* too strong in general:
  19.421 +lemma widen_PrimT: "G\<turnstile>PrimT x\<preceq>T \<Longrightarrow> T = PrimT x"
  19.422 +*)
  19.423 +lemma widen_PrimT: "G\<turnstile>PrimT x\<preceq>T \<Longrightarrow> (\<exists>y. T = PrimT y)"
  19.424 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.425 +by auto
  19.426 +
  19.427 +(* too strong in general:
  19.428 +lemma widen_PrimT2: "G\<turnstile>S\<preceq>PrimT x \<Longrightarrow> S = PrimT x"
  19.429 +*)
  19.430 +lemma widen_PrimT2: "G\<turnstile>S\<preceq>PrimT x \<Longrightarrow> \<exists>y. S = PrimT y"
  19.431 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.432 +by auto
  19.433 +
  19.434 +lemma widen_RefT: "G\<turnstile>RefT R\<preceq>T \<Longrightarrow> \<exists>t. T=RefT t"
  19.435 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.436 +by auto
  19.437 +
  19.438 +lemma widen_RefT2: "G\<turnstile>S\<preceq>RefT R \<Longrightarrow> \<exists>t. S=RefT t"
  19.439 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.440 +by auto
  19.441 +
  19.442 +lemma widen_Iface: "G\<turnstile>Iface I\<preceq>T \<Longrightarrow> T=Class Object \<or> (\<exists>J. T=Iface J)"
  19.443 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.444 +by auto
  19.445 +
  19.446 +lemma widen_Iface2: "G\<turnstile>S\<preceq> Iface J \<Longrightarrow> S = NT \<or> (\<exists>I. S = Iface I) \<or> (\<exists>D. S = Class D)"
  19.447 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.448 +by auto
  19.449 +
  19.450 +lemma widen_Iface_Iface: "G\<turnstile>Iface I\<preceq> Iface J \<Longrightarrow> G\<turnstile>I\<preceq>I J"
  19.451 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.452 +by auto
  19.453 +
  19.454 +lemma widen_Iface_Iface_eq [simp]: "G\<turnstile>Iface I\<preceq> Iface J = G\<turnstile>I\<preceq>I J"
  19.455 +apply (rule iffI)
  19.456 +apply  (erule widen_Iface_Iface)
  19.457 +apply (erule widen.subint)
  19.458 +done
  19.459 +
  19.460 +lemma widen_Class: "G\<turnstile>Class C\<preceq>T \<Longrightarrow>  (\<exists>D. T=Class D) \<or> (\<exists>I. T=Iface I)"
  19.461 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.462 +by auto
  19.463 +
  19.464 +lemma widen_Class2: "G\<turnstile>S\<preceq> Class C \<Longrightarrow> C = Object \<or> S = NT \<or> (\<exists>D. S = Class D)"
  19.465 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.466 +by auto
  19.467 +
  19.468 +lemma widen_Class_Class: "G\<turnstile>Class C\<preceq> Class cm \<Longrightarrow> G\<turnstile>C\<preceq>\<^sub>C cm"
  19.469 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.470 +by auto
  19.471 +
  19.472 +lemma widen_Class_Class_eq [simp]: "G\<turnstile>Class C\<preceq> Class cm = G\<turnstile>C\<preceq>\<^sub>C cm"
  19.473 +apply (rule iffI)
  19.474 +apply  (erule widen_Class_Class)
  19.475 +apply (erule widen.subcls)
  19.476 +done
  19.477 +
  19.478 +lemma widen_Class_Iface: "G\<turnstile>Class C\<preceq> Iface I \<Longrightarrow> G\<turnstile>C\<leadsto>I"
  19.479 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.480 +by auto
  19.481 +
  19.482 +lemma widen_Class_Iface_eq [simp]: "G\<turnstile>Class C\<preceq> Iface I = G\<turnstile>C\<leadsto>I"
  19.483 +apply (rule iffI)
  19.484 +apply  (erule widen_Class_Iface)
  19.485 +apply (erule widen.implmt)
  19.486 +done
  19.487 +
  19.488 +lemma widen_Array: "G\<turnstile>S.[]\<preceq>T \<Longrightarrow> T=Class Object \<or> (\<exists>T'. T=T'.[] \<and> G\<turnstile>S\<preceq>T')"
  19.489 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.490 +by auto
  19.491 +
  19.492 +lemma widen_Array2: "G\<turnstile>S\<preceq>T.[] \<Longrightarrow> S = NT \<or> (\<exists>S'. S=S'.[] \<and> G\<turnstile>S'\<preceq>T)"
  19.493 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.494 +by auto
  19.495 +
  19.496 +
  19.497 +lemma widen_ArrayPrimT: "G\<turnstile>PrimT t.[]\<preceq>T \<Longrightarrow> T=Class Object \<or> T=PrimT t.[]"
  19.498 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.499 +by auto
  19.500 +
  19.501 +lemma widen_ArrayRefT: 
  19.502 +  "G\<turnstile>RefT t.[]\<preceq>T \<Longrightarrow> T=Class Object \<or> (\<exists>s. T=RefT s.[] \<and> G\<turnstile>RefT t\<preceq>RefT s)"
  19.503 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.504 +by auto
  19.505 +
  19.506 +lemma widen_ArrayRefT_ArrayRefT_eq [simp]: 
  19.507 +  "G\<turnstile>RefT T.[]\<preceq>RefT T'.[] = G\<turnstile>RefT T\<preceq>RefT T'"
  19.508 +apply (rule iffI)
  19.509 +apply (drule widen_ArrayRefT)
  19.510 +apply simp
  19.511 +apply (erule widen.array)
  19.512 +done
  19.513 +
  19.514 +lemma widen_Array_Array: "G\<turnstile>T.[]\<preceq>T'.[] \<Longrightarrow> G\<turnstile>T\<preceq>T'"
  19.515 +apply (drule widen_Array)
  19.516 +apply auto
  19.517 +done
  19.518 +
  19.519 +lemma widen_Array_Class: "G\<turnstile>S.[] \<preceq> Class C \<Longrightarrow> C=Object"
  19.520 +by (auto dest: widen_Array)
  19.521 +
  19.522 +(*
  19.523 +qed_typerel "widen_NT2" "G\<turnstile>S\<preceq>NT \<Longrightarrow> S = NT"
  19.524 + [prove_widen_lemma "G\<turnstile>S\<preceq>T \<Longrightarrow> T = NT \<longrightarrow> S = NT"]
  19.525 +*)
  19.526 +lemma widen_NT2: "G\<turnstile>S\<preceq>NT \<Longrightarrow> S = NT"
  19.527 +apply (ind_cases "G\<turnstile>S\<preceq>T")
  19.528 +by auto
  19.529 +
  19.530 +lemma widen_Object:"\<lbrakk>isrtype G T;ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>RefT T \<preceq> Class Object"
  19.531 +apply (case_tac T)
  19.532 +apply (auto)
  19.533 +apply (subgoal_tac "G\<turnstile>pid_field_type\<preceq>\<^sub>C Object")
  19.534 +apply (auto intro: subcls_ObjectI)
  19.535 +done
  19.536 +
  19.537 +lemma widen_trans_lemma [rule_format (no_asm)]: 
  19.538 +  "\<lbrakk>G\<turnstile>S\<preceq>U; \<forall>C. is_class G C \<longrightarrow> G\<turnstile>C\<preceq>\<^sub>C Object\<rbrakk> \<Longrightarrow> \<forall>T. G\<turnstile>U\<preceq>T \<longrightarrow> G\<turnstile>S\<preceq>T"
  19.539 +apply (erule widen.induct)
  19.540 +apply        safe
  19.541 +prefer      5 apply (drule widen_RefT) apply clarsimp
  19.542 +apply      (frule_tac [1] widen_Iface)
  19.543 +apply      (frule_tac [2] widen_Class)
  19.544 +apply      (frule_tac [3] widen_Class)
  19.545 +apply      (frule_tac [4] widen_Iface)
  19.546 +apply      (frule_tac [5] widen_Class)
  19.547 +apply      (frule_tac [6] widen_Array)
  19.548 +apply      safe
  19.549 +apply            (rule widen.int_obj)
  19.550 +prefer          6 apply (drule implmt_is_class) apply simp
  19.551 +apply (tactic "ALLGOALS (etac thin_rl)")
  19.552 +prefer         6 apply simp
  19.553 +apply          (rule_tac [9] widen.arr_obj)
  19.554 +apply         (rotate_tac [9] -1)
  19.555 +apply         (frule_tac [9] widen_RefT)
  19.556 +apply         (auto elim!: rtrancl_trans subcls_implmt implmt_subint2)
  19.557 +done
  19.558 +
  19.559 +lemma ws_widen_trans: "\<lbrakk>G\<turnstile>S\<preceq>U; G\<turnstile>U\<preceq>T; ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>S\<preceq>T"
  19.560 +by (auto intro: widen_trans_lemma subcls_ObjectI)
  19.561 +
  19.562 +lemma widen_antisym_lemma [rule_format (no_asm)]: "\<lbrakk>G\<turnstile>S\<preceq>T;  
  19.563 + \<forall>I J. G\<turnstile>I\<preceq>I J \<and> G\<turnstile>J\<preceq>I I \<longrightarrow> I = J;  
  19.564 + \<forall>C D. G\<turnstile>C\<preceq>\<^sub>C D \<and> G\<turnstile>D\<preceq>\<^sub>C C \<longrightarrow> C = D;  
  19.565 + \<forall>I  . G\<turnstile>Object\<leadsto>I        \<longrightarrow> False\<rbrakk> \<Longrightarrow> G\<turnstile>T\<preceq>S \<longrightarrow> S = T"
  19.566 +apply (erule widen.induct)
  19.567 +apply (auto dest: widen_Iface widen_NT2 widen_Class)
  19.568 +done
  19.569 +
  19.570 +lemmas subint_antisym = 
  19.571 +       subint1_acyclic [THEN acyclic_impl_antisym_rtrancl, standard]
  19.572 +lemmas subcls_antisym = 
  19.573 +       subcls1_acyclic [THEN acyclic_impl_antisym_rtrancl, standard]
  19.574 +
  19.575 +lemma widen_antisym: "\<lbrakk>G\<turnstile>S\<preceq>T; G\<turnstile>T\<preceq>S; ws_prog G\<rbrakk> \<Longrightarrow> S=T"
  19.576 +by (fast elim: widen_antisym_lemma subint_antisym [THEN antisymD] 
  19.577 +                                   subcls_antisym [THEN antisymD])
  19.578 +
  19.579 +lemma widen_ObjectD [dest!]: "G\<turnstile>Class Object\<preceq>T \<Longrightarrow> T=Class Object"
  19.580 +apply (frule widen_Class)
  19.581 +apply (fast dest: widen_Class_Class widen_Class_Iface)
  19.582 +done
  19.583 +
  19.584 +constdefs
  19.585 +  widens :: "prog \<Rightarrow> [ty list, ty list] \<Rightarrow> bool" ("_\<turnstile>_[\<preceq>]_" [71,71,71] 70)
  19.586 + "G\<turnstile>Ts[\<preceq>]Ts' \<equiv> list_all2 (\<lambda>T T'. G\<turnstile>T\<preceq>T') Ts Ts'"
  19.587 +
  19.588 +lemma widens_Nil [simp]: "G\<turnstile>[][\<preceq>][]"
  19.589 +apply (unfold widens_def)
  19.590 +apply auto
  19.591 +done
  19.592 +
  19.593 +lemma widens_Cons [simp]: "G\<turnstile>(S#Ss)[\<preceq>](T#Ts) = (G\<turnstile>S\<preceq>T \<and> G\<turnstile>Ss[\<preceq>]Ts)"
  19.594 +apply (unfold widens_def)
  19.595 +apply auto
  19.596 +done
  19.597 +
  19.598 +
  19.599 +section "narrowing relation"
  19.600 +
  19.601 +(* all properties of narrowing and casting conversions we actually need *)
  19.602 +(* these can easily be proven from the definitions below *)
  19.603 +(*
  19.604 +rules
  19.605 +  cast_RefT2   "G\<turnstile>S\<preceq>? RefT R   \<Longrightarrow> \<exists>t. S=RefT t" 
  19.606 +  cast_PrimT2  "G\<turnstile>S\<preceq>? PrimT pt \<Longrightarrow> \<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
  19.607 +*)
  19.608 +
  19.609 +(* more detailed than necessary for type-safety, see above rules. *)
  19.610 +inductive "narrow G" intros (* narrowing reference conversion, cf. 5.1.5 *)
  19.611 +
  19.612 +  subcls:  "G\<turnstile>C\<preceq>\<^sub>C D \<Longrightarrow> G\<turnstile>     Class D\<succ>Class C"
  19.613 +  implmt:  "\<not>G\<turnstile>C\<leadsto>I \<Longrightarrow> G\<turnstile>     Class C\<succ>Iface I"
  19.614 +  obj_arr: "G\<turnstile>Class Object\<succ>T.[]"
  19.615 +  int_cls: "G\<turnstile>     Iface I\<succ>Class C"
  19.616 +  subint:  "imethds G I hidings imethds G J entails 
  19.617 +	   (\<lambda>(md, mh   ) (md',mh').\<spacespace>G\<turnstile>mrt mh\<preceq>mrt mh') \<Longrightarrow>
  19.618 +	   \<not>G\<turnstile>I\<preceq>I J         \<spacespace>\<spacespace>\<spacespace>\<Longrightarrow> G\<turnstile>     Iface I\<succ>Iface J"
  19.619 +  array:   "G\<turnstile>RefT S\<succ>RefT T   \<spacespace>\<Longrightarrow> G\<turnstile>   RefT S.[]\<succ>RefT T.[]"
  19.620 +
  19.621 +(*unused*)
  19.622 +lemma narrow_RefT: "G\<turnstile>RefT R\<succ>T \<Longrightarrow> \<exists>t. T=RefT t"
  19.623 +apply (ind_cases "G\<turnstile>S\<succ>T")
  19.624 +by auto
  19.625 +
  19.626 +lemma narrow_RefT2: "G\<turnstile>S\<succ>RefT R \<Longrightarrow> \<exists>t. S=RefT t"
  19.627 +apply (ind_cases "G\<turnstile>S\<succ>T")
  19.628 +by auto
  19.629 +
  19.630 +(*unused*)
  19.631 +lemma narrow_PrimT: "G\<turnstile>PrimT pt\<succ>T \<Longrightarrow> \<exists>t. T=PrimT t"
  19.632 +apply (ind_cases "G\<turnstile>S\<succ>T")
  19.633 +by auto
  19.634 +
  19.635 +lemma narrow_PrimT2: "G\<turnstile>S\<succ>PrimT pt \<Longrightarrow>  
  19.636 +				  \<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
  19.637 +apply (ind_cases "G\<turnstile>S\<succ>T")
  19.638 +by auto
  19.639 +
  19.640 +
  19.641 +section "casting relation"
  19.642 +
  19.643 +inductive "cast G" intros (* casting conversion, cf. 5.5 *)
  19.644 +
  19.645 +  widen:   "G\<turnstile>S\<preceq>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
  19.646 +  narrow:  "G\<turnstile>S\<succ>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
  19.647 +
  19.648 +(*
  19.649 +lemma ??unknown??: "\<lbrakk>G\<turnstile>S\<preceq>T; G\<turnstile>S\<succ>T\<rbrakk> \<Longrightarrow> R"
  19.650 + deferred *)
  19.651 +
  19.652 +(*unused*)
  19.653 +lemma cast_RefT: "G\<turnstile>RefT R\<preceq>? T \<Longrightarrow> \<exists>t. T=RefT t"
  19.654 +apply (ind_cases "G\<turnstile>S\<preceq>? T")
  19.655 +by (auto dest: widen_RefT narrow_RefT)
  19.656 +
  19.657 +lemma cast_RefT2: "G\<turnstile>S\<preceq>? RefT R \<Longrightarrow> \<exists>t. S=RefT t"
  19.658 +apply (ind_cases "G\<turnstile>S\<preceq>? T")
  19.659 +by (auto dest: widen_RefT2 narrow_RefT2)
  19.660 +
  19.661 +(*unused*)
  19.662 +lemma cast_PrimT: "G\<turnstile>PrimT pt\<preceq>? T \<Longrightarrow> \<exists>t. T=PrimT t"
  19.663 +apply (ind_cases "G\<turnstile>S\<preceq>? T")
  19.664 +by (auto dest: widen_PrimT narrow_PrimT)
  19.665 +
  19.666 +lemma cast_PrimT2: "G\<turnstile>S\<preceq>? PrimT pt \<Longrightarrow> \<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
  19.667 +apply (ind_cases "G\<turnstile>S\<preceq>? T")
  19.668 +by (auto dest: widen_PrimT2 narrow_PrimT2)
  19.669 +
  19.670 +end
    20.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    20.2 +++ b/src/HOL/Bali/TypeSafe.thy	Mon Jan 28 17:00:19 2002 +0100
    20.3 @@ -0,0 +1,1211 @@
    20.4 +(*  Title:      isabelle/Bali/TypeSafe.thy
    20.5 +    ID:         $Id$
    20.6 +    Author:     David von Oheimb
    20.7 +    Copyright   1997 Technische Universitaet Muenchen
    20.8 +*)
    20.9 +header {* The type soundness proof for Java *}
   20.10 +
   20.11 +
   20.12 +theory TypeSafe = Eval + WellForm + Conform:
   20.13 +
   20.14 +section "result conformance"
   20.15 +
   20.16 +constdefs
   20.17 +  assign_conforms :: "st \<Rightarrow> (val \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> ty \<Rightarrow> env_ \<Rightarrow> bool"
   20.18 +          ("_\<le>|_\<preceq>_\<Colon>\<preceq>_"                                        [71,71,71,71] 70)
   20.19 + "s\<le>|f\<preceq>T\<Colon>\<preceq>E \<equiv>
   20.20 +  \<forall>s' w. Norm s'\<Colon>\<preceq>E \<longrightarrow> fst E,s'\<turnstile>w\<Colon>\<preceq>T \<longrightarrow> s\<le>|s' \<longrightarrow> assign f w (Norm s')\<Colon>\<preceq>E"
   20.21 +
   20.22 +  rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool"
   20.23 +          ("_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_"                               [71,71,71,71,71,71] 70)
   20.24 +  "G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T 
   20.25 +    \<equiv> case T of
   20.26 +        Inl T  \<Rightarrow> if (\<exists>vf. t=In2 vf)
   20.27 +                  then G,s\<turnstile>fst (the_In2 v)\<Colon>\<preceq>T \<and> s\<le>|snd (the_In2 v)\<preceq>T\<Colon>\<preceq>(G,L)
   20.28 +                  else G,s\<turnstile>the_In1 v\<Colon>\<preceq>T
   20.29 +      | Inr Ts \<Rightarrow> list_all2 (conf G s) (the_In3 v) Ts"
   20.30 +
   20.31 +lemma rconf_In1 [simp]: 
   20.32 + "G,L,s\<turnstile>In1 ec\<succ>In1 v \<Colon>\<preceq>Inl T  =  G,s\<turnstile>v\<Colon>\<preceq>T"
   20.33 +apply (unfold rconf_def)
   20.34 +apply (simp (no_asm))
   20.35 +done
   20.36 +
   20.37 +lemma rconf_In2 [simp]: 
   20.38 + "G,L,s\<turnstile>In2 va\<succ>In2 vf\<Colon>\<preceq>Inl T  = (G,s\<turnstile>fst vf\<Colon>\<preceq>T \<and> s\<le>|snd vf\<preceq>T\<Colon>\<preceq>(G,L))"
   20.39 +apply (unfold rconf_def)
   20.40 +apply (simp (no_asm))
   20.41 +done
   20.42 +
   20.43 +lemma rconf_In3 [simp]: 
   20.44 + "G,L,s\<turnstile>In3 es\<succ>In3 vs\<Colon>\<preceq>Inr Ts = list_all2 (\<lambda>v T. G,s\<turnstile>v\<Colon>\<preceq>T) vs Ts"
   20.45 +apply (unfold rconf_def)
   20.46 +apply (simp (no_asm))
   20.47 +done
   20.48 +
   20.49 +section "fits and conf"
   20.50 +
   20.51 +(* unused *)
   20.52 +lemma conf_fits: "G,s\<turnstile>v\<Colon>\<preceq>T \<Longrightarrow> G,s\<turnstile>v fits T"
   20.53 +apply (unfold fits_def)
   20.54 +apply clarify
   20.55 +apply (erule swap, simp (no_asm_use))
   20.56 +apply (drule conf_RefTD)
   20.57 +apply auto
   20.58 +done
   20.59 +
   20.60 +lemma fits_conf: 
   20.61 +  "\<lbrakk>G,s\<turnstile>v\<Colon>\<preceq>T; G\<turnstile>T\<preceq>? T'; G,s\<turnstile>v fits T'; ws_prog G\<rbrakk> \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T'"
   20.62 +apply (auto dest!: fitsD cast_PrimT2 cast_RefT2)
   20.63 +apply (force dest: conf_RefTD intro: conf_AddrI)
   20.64 +done
   20.65 +
   20.66 +lemma fits_Array: 
   20.67 + "\<lbrakk>G,s\<turnstile>v\<Colon>\<preceq>T; G\<turnstile>T'.[]\<preceq>T.[]; G,s\<turnstile>v fits T'; ws_prog G\<rbrakk> \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T'"
   20.68 +apply (auto dest!: fitsD widen_ArrayPrimT widen_ArrayRefT)
   20.69 +apply (force dest: conf_RefTD intro: conf_AddrI)
   20.70 +done
   20.71 +
   20.72 +
   20.73 +
   20.74 +section "gext"
   20.75 +
   20.76 +lemma halloc_gext: "\<And>s1 s2. G\<turnstile>s1 \<midarrow>halloc oi\<succ>a\<rightarrow> s2 \<Longrightarrow> snd s1\<le>|snd s2"
   20.77 +apply (simp (no_asm_simp) only: split_tupled_all)
   20.78 +apply (erule halloc.induct)
   20.79 +apply  (auto dest!: new_AddrD)
   20.80 +done
   20.81 +
   20.82 +lemma sxalloc_gext: "\<And>s1 s2. G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2 \<Longrightarrow> snd s1\<le>|snd s2"
   20.83 +apply (simp (no_asm_simp) only: split_tupled_all)
   20.84 +apply (erule sxalloc.induct)
   20.85 +apply   (auto dest!: halloc_gext)
   20.86 +done
   20.87 +
   20.88 +lemma eval_gext_lemma [rule_format (no_asm)]: 
   20.89 + "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<Longrightarrow> snd s\<le>|snd s' \<and> (case w of  
   20.90 +    In1 v \<Rightarrow> True  
   20.91 +  | In2 vf \<Rightarrow> normal s \<longrightarrow> (\<forall>v x s. s\<le>|snd (assign (snd vf) v (x,s)))  
   20.92 +  | In3 vs \<Rightarrow> True)"
   20.93 +apply (erule eval_induct)
   20.94 +prefer 24 
   20.95 +  apply (case_tac "inited C (globs s0)", clarsimp, erule thin_rl) (* Init *)
   20.96 +apply (auto del: conjI  dest!: not_initedD gext_new sxalloc_gext halloc_gext
   20.97 + simp  add: lvar_def fvar_def2 avar_def2 init_lvars_def2
   20.98 + split del: split_if_asm split add: sum3.split)
   20.99 +(* 6 subgoals *)
  20.100 +apply force+
  20.101 +done
  20.102 +
  20.103 +lemma evar_gext_f: 
  20.104 +  "G\<turnstile>Norm s1 \<midarrow>e=\<succ>vf \<rightarrow> s2 \<Longrightarrow> s\<le>|snd (assign (snd vf) v (x,s))"
  20.105 +apply (drule eval_gext_lemma [THEN conjunct2])
  20.106 +apply auto
  20.107 +done
  20.108 +
  20.109 +lemmas eval_gext = eval_gext_lemma [THEN conjunct1]
  20.110 +
  20.111 +lemma eval_gext': "G\<turnstile>(x1,s1) \<midarrow>t\<succ>\<rightarrow> (w,x2,s2) \<Longrightarrow> s1\<le>|s2"
  20.112 +apply (drule eval_gext)
  20.113 +apply auto
  20.114 +done
  20.115 +
  20.116 +lemma init_yields_initd: "G\<turnstile>Norm s1 \<midarrow>Init C\<rightarrow> s2 \<Longrightarrow> initd C s2"
  20.117 +apply (erule eval_cases , auto split del: split_if_asm)
  20.118 +apply (case_tac "inited C (globs s1)")
  20.119 +apply  (clarsimp split del: split_if_asm)+
  20.120 +apply (drule eval_gext')+
  20.121 +apply (drule init_class_obj_inited)
  20.122 +apply (erule inited_gext)
  20.123 +apply (simp (no_asm_use))
  20.124 +done
  20.125 +
  20.126 +
  20.127 +section "Lemmas"
  20.128 +
  20.129 +lemma obj_ty_obj_class1: 
  20.130 + "\<lbrakk>wf_prog G; is_type G (obj_ty obj)\<rbrakk> \<Longrightarrow> is_class G (obj_class obj)"
  20.131 +apply (case_tac "tag obj")
  20.132 +apply (auto simp add: obj_ty_def obj_class_def)
  20.133 +done
  20.134 +
  20.135 +lemma oconf_init_obj: 
  20.136 + "\<lbrakk>wf_prog G;  
  20.137 + (case r of Heap a \<Rightarrow> is_type G (obj_ty obj) | Stat C \<Rightarrow> is_class G C)
  20.138 +\<rbrakk> \<Longrightarrow> G,s\<turnstile>obj \<lparr>values:=init_vals (var_tys G (tag obj) r)\<rparr>\<Colon>\<preceq>\<surd>r"
  20.139 +apply (auto intro!: oconf_init_obj_lemma unique_fields)
  20.140 +done
  20.141 +
  20.142 +(*
  20.143 +lemma obj_split: "P obj = (\<forall> oi vs. obj = \<lparr>tag=oi,values=vs\<rparr> \<longrightarrow> ?P \<lparr>tag=oi,values=vs\<rparr>)"
  20.144 +apply auto
  20.145 +apply (case_tac "obj")
  20.146 +apply auto
  20.147 +*)
  20.148 +
  20.149 +lemma conforms_newG: "\<lbrakk>globs s oref = None; (x, s)\<Colon>\<preceq>(G,L);   
  20.150 +  wf_prog G; case oref of Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=vs\<rparr>)  
  20.151 +                        | Stat C \<Rightarrow> is_class G C\<rbrakk> \<Longrightarrow>  
  20.152 +  (x, init_obj G oi oref s)\<Colon>\<preceq>(G, L)"
  20.153 +apply (unfold init_obj_def)
  20.154 +apply (auto elim!: conforms_gupd dest!: oconf_init_obj 
  20.155 +            simp add: obj.update_defs)
  20.156 +done
  20.157 +
  20.158 +lemma conforms_init_class_obj: 
  20.159 + "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L); wf_prog G; class G C=Some y; \<not> inited C (globs s)\<rbrakk> \<Longrightarrow> 
  20.160 +  (x,init_class_obj G C s)\<Colon>\<preceq>(G, L)"
  20.161 +apply (rule not_initedD [THEN conforms_newG])
  20.162 +apply    (auto)
  20.163 +done
  20.164 +
  20.165 +
  20.166 +lemma fst_init_lvars[simp]: 
  20.167 + "fst (init_lvars G C sig (invmode m e) a' pvs (x,s)) = 
  20.168 +  (if static m then x else (np a') x)"
  20.169 +apply (simp (no_asm) add: init_lvars_def2)
  20.170 +done
  20.171 +
  20.172 +
  20.173 +lemma halloc_conforms: "\<And>s1. \<lbrakk>G\<turnstile>s1 \<midarrow>halloc oi\<succ>a\<rightarrow> s2; wf_prog G; s1\<Colon>\<preceq>(G, L); 
  20.174 +  is_type G (obj_ty \<lparr>tag=oi,values=fs\<rparr>)\<rbrakk> \<Longrightarrow> s2\<Colon>\<preceq>(G, L)"
  20.175 +apply (simp (no_asm_simp) only: split_tupled_all)
  20.176 +apply (case_tac "aa")
  20.177 +apply  (auto elim!: halloc_elim_cases dest!: new_AddrD 
  20.178 +       intro!: conforms_newG [THEN conforms_xconf] conf_AddrI)
  20.179 +done
  20.180 +
  20.181 +lemma halloc_type_sound: "\<And>s1. \<lbrakk>G\<turnstile>s1 \<midarrow>halloc oi\<succ>a\<rightarrow> (x,s); wf_prog G; s1\<Colon>\<preceq>(G, L);
  20.182 +  T = obj_ty \<lparr>tag=oi,values=fs\<rparr>; is_type G T\<rbrakk> \<Longrightarrow>  
  20.183 +  (x,s)\<Colon>\<preceq>(G, L) \<and> (x = None \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq>T)"
  20.184 +apply (auto elim!: halloc_conforms)
  20.185 +apply (case_tac "aa")
  20.186 +apply (subst obj_ty_eq)
  20.187 +apply  (auto elim!: halloc_elim_cases dest!: new_AddrD intro!: conf_AddrI)
  20.188 +done
  20.189 +
  20.190 +lemma sxalloc_type_sound: 
  20.191 + "\<And>s1 s2. \<lbrakk>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; wf_prog G\<rbrakk> \<Longrightarrow> case fst s1 of  
  20.192 +  None \<Rightarrow> s2 = s1 | Some x \<Rightarrow>  
  20.193 +  (\<exists>a. fst s2 = Some(Xcpt (Loc a)) \<and> (\<forall>L. s1\<Colon>\<preceq>(G,L) \<longrightarrow> s2\<Colon>\<preceq>(G,L)))"
  20.194 +apply (simp (no_asm_simp) only: split_tupled_all)
  20.195 +apply (erule sxalloc.induct)
  20.196 +apply   auto
  20.197 +apply (rule halloc_conforms [THEN conforms_xconf])
  20.198 +apply     (auto elim!: halloc_elim_cases dest!: new_AddrD intro!: conf_AddrI)
  20.199 +done
  20.200 +
  20.201 +lemma wt_init_comp_ty: 
  20.202 +"is_acc_type G (pid C) T \<Longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>"
  20.203 +apply (unfold init_comp_ty_def)
  20.204 +apply (clarsimp simp add: accessible_in_RefT_simp 
  20.205 +                          is_acc_type_def is_acc_class_def)
  20.206 +done
  20.207 +
  20.208 +
  20.209 +declare fun_upd_same [simp]
  20.210 +
  20.211 +declare fun_upd_apply [simp del]
  20.212 +
  20.213 +
  20.214 +constdefs
  20.215 +  DynT_prop::"[prog,inv_mode,qtname,ref_ty] \<Rightarrow> bool" 
  20.216 +                                              ("_\<turnstile>_\<rightarrow>_\<preceq>_"[71,71,71,71]70)
  20.217 + "G\<turnstile>mode\<rightarrow>D\<preceq>t \<equiv> mode = IntVir \<longrightarrow> is_class G D \<and> 
  20.218 +                     (if (\<exists>T. t=ArrayT T) then D=Object else G\<turnstile>Class D\<preceq>RefT t)"
  20.219 +
  20.220 +lemma DynT_propI: 
  20.221 + "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L); G,s\<turnstile>a'\<Colon>\<preceq>RefT statT; wf_prog G; mode = IntVir \<longrightarrow> a' \<noteq> Null\<rbrakk> 
  20.222 +  \<Longrightarrow>  G\<turnstile>mode\<rightarrow>invocation_class mode s a' statT\<preceq>statT"
  20.223 +proof (unfold DynT_prop_def)
  20.224 +  assume state_conform: "(x,s)\<Colon>\<preceq>(G, L)"
  20.225 +     and      statT_a': "G,s\<turnstile>a'\<Colon>\<preceq>RefT statT"
  20.226 +     and            wf: "wf_prog G"
  20.227 +     and          mode: "mode = IntVir \<longrightarrow> a' \<noteq> Null"
  20.228 +  let ?invCls = "(invocation_class mode s a' statT)"
  20.229 +  let ?IntVir = "mode = IntVir"
  20.230 +  let ?Concl  = "\<lambda>invCls. is_class G invCls \<and>
  20.231 +                          (if \<exists>T. statT = ArrayT T
  20.232 +                                  then invCls = Object
  20.233 +                                  else G\<turnstile>Class invCls\<preceq>RefT statT)"
  20.234 +  show "?IntVir \<longrightarrow> ?Concl ?invCls"
  20.235 +  proof  
  20.236 +    assume modeIntVir: ?IntVir 
  20.237 +    with mode have not_Null: "a' \<noteq> Null" ..
  20.238 +    from statT_a' not_Null state_conform 
  20.239 +    obtain a obj 
  20.240 +      where obj_props:  "a' = Addr a" "globs s (Inl a) = Some obj"   
  20.241 +                        "G\<turnstile>obj_ty obj\<preceq>RefT statT" "is_type G (obj_ty obj)"
  20.242 +      by (blast dest: conforms_RefTD)
  20.243 +    show "?Concl ?invCls"
  20.244 +    proof (cases "tag obj")
  20.245 +      case CInst
  20.246 +      with modeIntVir obj_props
  20.247 +      show ?thesis 
  20.248 +	by (auto dest!: widen_Array2 split add: split_if)
  20.249 +    next
  20.250 +      case Arr
  20.251 +      from Arr obtain T where "obj_ty obj = T.[]" by (blast dest: obj_ty_Arr1)
  20.252 +      moreover from Arr have "obj_class obj = Object" 
  20.253 +	by (blast dest: obj_class_Arr1)
  20.254 +      moreover note modeIntVir obj_props wf 
  20.255 +      ultimately show ?thesis by (auto dest!: widen_Array )
  20.256 +    qed
  20.257 +  qed
  20.258 +qed
  20.259 +
  20.260 +lemma invocation_methd:
  20.261 +"\<lbrakk>wf_prog G; statT \<noteq> NullT; 
  20.262 + (\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC);
  20.263 + (\<forall>     I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> mode \<noteq> SuperM);
  20.264 + (\<forall>     T. statT = ArrayT T \<longrightarrow> mode \<noteq> SuperM);
  20.265 + G\<turnstile>mode\<rightarrow>invocation_class mode s a' statT\<preceq>statT;  
  20.266 + dynlookup G statT (invocation_class mode s a' statT) sig = Some m \<rbrakk> 
  20.267 +\<Longrightarrow> methd G (invocation_declclass G mode s a' statT sig) sig = Some m"
  20.268 +proof -
  20.269 +  assume         wf: "wf_prog G"
  20.270 +     and  not_NullT: "statT \<noteq> NullT"
  20.271 +     and statC_prop: "(\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC)"
  20.272 +     and statI_prop: "(\<forall> I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> mode \<noteq> SuperM)"
  20.273 +     and statA_prop: "(\<forall>     T. statT = ArrayT T \<longrightarrow> mode \<noteq> SuperM)"
  20.274 +     and  invC_prop: "G\<turnstile>mode\<rightarrow>invocation_class mode s a' statT\<preceq>statT"
  20.275 +     and  dynlookup: "dynlookup G statT (invocation_class mode s a' statT) sig 
  20.276 +                      = Some m"
  20.277 +  show ?thesis
  20.278 +  proof (cases statT)
  20.279 +    case NullT
  20.280 +    with not_NullT show ?thesis by simp
  20.281 +  next
  20.282 +    case IfaceT
  20.283 +    with statI_prop obtain I 
  20.284 +      where    statI: "statT = IfaceT I" and 
  20.285 +            is_iface: "is_iface G I"     and
  20.286 +          not_SuperM: "mode \<noteq> SuperM" by blast            
  20.287 +    
  20.288 +    show ?thesis 
  20.289 +    proof (cases mode)
  20.290 +      case Static
  20.291 +      with wf dynlookup statI is_iface 
  20.292 +      show ?thesis
  20.293 +	by (auto simp add: invocation_declclass_def dynlookup_def 
  20.294 +                           dynimethd_def dynmethd_C_C 
  20.295 +	            intro: dynmethd_declclass
  20.296 +	            dest!: wf_imethdsD
  20.297 +                     dest: table_of_map_SomeI
  20.298 +                    split: split_if_asm)
  20.299 +    next	
  20.300 +      case SuperM
  20.301 +      with not_SuperM show ?thesis ..
  20.302 +    next
  20.303 +      case IntVir
  20.304 +      with wf dynlookup IfaceT invC_prop show ?thesis 
  20.305 +	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
  20.306 +                           DynT_prop_def
  20.307 +	            intro: methd_declclass dynmethd_declclass
  20.308 +	            split: split_if_asm)
  20.309 +    qed
  20.310 +  next
  20.311 +    case ClassT
  20.312 +    show ?thesis
  20.313 +    proof (cases mode)
  20.314 +      case Static
  20.315 +      with wf ClassT dynlookup statC_prop 
  20.316 +      show ?thesis by (auto simp add: invocation_declclass_def dynlookup_def
  20.317 +                               intro: dynmethd_declclass)
  20.318 +    next
  20.319 +      case SuperM
  20.320 +      with wf ClassT dynlookup statC_prop 
  20.321 +      show ?thesis by (auto simp add: invocation_declclass_def dynlookup_def
  20.322 +                               intro: dynmethd_declclass)
  20.323 +    next
  20.324 +      case IntVir
  20.325 +      with wf ClassT dynlookup statC_prop invC_prop 
  20.326 +      show ?thesis
  20.327 +	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
  20.328 +                           DynT_prop_def
  20.329 +	            intro: dynmethd_declclass)
  20.330 +    qed
  20.331 +  next
  20.332 +    case ArrayT
  20.333 +    show ?thesis
  20.334 +    proof (cases mode)
  20.335 +      case Static
  20.336 +      with wf ArrayT dynlookup show ?thesis
  20.337 +	by (auto simp add: invocation_declclass_def dynlookup_def 
  20.338 +                           dynimethd_def dynmethd_C_C
  20.339 +                    intro: dynmethd_declclass
  20.340 +                     dest: table_of_map_SomeI)
  20.341 +    next
  20.342 +      case SuperM
  20.343 +      with ArrayT statA_prop show ?thesis by blast
  20.344 +    next
  20.345 +      case IntVir
  20.346 +      with wf ArrayT dynlookup invC_prop show ?thesis
  20.347 +	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
  20.348 +                           DynT_prop_def dynmethd_C_C
  20.349 +                    intro: dynmethd_declclass
  20.350 +                     dest: table_of_map_SomeI)
  20.351 +    qed
  20.352 +  qed
  20.353 +qed
  20.354 +   
  20.355 +declare split_paired_All [simp del] split_paired_Ex [simp del] 
  20.356 +ML_setup {*
  20.357 +simpset_ref() := simpset() delloop "split_all_tac";
  20.358 +claset_ref () := claset () delSWrapper "split_all_tac"
  20.359 +*}
  20.360 +lemma DynT_mheadsD: 
  20.361 +"\<lbrakk>G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT; 
  20.362 +  wf_prog G; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT; 
  20.363 +  sm \<in> mheads G C statT sig; 
  20.364 +  invC = invocation_class (invmode (mhd sm) e) s a' statT;
  20.365 +  declC =invocation_declclass G (invmode (mhd sm) e) s a' statT sig
  20.366 + \<rbrakk> \<Longrightarrow> 
  20.367 +  \<exists> dm. 
  20.368 +  methd G declC sig = Some dm  \<and> G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm) \<and> 
  20.369 +  wf_mdecl G declC (sig, mthd dm) \<and>
  20.370 +  declC = declclass dm \<and>
  20.371 +  is_static dm = is_static sm \<and>  
  20.372 +  is_class G invC \<and> is_class G declC  \<and> G\<turnstile>invC\<preceq>\<^sub>C declC \<and>  
  20.373 +  (if invmode (mhd sm) e = IntVir
  20.374 +      then (\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)
  20.375 +      else (  (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
  20.376 +            \<or> (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object)) \<and> 
  20.377 +           (declrefT sm) = ClassT (declclass dm))"
  20.378 +proof -
  20.379 +  assume invC_prop: "G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT" 
  20.380 +     and        wf: "wf_prog G" 
  20.381 +     and      wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
  20.382 +     and        sm: "sm \<in> mheads G C statT sig" 
  20.383 +     and      invC: "invC = invocation_class (invmode (mhd sm) e) s a' statT"
  20.384 +     and     declC: "declC = 
  20.385 +                    invocation_declclass G (invmode (mhd sm) e) s a' statT sig"
  20.386 +  from wt_e wf have type_statT: "is_type G (RefT statT)"
  20.387 +    by (auto dest: ty_expr_is_type)
  20.388 +  from sm have not_Null: "statT \<noteq> NullT" by auto
  20.389 +  from type_statT 
  20.390 +  have wf_C: "(\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC)"
  20.391 +    by (auto)
  20.392 +  from type_statT wt_e 
  20.393 +  have wf_I: "(\<forall>I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> 
  20.394 +                                        invmode (mhd sm) e \<noteq> SuperM)"
  20.395 +    by (auto dest: invocationTypeExpr_noClassD)
  20.396 +  from wt_e
  20.397 +  have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode (mhd sm) e \<noteq> SuperM)"
  20.398 +    by (auto dest: invocationTypeExpr_noClassD)
  20.399 +  show ?thesis
  20.400 +  proof (cases "invmode (mhd sm) e = IntVir")
  20.401 +    case True
  20.402 +    with invC_prop not_Null
  20.403 +    have invC_prop': " is_class G invC \<and> 
  20.404 +                      (if (\<exists>T. statT=ArrayT T) then invC=Object 
  20.405 +                                              else G\<turnstile>Class invC\<preceq>RefT statT)"
  20.406 +      by (auto simp add: DynT_prop_def) 
  20.407 +    from True 
  20.408 +    have "\<not> is_static sm"
  20.409 +      by (simp add: invmode_IntVir_eq)
  20.410 +    with invC_prop' not_Null
  20.411 +    have "G,statT \<turnstile> invC valid_lookup_cls_for (is_static sm)"
  20.412 +      by (cases statT) auto
  20.413 +    with sm wf type_statT obtain dm where
  20.414 +           dm: "dynlookup G statT invC sig = Some dm" and
  20.415 +      resT_dm: "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm)"      and
  20.416 +       static: "is_static dm = is_static sm"
  20.417 +      by - (drule dynamic_mheadsD,auto)
  20.418 +    with declC invC not_Null 
  20.419 +    have declC': "declC = (declclass dm)" 
  20.420 +      by (auto simp add: invocation_declclass_def)
  20.421 +    with wf invC declC not_Null wf_C wf_I wf_A invC_prop dm 
  20.422 +    have dm': "methd G declC sig = Some dm"
  20.423 +      by - (drule invocation_methd,auto)
  20.424 +    from wf dm invC_prop' declC' type_statT 
  20.425 +    have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
  20.426 +      by (auto dest: dynlookup_declC)
  20.427 +    from wf dm' declC_prop declC' 
  20.428 +    have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
  20.429 +      by (auto dest: methd_wf_mdecl)
  20.430 +    from invC_prop' 
  20.431 +    have statC_prop: "(\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)"
  20.432 +      by auto
  20.433 +    from True dm' resT_dm wf_dm invC_prop' declC_prop statC_prop declC' static
  20.434 +    show ?thesis by auto
  20.435 +  next
  20.436 +    case False
  20.437 +    with type_statT wf invC not_Null wf_I wf_A
  20.438 +    have invC_prop': "is_class G invC \<and>  
  20.439 +                     ((\<exists> statC. statT=ClassT statC \<and> invC=statC) \<or>
  20.440 +                      (\<forall> statC. statT\<noteq>ClassT statC \<and> invC=Object)) "
  20.441 +        by (case_tac "statT") (auto simp add: invocation_class_def 
  20.442 +                                       split: inv_mode.splits)
  20.443 +    with not_Null wf
  20.444 +    have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig"
  20.445 +      by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_C_C
  20.446 +                                            dynimethd_def)
  20.447 +    from sm wf wt_e not_Null False invC_prop' obtain "dm" where
  20.448 +                    dm: "methd G invC sig = Some dm"          and
  20.449 +	eq_declC_sm_dm:"declrefT sm = ClassT (declclass dm)"  and
  20.450 +	     eq_mheads:"mhd sm=mhead (mthd dm) "
  20.451 +      by - (drule static_mheadsD, auto dest: accmethd_SomeD)
  20.452 +    then have static: "is_static dm = is_static sm" by - (case_tac "sm",auto)
  20.453 +    with declC invC dynlookup_static dm
  20.454 +    have declC': "declC = (declclass dm)"  
  20.455 +      by (auto simp add: invocation_declclass_def)
  20.456 +    from invC_prop' wf declC' dm 
  20.457 +    have dm': "methd G declC sig = Some dm"
  20.458 +      by (auto intro: methd_declclass)
  20.459 +    from wf dm invC_prop' declC' type_statT 
  20.460 +    have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
  20.461 +      by (auto dest: methd_declC )
  20.462 +    then have declC_prop1: "invC=Object \<longrightarrow> declC=Object"  by auto
  20.463 +    from wf dm' declC_prop declC' 
  20.464 +    have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
  20.465 +      by (auto dest: methd_wf_mdecl)
  20.466 +    from invC_prop' declC_prop declC_prop1
  20.467 +    have statC_prop: "(   (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
  20.468 +                       \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object))" 
  20.469 +      by auto
  20.470 +    from False dm' wf_dm invC_prop' declC_prop statC_prop declC' 
  20.471 +         eq_declC_sm_dm eq_mheads static
  20.472 +    show ?thesis by auto
  20.473 +  qed
  20.474 +qed
  20.475 +
  20.476 +(*
  20.477 +lemma DynT_mheadsD: 
  20.478 +"\<lbrakk>G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT; 
  20.479 +  wf_prog G; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT; 
  20.480 +  sm \<in> mheads G C statT sig; 
  20.481 +  invC = invocation_class (invmode (mhd sm) e) s a' statT;
  20.482 +  declC =invocation_declclass G (invmode (mhd sm) e) s a' statT sig
  20.483 + \<rbrakk> \<Longrightarrow> 
  20.484 +  \<exists> dm. 
  20.485 +  methd G declC sig = Some dm  \<and> G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm) \<and> 
  20.486 +  wf_mdecl G declC (sig, mthd dm) \<and>  
  20.487 +  is_class G invC \<and> is_class G declC  \<and> G\<turnstile>invC\<preceq>\<^sub>C declC \<and>  
  20.488 +  (if invmode (mhd sm) e = IntVir
  20.489 +      then (\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)
  20.490 +      else (\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>statC \<preceq>\<^sub>C declC) \<and> 
  20.491 +           (declrefT sm) = ClassT (declclass dm))"
  20.492 +proof -
  20.493 +  assume invC_prop: "G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT" 
  20.494 +     and        wf: "wf_prog G" 
  20.495 +     and      wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
  20.496 +     and        sm: "sm \<in> mheads G C statT sig" 
  20.497 +     and      invC: "invC = invocation_class (invmode (mhd sm) e) s a' statT"
  20.498 +     and     declC: "declC = 
  20.499 +                    invocation_declclass G (invmode (mhd sm) e) s a' statT sig"
  20.500 +  from wt_e wf have type_statT: "is_type G (RefT statT)"
  20.501 +    by (auto dest: ty_expr_is_type)
  20.502 +  from sm have not_Null: "statT \<noteq> NullT" by auto
  20.503 +  from type_statT 
  20.504 +  have wf_C: "(\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC)"
  20.505 +    by (auto)
  20.506 +  from type_statT wt_e 
  20.507 +  have wf_I: "(\<forall>I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> 
  20.508 +                                        invmode (mhd sm) e \<noteq> SuperM)"
  20.509 +    by (auto dest: invocationTypeExpr_noClassD)
  20.510 +  from wt_e
  20.511 +  have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode (mhd sm) e \<noteq> SuperM)"
  20.512 +    by (auto dest: invocationTypeExpr_noClassD)
  20.513 +  show ?thesis
  20.514 +  proof (cases "invmode (mhd sm) e = IntVir")
  20.515 +    case True
  20.516 +    with invC_prop not_Null
  20.517 +    have invC_prop': "is_class G invC \<and>  
  20.518 +                      (if (\<exists>T. statT=ArrayT T) then invC=Object 
  20.519 +                                              else G\<turnstile>Class invC\<preceq>RefT statT)"
  20.520 +      by (auto simp add: DynT_prop_def) 
  20.521 +    with sm wf type_statT not_Null obtain dm where
  20.522 +           dm: "dynlookup G statT invC sig = Some dm" and
  20.523 +      resT_dm: "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm)"
  20.524 +      by - (clarify,drule dynamic_mheadsD,auto)
  20.525 +    with declC invC not_Null 
  20.526 +    have declC': "declC = (declclass dm)" 
  20.527 +      by (auto simp add: invocation_declclass_def)
  20.528 +    with wf invC declC not_Null wf_C wf_I wf_A invC_prop dm 
  20.529 +    have dm': "methd G declC sig = Some dm"
  20.530 +      by - (drule invocation_methd,auto)
  20.531 +    from wf dm invC_prop' declC' type_statT 
  20.532 +    have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
  20.533 +      by (auto dest: dynlookup_declC)
  20.534 +    from wf dm' declC_prop declC' 
  20.535 +    have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
  20.536 +      by (auto dest: methd_wf_mdecl)
  20.537 +    from invC_prop' 
  20.538 +    have statC_prop: "(\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)"
  20.539 +      by auto
  20.540 +    from True dm' resT_dm wf_dm invC_prop' declC_prop statC_prop
  20.541 +    show ?thesis by auto
  20.542 +  next
  20.543 +    case False
  20.544 +    
  20.545 +    with type_statT wf invC not_Null wf_I wf_A
  20.546 +    have invC_prop': "is_class G invC \<and>  
  20.547 +                     ((\<exists> statC. statT=ClassT statC \<and> invC=statC) \<or>
  20.548 +                      (\<forall> statC. statT\<noteq>ClassT statC \<and> invC=Object)) "
  20.549 +        
  20.550 +        by (case_tac "statT") (auto simp add: invocation_class_def 
  20.551 +                                       split: inv_mode.splits)
  20.552 +    with not_Null 
  20.553 +    have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig"
  20.554 +      by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_def 
  20.555 +                                            dynimethd_def)
  20.556 +    from sm wf wt_e not_Null False invC_prop' obtain "dm" where
  20.557 +                    dm: "methd G invC sig = Some dm"          and
  20.558 +	eq_declC_sm_dm:"declrefT sm = ClassT (declclass dm)"  and
  20.559 +	     eq_mheads:"mhd sm=mhead (mthd dm) "
  20.560 +      by - (drule static_mheadsD, auto dest: accmethd_SomeD)
  20.561 +    with declC invC dynlookup_static dm
  20.562 +    have declC': "declC = (declclass dm)"  
  20.563 +      by (auto simp add: invocation_declclass_def)
  20.564 +    from invC_prop' wf declC' dm 
  20.565 +    have dm': "methd G declC sig = Some dm"
  20.566 +      by (auto intro: methd_declclass)
  20.567 +    from wf dm invC_prop' declC' type_statT 
  20.568 +    have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
  20.569 +      by (auto dest: methd_declC )   
  20.570 +    from wf dm' declC_prop declC' 
  20.571 +    have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
  20.572 +      by (auto dest: methd_wf_mdecl)
  20.573 +    from invC_prop' declC_prop
  20.574 +    have statC_prop: "(\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>statC \<preceq>\<^sub>C declC)" 
  20.575 +      by auto
  20.576 +    from False dm' wf_dm invC_prop' declC_prop statC_prop 
  20.577 +         eq_declC_sm_dm eq_mheads
  20.578 +    show ?thesis by auto
  20.579 +  qed
  20.580 +qed	
  20.581 +*)
  20.582 +
  20.583 +declare split_paired_All [simp del] split_paired_Ex [simp del] 
  20.584 +declare split_if     [split del] split_if_asm     [split del] 
  20.585 +        option.split [split del] option.split_asm [split del]
  20.586 +ML_setup {*
  20.587 +simpset_ref() := simpset() delloop "split_all_tac";
  20.588 +claset_ref () := claset () delSWrapper "split_all_tac"
  20.589 +*}
  20.590 +
  20.591 +lemma DynT_conf: "\<lbrakk>G\<turnstile>invocation_class mode s a' statT \<preceq>\<^sub>C declC; wf_prog G;
  20.592 + isrtype G (statT);
  20.593 + G,s\<turnstile>a'\<Colon>\<preceq>RefT statT; mode = IntVir \<longrightarrow> a' \<noteq> Null;  
  20.594 +  mode \<noteq> IntVir \<longrightarrow>    (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
  20.595 +                    \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object)\<rbrakk> 
  20.596 + \<Longrightarrow>G,s\<turnstile>a'\<Colon>\<preceq> Class declC"
  20.597 +apply (case_tac "mode = IntVir")
  20.598 +apply (drule conf_RefTD)
  20.599 +apply (force intro!: conf_AddrI 
  20.600 +            simp add: obj_class_def split add: obj_tag.split_asm)
  20.601 +apply  clarsimp
  20.602 +apply  safe
  20.603 +apply    (erule (1) widen.subcls [THEN conf_widen])
  20.604 +apply    (erule wf_ws_prog)
  20.605 +
  20.606 +apply    (frule widen_Object) apply (erule wf_ws_prog)
  20.607 +apply    (erule (1) conf_widen) apply (erule wf_ws_prog)
  20.608 +done
  20.609 +
  20.610 +
  20.611 +lemma Ass_lemma: 
  20.612 + "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w, f)\<rightarrow> Norm s1; G\<turnstile>Norm s1 \<midarrow>e-\<succ>v\<rightarrow> Norm s2; G,s2\<turnstile>v\<Colon>\<preceq>T'; 
  20.613 +   s1\<le>|s2 \<longrightarrow> assign f v (Norm s2)\<Colon>\<preceq>(G, L)
  20.614 +  \<rbrakk> \<Longrightarrow> assign f v (Norm s2)\<Colon>\<preceq>(G, L) \<and> 
  20.615 +        (\<lambda>(x',s'). x' = None \<longrightarrow> G,s'\<turnstile>v\<Colon>\<preceq>T') (assign f v (Norm s2))"
  20.616 +apply (drule_tac x = "None" and s = "s2" and v = "v" in evar_gext_f)
  20.617 +apply (drule eval_gext', clarsimp)
  20.618 +apply (erule conf_gext)
  20.619 +apply simp
  20.620 +done
  20.621 +
  20.622 +lemma Throw_lemma: "\<lbrakk>G\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable; wf_prog G; (x1,s1)\<Colon>\<preceq>(G, L);  
  20.623 +    x1 = None \<longrightarrow> G,s1\<turnstile>a'\<Colon>\<preceq> Class tn\<rbrakk> \<Longrightarrow> (throw a' x1, s1)\<Colon>\<preceq>(G, L)"
  20.624 +apply (auto split add: split_abrupt_if simp add: throw_def2)
  20.625 +apply (erule conforms_xconf)
  20.626 +apply (frule conf_RefTD)
  20.627 +apply (auto elim: widen.subcls [THEN conf_widen])
  20.628 +done
  20.629 +
  20.630 +lemma Try_lemma: "\<lbrakk>G\<turnstile>obj_ty (the (globs s1' (Heap a)))\<preceq> Class tn; 
  20.631 + (Some (Xcpt (Loc a)), s1')\<Colon>\<preceq>(G, L); wf_prog G\<rbrakk> 
  20.632 + \<Longrightarrow> Norm (lupd(vn\<mapsto>Addr a) s1')\<Colon>\<preceq>(G, L(vn\<mapsto>Class tn))"
  20.633 +apply (rule conforms_allocL)
  20.634 +apply  (erule conforms_NormI)
  20.635 +apply (drule conforms_XcptLocD [THEN conf_RefTD],rule HOL.refl)
  20.636 +apply (auto intro!: conf_AddrI)
  20.637 +done
  20.638 +
  20.639 +lemma Fin_lemma: 
  20.640 +"\<lbrakk>G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> (x2,s2); wf_prog G; (Some a, s1)\<Colon>\<preceq>(G, L); (x2,s2)\<Colon>\<preceq>(G, L)\<rbrakk> 
  20.641 +\<Longrightarrow>  (abrupt_if True (Some a) x2, s2)\<Colon>\<preceq>(G, L)"
  20.642 +apply (force elim: eval_gext' conforms_xgext split add: split_abrupt_if)
  20.643 +done
  20.644 +
  20.645 +lemma FVar_lemma1: "\<lbrakk>table_of (DeclConcepts.fields G Ca) (fn, C) = Some f ; 
  20.646 +  x2 = None \<longrightarrow> G,s2\<turnstile>a\<Colon>\<preceq> Class Ca; wf_prog G; G\<turnstile>Ca\<preceq>\<^sub>C C; C \<noteq> Object; 
  20.647 +  class G C = Some y; (x2,s2)\<Colon>\<preceq>(G, L); s1\<le>|s2; inited C (globs s1); 
  20.648 +  (if static f then id else np a) x2 = None\<rbrakk> 
  20.649 + \<Longrightarrow>  
  20.650 +  \<exists>obj. globs s2 (if static f then Inr C else Inl (the_Addr a)) = Some obj \<and> 
  20.651 +  var_tys G (tag obj)  (if static f then Inr C else Inl(the_Addr a)) 
  20.652 +          (Inl(fn,C)) = Some (type f)"
  20.653 +apply (drule initedD)
  20.654 +apply (frule subcls_is_class2, simp (no_asm_simp))
  20.655 +apply (case_tac "static f")
  20.656 +apply  clarsimp
  20.657 +apply  (drule (1) rev_gext_objD, clarsimp)
  20.658 +apply  (frule fields_declC, erule wf_ws_prog, simp (no_asm_simp))
  20.659 +apply  (rule var_tys_Some_eq [THEN iffD2])
  20.660 +apply  clarsimp
  20.661 +apply  (erule fields_table_SomeI, simp (no_asm))
  20.662 +apply clarsimp
  20.663 +apply (drule conf_RefTD, clarsimp, rule var_tys_Some_eq [THEN iffD2])
  20.664 +apply (auto dest!: widen_Array split add: obj_tag.split)
  20.665 +apply (rotate_tac -1) (* to front: tag obja = CInst pid_field_type to enable
  20.666 +                         conditional rewrite *)
  20.667 +apply (rule fields_table_SomeI)
  20.668 +apply (auto elim!: fields_mono subcls_is_class2)
  20.669 +done
  20.670 +
  20.671 +lemma FVar_lemma: 
  20.672 +"\<lbrakk>((v, f), Norm s2') = fvar C (static field) fn a (x2, s2); G\<turnstile>Ca\<preceq>\<^sub>C C;  
  20.673 +  table_of (DeclConcepts.fields G Ca) (fn, C) = Some field; wf_prog G;   
  20.674 +  x2 = None \<longrightarrow> G,s2\<turnstile>a\<Colon>\<preceq>Class Ca; C \<noteq> Object; class G C = Some y;   
  20.675 +  (x2, s2)\<Colon>\<preceq>(G, L); s1\<le>|s2; inited C (globs s1)\<rbrakk> \<Longrightarrow>  
  20.676 +  G,s2'\<turnstile>v\<Colon>\<preceq>type field \<and> s2'\<le>|f\<preceq>type field\<Colon>\<preceq>(G, L)"
  20.677 +apply (unfold assign_conforms_def)
  20.678 +apply (drule sym)
  20.679 +apply (clarsimp simp add: fvar_def2)
  20.680 +apply (drule (9) FVar_lemma1)
  20.681 +apply (clarsimp)
  20.682 +apply (drule (2) conforms_globsD [THEN oconf_lconf, THEN lconfD])
  20.683 +apply clarsimp
  20.684 +apply (drule (1) rev_gext_objD)
  20.685 +apply (auto elim!: conforms_upd_gobj)
  20.686 +done
  20.687 +
  20.688 +
  20.689 +lemma AVar_lemma1: "\<lbrakk>globs s (Inl a) = Some obj;tag obj=Arr ty i; 
  20.690 + the_Intg i' in_bounds i; wf_prog G; G\<turnstile>ty.[]\<preceq>Tb.[]; Norm s\<Colon>\<preceq>(G, L)
  20.691 +\<rbrakk> \<Longrightarrow> G,s\<turnstile>the ((values obj) (Inr (the_Intg i')))\<Colon>\<preceq>Tb"
  20.692 +apply (erule widen_Array_Array [THEN conf_widen])
  20.693 +apply  (erule_tac [2] wf_ws_prog)
  20.694 +apply (drule (1) conforms_globsD [THEN oconf_lconf, THEN lconfD])
  20.695 +defer apply assumption
  20.696 +apply (force intro: var_tys_Some_eq [THEN iffD2])
  20.697 +done
  20.698 +
  20.699 +lemma obj_split: "\<And> obj. \<exists> t vs. obj = \<lparr>tag=t,values=vs\<rparr>"
  20.700 +proof record_split
  20.701 +  fix tag values more
  20.702 +  show "\<exists>t vs. \<lparr>tag = tag, values = values, \<dots> = more\<rparr> = \<lparr>tag = t, values = vs\<rparr>"
  20.703 +    by auto
  20.704 +qed
  20.705 + 
  20.706 +lemma AVar_lemma: "\<lbrakk>wf_prog G; G\<turnstile>(x1, s1) \<midarrow>e2-\<succ>i\<rightarrow> (x2, s2);  
  20.707 +  ((v,f), Norm s2') = avar G i a (x2, s2); x1 = None \<longrightarrow> G,s1\<turnstile>a\<Colon>\<preceq>Ta.[];  
  20.708 +  (x2, s2)\<Colon>\<preceq>(G, L); s1\<le>|s2\<rbrakk> \<Longrightarrow> G,s2'\<turnstile>v\<Colon>\<preceq>Ta \<and> s2'\<le>|f\<preceq>Ta\<Colon>\<preceq>(G, L)"
  20.709 +apply (unfold assign_conforms_def)
  20.710 +apply (drule sym)
  20.711 +apply (clarsimp simp add: avar_def2)
  20.712 +apply (drule (1) conf_gext)
  20.713 +apply (drule conf_RefTD, clarsimp)
  20.714 +apply (subgoal_tac "\<exists> t vs. obj = \<lparr>tag=t,values=vs\<rparr>")
  20.715 +defer
  20.716 +apply   (rule obj_split)
  20.717 +apply clarify
  20.718 +apply (frule obj_ty_widenD)
  20.719 +apply (auto dest!: widen_Class)
  20.720 +apply  (force dest: AVar_lemma1)
  20.721 +apply (auto split add: split_if)
  20.722 +apply (force elim!: fits_Array dest: gext_objD 
  20.723 +       intro: var_tys_Some_eq [THEN iffD2] conforms_upd_gobj)
  20.724 +done
  20.725 +
  20.726 +
  20.727 +section "Call"
  20.728 +lemma conforms_init_lvars_lemma: "\<lbrakk>wf_prog G;  
  20.729 +  wf_mhead G P sig mh; 
  20.730 +  Ball (set lvars) (split (\<lambda>vn. is_type G)); 
  20.731 +  list_all2 (conf G s) pvs pTsa; G\<turnstile>pTsa[\<preceq>](parTs sig)\<rbrakk> \<Longrightarrow>  
  20.732 +  G,s\<turnstile>init_vals (table_of lvars)(pars mh[\<mapsto>]pvs)
  20.733 +      [\<Colon>\<preceq>]table_of lvars(pars mh[\<mapsto>]parTs sig)"
  20.734 +apply (unfold wf_mhead_def)
  20.735 +apply clarify
  20.736 +apply (erule (2) Ball_set_table [THEN lconf_init_vals, THEN lconf_ext_list])
  20.737 +apply (drule wf_ws_prog)
  20.738 +apply (erule (2) conf_list_widen)
  20.739 +done
  20.740 +
  20.741 +
  20.742 +lemma lconf_map_lname [simp]: 
  20.743 +  "G,s\<turnstile>(lname_case l1 l2)[\<Colon>\<preceq>](lname_case L1 L2)
  20.744 +   =
  20.745 +  (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit . l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
  20.746 +apply (unfold lconf_def)
  20.747 +apply safe
  20.748 +apply (case_tac "n")
  20.749 +apply (force split add: lname.split)+
  20.750 +done
  20.751 +
  20.752 +lemma lconf_map_ename [simp]:
  20.753 +  "G,s\<turnstile>(ename_case l1 l2)[\<Colon>\<preceq>](ename_case L1 L2)
  20.754 +   =
  20.755 +  (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit. l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
  20.756 +apply (unfold lconf_def)
  20.757 +apply safe
  20.758 +apply (force split add: ename.split)+
  20.759 +done
  20.760 +
  20.761 +
  20.762 +lemma defval_conf1 [rule_format (no_asm), elim]: 
  20.763 +  "is_type G T \<longrightarrow> (\<exists>v\<in>Some (default_val T): G,s\<turnstile>v\<Colon>\<preceq>T)"
  20.764 +apply (unfold conf_def)
  20.765 +apply (induct "T")
  20.766 +apply (auto intro: prim_ty.induct)
  20.767 +done
  20.768 +
  20.769 +
  20.770 +lemma conforms_init_lvars: 
  20.771 +"\<lbrakk>wf_mhead G (pid declC) sig (mhead (mthd dm)); wf_prog G;  
  20.772 +  list_all2 (conf G s) pvs pTsa; G\<turnstile>pTsa[\<preceq>](parTs sig);  
  20.773 +  (x, s)\<Colon>\<preceq>(G, L); 
  20.774 +  methd G declC sig = Some dm;  
  20.775 +  isrtype G statT;
  20.776 +  G\<turnstile>invC\<preceq>\<^sub>C declC; 
  20.777 +  G,s\<turnstile>a'\<Colon>\<preceq>RefT statT;  
  20.778 +  invmode (mhd sm) e = IntVir \<longrightarrow> a' \<noteq> Null; 
  20.779 +  invmode (mhd sm) e \<noteq> IntVir \<longrightarrow>  
  20.780 +       (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
  20.781 +    \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object);
  20.782 +  invC  = invocation_class (invmode (mhd sm) e) s a' statT;
  20.783 +  declC = invocation_declclass G (invmode (mhd sm) e) s a' statT sig;
  20.784 +  Ball (set (lcls (mbody (mthd dm)))) 
  20.785 +       (split (\<lambda>vn. is_type G)) 
  20.786 + \<rbrakk> \<Longrightarrow>  
  20.787 +  init_lvars G declC sig (invmode (mhd sm) e) a'  
  20.788 +  pvs (x,s)\<Colon>\<preceq>(G,\<lambda> k. 
  20.789 +                (case k of
  20.790 +                   EName e \<Rightarrow> (case e of 
  20.791 +                                 VNam v 
  20.792 +                                  \<Rightarrow> (table_of (lcls (mbody (mthd dm)))
  20.793 +                                        (pars (mthd dm)[\<mapsto>]parTs sig)) v
  20.794 +                               | Res \<Rightarrow> Some (resTy (mthd dm)))
  20.795 +                 | This \<Rightarrow> if (static (mthd sm)) 
  20.796 +                              then None else Some (Class declC)))"
  20.797 +apply (simp add: init_lvars_def2)
  20.798 +apply (rule conforms_set_locals)
  20.799 +apply  (simp (no_asm_simp) split add: split_if)
  20.800 +apply (drule  (4) DynT_conf)
  20.801 +apply clarsimp
  20.802 +(* apply intro *)
  20.803 +apply (drule (4) conforms_init_lvars_lemma)
  20.804 +apply (case_tac "dm",simp)
  20.805 +apply (rule conjI)
  20.806 +apply (unfold lconf_def, clarify)
  20.807 +apply (rule defval_conf1)
  20.808 +apply (clarsimp simp add: wf_mhead_def is_acc_type_def)
  20.809 +apply (case_tac "is_static sm")
  20.810 +apply simp_all
  20.811 +done
  20.812 +
  20.813 +
  20.814 +lemma Call_type_sound: "\<lbrakk>wf_prog G; G\<turnstile>(x1, s1) \<midarrow>ps\<doteq>\<succ>pvs\<rightarrow> (x2, s2);  
  20.815 + declC 
  20.816 + = invocation_declclass G (invmode (mhd esm) e) s2 a' statT \<lparr>name=mn,parTs=pTs\<rparr>;
  20.817 +s2'=init_lvars G declC \<lparr>name=mn,parTs=pTs\<rparr> (invmode (mhd esm) e) a' pvs (x2,s2);
  20.818 + G\<turnstile>s2' \<midarrow>Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<rightarrow> (x3, s3);    
  20.819 + \<forall>L. s2'\<Colon>\<preceq>(G, L) 
  20.820 +     \<longrightarrow> (\<forall>T. \<lparr>prg=G,cls=declC,lcl=L\<rparr>\<turnstile> Methd declC \<lparr>name=mn,parTs=pTs\<rparr>\<Colon>-T 
  20.821 +     \<longrightarrow> (x3, s3)\<Colon>\<preceq>(G, L) \<and> (x3 = None \<longrightarrow> G,s3\<turnstile>v\<Colon>\<preceq>T));  
  20.822 + Norm s0\<Colon>\<preceq>(G, L); 
  20.823 + \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>ps\<Colon>\<doteq>pTsa;  
  20.824 + max_spec G C statT \<lparr>name=mn,parTs=pTsa\<rparr> = {(esm, pTs)}; 
  20.825 + (x1, s1)\<Colon>\<preceq>(G, L); 
  20.826 + x1 = None \<longrightarrow> G,s1\<turnstile>a'\<Colon>\<preceq>RefT statT; (x2, s2)\<Colon>\<preceq>(G, L);  
  20.827 + x2 = None \<longrightarrow> list_all2 (conf G s2) pvs pTsa;
  20.828 + sm=(mhd esm)\<rbrakk> \<Longrightarrow>     
  20.829 + (x3, set_locals (locals s2) s3)\<Colon>\<preceq>(G, L) \<and> 
  20.830 + (x3 = None \<longrightarrow> G,s3\<turnstile>v\<Colon>\<preceq>resTy sm)"
  20.831 +apply clarify
  20.832 +apply (case_tac "x2")
  20.833 +defer
  20.834 +apply  (clarsimp split add: split_if_asm simp add: init_lvars_def2)
  20.835 +apply (case_tac "a' = Null \<and> \<not> (static (mhd esm)) \<and> e \<noteq> Super")
  20.836 +apply  (clarsimp simp add: init_lvars_def2)
  20.837 +apply clarsimp
  20.838 +apply (drule eval_gext')
  20.839 +apply (frule (1) conf_gext)
  20.840 +apply (drule max_spec2mheads, clarsimp)
  20.841 +apply (subgoal_tac "invmode (mhd esm) e = IntVir \<longrightarrow> a' \<noteq> Null")
  20.842 +defer  
  20.843 +apply  (clarsimp simp add: invmode_IntVir_eq)
  20.844 +apply (frule (6) DynT_mheadsD [OF DynT_propI,of _ "s2"],(rule HOL.refl)+)
  20.845 +apply clarify
  20.846 +apply (drule wf_mdeclD1, clarsimp) 
  20.847 +apply (frule  ty_expr_is_type) apply simp
  20.848 +apply (frule (2) conforms_init_lvars)
  20.849 +apply   simp
  20.850 +apply   assumption+
  20.851 +apply   simp
  20.852 +apply   assumption+
  20.853 +apply   clarsimp
  20.854 +apply   (rule HOL.refl)
  20.855 +apply   simp
  20.856 +apply   (rule Ball_weaken)
  20.857 +apply     assumption
  20.858 +apply     (force simp add: is_acc_type_def)
  20.859 +apply (tactic "smp_tac 1 1")
  20.860 +apply (frule (2) wt_MethdI, clarsimp)
  20.861 +apply (subgoal_tac "is_static dm = (static (mthd esm))") 
  20.862 +apply   (simp only:)
  20.863 +apply   (tactic "smp_tac 1 1")
  20.864 +apply   (rule conjI)
  20.865 +apply     (erule  conforms_return)
  20.866 +apply     blast
  20.867 +
  20.868 +apply     (force dest!: eval_gext del: impCE simp add: init_lvars_def2)
  20.869 +apply     clarsimp
  20.870 +apply     (drule (2) widen_trans, erule (1) conf_widen)
  20.871 +apply     (erule wf_ws_prog)
  20.872 +
  20.873 +apply   auto
  20.874 +done
  20.875 +
  20.876 +
  20.877 +subsection "accessibility"
  20.878 +
  20.879 +theorem dynamic_field_access_ok:
  20.880 +  (assumes wf: "wf_prog G" and
  20.881 +       eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2" and
  20.882 +     not_Null: "a\<noteq>Null" and
  20.883 +    conform_a: "G,(store s2)\<turnstile>a\<Colon>\<preceq> Class statC" and
  20.884 +   conform_s2: "s2\<Colon>\<preceq>(G, L)" and 
  20.885 +    normal_s2: "normal s2" and
  20.886 +         wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>,dt\<Turnstile>e\<Colon>-Class statC" and
  20.887 +            f: "accfield G accC statC fn = Some f" and
  20.888 +         dynC: "if stat then dynC=statC  
  20.889 +                        else dynC=obj_class (lookup_obj (store s2) a)"
  20.890 +  ) "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f) \<and> 
  20.891 +     G\<turnstile>Field fn f in dynC dyn_accessible_from accC"
  20.892 +proof (cases "stat")
  20.893 +  case True
  20.894 +  with dynC 
  20.895 +  have dynC': "dynC=statC" by simp
  20.896 +  with f
  20.897 +  have "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)"
  20.898 +    by (auto simp add: accfield_def Let_def intro!: table_of_remap_SomeD)
  20.899 +  with dynC' f
  20.900 +  show ?thesis
  20.901 +    by (auto intro!: static_to_dynamic_accessible_from
  20.902 +         dest: accfield_accessibleD accessible_from_commonD)
  20.903 +next
  20.904 +  case False
  20.905 +  with wf conform_a not_Null conform_s2 dynC
  20.906 +  obtain subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
  20.907 +    "is_class G dynC"
  20.908 +    by (auto dest!: conforms_RefTD [of _ _ _ _ "(fst s2)" L]
  20.909 +              dest: obj_ty_obj_class1
  20.910 +          simp add: obj_ty_obj_class )
  20.911 +  with wf f
  20.912 +  have "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)"
  20.913 +    by (auto simp add: accfield_def Let_def
  20.914 +                 dest: fields_mono
  20.915 +                dest!: table_of_remap_SomeD)
  20.916 +  moreover
  20.917 +  from f subclseq
  20.918 +  have "G\<turnstile>Field fn f in dynC dyn_accessible_from accC"
  20.919 +    by (auto intro!: static_to_dynamic_accessible_from 
  20.920 +               dest: accfield_accessibleD)
  20.921 +  ultimately show ?thesis
  20.922 +    by blast
  20.923 +qed
  20.924 +
  20.925 +lemma call_access_ok: 
  20.926 +(assumes invC_prop: "G\<turnstile>invmode (mhd statM) e\<rightarrow>invC\<preceq>statT" 
  20.927 +     and        wf: "wf_prog G" 
  20.928 +     and      wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
  20.929 +     and     statM: "statM \<in> mheads G accC statT sig" 
  20.930 +     and      invC: "invC = invocation_class (invmode (mhd statM) e) s a statT"
  20.931 +)"\<exists> dynM. dynlookup G statT invC sig = Some dynM \<and>
  20.932 +  G\<turnstile>Methd sig dynM in invC dyn_accessible_from accC"
  20.933 +proof -
  20.934 +  from wt_e wf have type_statT: "is_type G (RefT statT)"
  20.935 +    by (auto dest: ty_expr_is_type)
  20.936 +  from statM have not_Null: "statT \<noteq> NullT" by auto
  20.937 +  from type_statT wt_e 
  20.938 +  have wf_I: "(\<forall>I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> 
  20.939 +                                        invmode (mhd statM) e \<noteq> SuperM)"
  20.940 +    by (auto dest: invocationTypeExpr_noClassD)
  20.941 +  from wt_e
  20.942 +  have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode (mhd statM) e \<noteq> SuperM)"
  20.943 +    by (auto dest: invocationTypeExpr_noClassD)
  20.944 +  show ?thesis
  20.945 +  proof (cases "invmode (mhd statM) e = IntVir")
  20.946 +    case True
  20.947 +    with invC_prop not_Null
  20.948 +    have invC_prop': "is_class G invC \<and>  
  20.949 +                      (if (\<exists>T. statT=ArrayT T) then invC=Object 
  20.950 +                                              else G\<turnstile>Class invC\<preceq>RefT statT)"
  20.951 +      by (auto simp add: DynT_prop_def)
  20.952 +    with True not_Null
  20.953 +    have "G,statT \<turnstile> invC valid_lookup_cls_for is_static statM"
  20.954 +     by (cases statT) (auto simp add: invmode_def 
  20.955 +                         split: split_if split_if_asm) (*  was deleted above *)
  20.956 +    with statM type_statT wf 
  20.957 +    show ?thesis
  20.958 +      by - (rule dynlookup_access,auto)
  20.959 +  next
  20.960 +    case False
  20.961 +    with type_statT wf invC not_Null wf_I wf_A
  20.962 +    have invC_prop': " is_class G invC \<and>
  20.963 +                      ((\<exists> statC. statT=ClassT statC \<and> invC=statC) \<or>
  20.964 +                      (\<forall> statC. statT\<noteq>ClassT statC \<and> invC=Object)) "
  20.965 +        by (case_tac "statT") (auto simp add: invocation_class_def 
  20.966 +                                       split: inv_mode.splits)
  20.967 +    with not_Null wf
  20.968 +    have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig"
  20.969 +      by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_C_C
  20.970 +                                            dynimethd_def)
  20.971 +   from statM wf wt_e not_Null False invC_prop' obtain dynM where
  20.972 +                "accmethd G accC invC sig = Some dynM" 
  20.973 +     by (auto dest!: static_mheadsD)
  20.974 +   from invC_prop' False not_Null wf_I
  20.975 +   have "G,statT \<turnstile> invC valid_lookup_cls_for is_static statM"
  20.976 +     by (cases statT) (auto simp add: invmode_def
  20.977 +                        split: split_if split_if_asm) (*  was deleted above *)
  20.978 +   with statM type_statT wf 
  20.979 +    show ?thesis
  20.980 +      by - (rule dynlookup_access,auto)
  20.981 +  qed
  20.982 +qed
  20.983 +
  20.984 +section "main proof of type safety"
  20.985 +
  20.986 +ML {*
  20.987 +val forward_hyp_tac = EVERY' [smp_tac 1,
  20.988 +	FIRST'[mp_tac,etac exI,smp_tac 2,smp_tac 1,EVERY'[etac impE,etac exI]],
  20.989 +	REPEAT o (etac conjE)];
  20.990 +val typD_tac = eresolve_tac (thms "wt_elim_cases") THEN_ALL_NEW 
  20.991 +	EVERY' [full_simp_tac (simpset() setloop (K no_tac)), 
  20.992 +         clarify_tac(claset() addSEs[])]
  20.993 +*}
  20.994 +
  20.995 +lemma conforms_locals [rule_format]: 
  20.996 +  "(a,b)\<Colon>\<preceq>(G, L) \<longrightarrow> L x = Some T \<longrightarrow> G,b\<turnstile>the (locals b x)\<Colon>\<preceq>T"
  20.997 +apply (force simp: conforms_def Let_def lconf_def)
  20.998 +done
  20.999 +
 20.1000 +lemma eval_type_sound [rule_format (no_asm)]: 
 20.1001 + "wf_prog G \<Longrightarrow> G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1) \<Longrightarrow> (\<forall>L. s0\<Colon>\<preceq>(G,L) \<longrightarrow>    
 20.1002 +  (\<forall>C T. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<longrightarrow> s1\<Colon>\<preceq>(G,L) \<and>  
 20.1003 +  (let (x,s) = s1 in x = None \<longrightarrow> G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T)))"
 20.1004 +apply (erule eval_induct)
 20.1005 +
 20.1006 +(* 29 subgoals *)
 20.1007 +(* Xcpt, Inst, Methd, Nil, Skip, Expr, Comp *)
 20.1008 +apply         (simp_all (no_asm_use) add: Let_def body_def)
 20.1009 +apply       (tactic "ALLGOALS (EVERY'[Clarify_tac, TRY o typD_tac, 
 20.1010 +                     TRY o forward_hyp_tac])")
 20.1011 +apply      (tactic"ALLGOALS(EVERY'[asm_simp_tac(simpset()),TRY o Clarify_tac])")
 20.1012 +
 20.1013 +(* 20 subgoals *)
 20.1014 +
 20.1015 +(* Break *)
 20.1016 +apply (erule conforms_absorb)
 20.1017 +
 20.1018 +(* Cons *)
 20.1019 +apply (erule_tac V = "G\<turnstile>Norm s0 \<midarrow>?ea\<succ>\<rightarrow> ?vs1" in thin_rl)
 20.1020 +apply (frule eval_gext')
 20.1021 +apply force
 20.1022 +
 20.1023 +(* LVar *)
 20.1024 +apply (force elim: conforms_localD [THEN lconfD] conforms_lupd 
 20.1025 +       simp add: assign_conforms_def lvar_def)
 20.1026 +
 20.1027 +(* Cast *)
 20.1028 +apply (force dest: fits_conf)
 20.1029 +
 20.1030 +(* Lit *)
 20.1031 +apply (rule conf_litval)
 20.1032 +apply (simp add: empty_dt_def)
 20.1033 +
 20.1034 +(* Super *)
 20.1035 +apply (rule conf_widen)
 20.1036 +apply   (erule (1) subcls_direct [THEN widen.subcls])
 20.1037 +apply  (erule (1) conforms_localD [THEN lconfD])
 20.1038 +apply (erule wf_ws_prog)
 20.1039 +
 20.1040 +(* Acc *)
 20.1041 +apply fast
 20.1042 +
 20.1043 +(* Body *)
 20.1044 +apply (rule conjI)
 20.1045 +apply (rule conforms_absorb)
 20.1046 +apply (fast)
 20.1047 +apply (fast intro: conforms_locals)
 20.1048 +
 20.1049 +(* Cond *)
 20.1050 +apply (simp split: split_if_asm)
 20.1051 +apply  (tactic "forward_hyp_tac 1", force)
 20.1052 +apply (tactic "forward_hyp_tac 1", force)
 20.1053 +
 20.1054 +(* If *)
 20.1055 +apply (force split add: split_if_asm)
 20.1056 +
 20.1057 +(* Loop *)
 20.1058 +apply (drule (1) wt.Loop)
 20.1059 +apply (clarsimp split: split_if_asm)
 20.1060 +apply (fast intro: conforms_absorb)
 20.1061 +
 20.1062 +(* Fin *)
 20.1063 +apply (case_tac "x1", force)
 20.1064 +apply (drule spec, erule impE, erule conforms_NormI)
 20.1065 +apply (erule impE)
 20.1066 +apply   blast
 20.1067 +apply (clarsimp)
 20.1068 +apply (erule (3) Fin_lemma)
 20.1069 +
 20.1070 +(* Throw *)
 20.1071 +apply (erule (3) Throw_lemma)
 20.1072 +
 20.1073 +(* NewC *)
 20.1074 +apply (clarsimp simp add: is_acc_class_def)
 20.1075 +apply (drule (1) halloc_type_sound,blast, rule HOL.refl, simp, simp)
 20.1076 +
 20.1077 +(* NewA *)
 20.1078 +apply (tactic "smp_tac 1 1",frule wt_init_comp_ty,erule impE,blast)
 20.1079 +apply (tactic "forward_hyp_tac 1")
 20.1080 +apply (case_tac "check_neg i' ab")
 20.1081 +apply  (clarsimp simp add: is_acc_type_def)
 20.1082 +apply  (drule (2) halloc_type_sound, rule HOL.refl, simp, simp)
 20.1083 +apply force
 20.1084 +
 20.1085 +(* Level 34, 6 subgoals *)
 20.1086 +
 20.1087 +(* Init *)
 20.1088 +apply (case_tac "inited C (globs s0)")
 20.1089 +apply  (clarsimp)
 20.1090 +apply (clarsimp)
 20.1091 +apply (frule (1) wf_prog_cdecl)
 20.1092 +apply (drule spec, erule impE, erule (3) conforms_init_class_obj)
 20.1093 +apply (drule_tac "psi" = "class G C = ?x" in asm_rl,erule impE,
 20.1094 +      force dest!: wf_cdecl_supD split add: split_if simp add: is_acc_class_def)
 20.1095 +apply (drule spec, erule impE, erule conforms_set_locals, rule lconf_empty)
 20.1096 +apply (erule impE) apply (rule exI) apply (erule wf_cdecl_wt_init)
 20.1097 +apply (drule (1) conforms_return, force dest: eval_gext', assumption)
 20.1098 +
 20.1099 +
 20.1100 +(* Ass *)
 20.1101 +apply (tactic "forward_hyp_tac 1")
 20.1102 +apply (rename_tac x1 s1 x2 s2 v va w L C Ta T', case_tac x1)
 20.1103 +prefer 2 apply force
 20.1104 +apply (case_tac x2)
 20.1105 +prefer 2 apply force
 20.1106 +apply (simp, drule conjunct2)
 20.1107 +apply (drule (1) conf_widen)
 20.1108 +apply  (erule wf_ws_prog)
 20.1109 +apply (erule (2) Ass_lemma)
 20.1110 +apply (clarsimp simp add: assign_conforms_def)
 20.1111 +
 20.1112 +(* Try *)
 20.1113 +apply (drule (1) sxalloc_type_sound, simp (no_asm_use))
 20.1114 +apply (case_tac a)
 20.1115 +apply  clarsimp
 20.1116 +apply clarsimp
 20.1117 +apply (tactic "smp_tac 1 1")
 20.1118 +apply (simp split add: split_if_asm)
 20.1119 +apply (fast dest: conforms_deallocL Try_lemma)
 20.1120 +
 20.1121 +(* FVar *)
 20.1122 +
 20.1123 +apply (frule accfield_fields)
 20.1124 +apply (frule ty_expr_is_type [THEN type_is_class],simp)
 20.1125 +apply simp
 20.1126 +apply (frule wf_ws_prog)
 20.1127 +apply (frule (1) fields_declC,simp)
 20.1128 +apply clarsimp 
 20.1129 +(*b y EVERY'[datac cfield_defpl_is_class 2, Clarsimp_tac] 1; not useful here*)
 20.1130 +apply (tactic "smp_tac 1 1")
 20.1131 +apply (tactic "forward_hyp_tac 1")
 20.1132 +apply (rule conjI, force split add: split_if simp add: fvar_def2)
 20.1133 +apply (drule init_yields_initd, frule eval_gext')
 20.1134 +apply clarsimp
 20.1135 +apply (case_tac "C=Object")
 20.1136 +apply  clarsimp
 20.1137 +apply (erule (9) FVar_lemma)
 20.1138 +
 20.1139 +(* AVar *)
 20.1140 +apply (tactic "forward_hyp_tac 1")
 20.1141 +apply (erule_tac V = "G\<turnstile>Norm s0 \<midarrow>?e1-\<succ>?a'\<rightarrow> (?x1 1, ?s1)" in thin_rl, 
 20.1142 +         frule eval_gext')
 20.1143 +apply (rule conjI)
 20.1144 +apply  (clarsimp simp add: avar_def2)
 20.1145 +apply clarsimp
 20.1146 +apply (erule (5) AVar_lemma)
 20.1147 +
 20.1148 +(* Call *)
 20.1149 +apply (tactic "forward_hyp_tac 1")
 20.1150 +apply (rule Call_type_sound)
 20.1151 +apply auto
 20.1152 +done
 20.1153 +
 20.1154 +
 20.1155 +declare fun_upd_apply [simp]
 20.1156 +declare split_paired_All [simp] split_paired_Ex [simp]
 20.1157 +declare split_if     [split] split_if_asm     [split] 
 20.1158 +        option.split [split] option.split_asm [split]
 20.1159 +ML_setup {* 
 20.1160 +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac);
 20.1161 +claset_ref()  := claset () addSbefore ("split_all_tac", split_all_tac)
 20.1162 +*}
 20.1163 +
 20.1164 +theorem eval_ts: 
 20.1165 + "\<lbrakk>G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow> (x',s'); wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-T\<rbrakk> 
 20.1166 +\<Longrightarrow>  (x',s')\<Colon>\<preceq>(G,L) \<and> (x'=None \<longrightarrow> G,s'\<turnstile>v\<Colon>\<preceq>T)"
 20.1167 +apply (drule (3) eval_type_sound)
 20.1168 +apply (unfold Let_def)
 20.1169 +apply clarsimp
 20.1170 +done
 20.1171 +
 20.1172 +theorem evals_ts: 
 20.1173 +"\<lbrakk>G\<turnstile>s \<midarrow>es\<doteq>\<succ>vs\<rightarrow> (x',s'); wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>es\<Colon>\<doteq>Ts\<rbrakk> 
 20.1174 +\<Longrightarrow>  (x',s')\<Colon>\<preceq>(G,L) \<and> (x'=None \<longrightarrow> list_all2 (conf G s') vs Ts)"
 20.1175 +apply (drule (3) eval_type_sound)
 20.1176 +apply (unfold Let_def)
 20.1177 +apply clarsimp
 20.1178 +done
 20.1179 +
 20.1180 +theorem evar_ts: 
 20.1181 +"\<lbrakk>G\<turnstile>s \<midarrow>v=\<succ>vf\<rightarrow> (x',s'); wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>v\<Colon>=T\<rbrakk> \<Longrightarrow>  
 20.1182 +  (x',s')\<Colon>\<preceq>(G,L) \<and> (x'=None \<longrightarrow> G,L,s'\<turnstile>In2 v\<succ>In2 vf\<Colon>\<preceq>Inl T)"
 20.1183 +apply (drule (3) eval_type_sound)
 20.1184 +apply (unfold Let_def)
 20.1185 +apply clarsimp
 20.1186 +done
 20.1187 +
 20.1188 +theorem exec_ts: 
 20.1189 +"\<lbrakk>G\<turnstile>s \<midarrow>s0\<rightarrow> s'; wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>s0\<Colon>\<surd>\<rbrakk> \<Longrightarrow> s'\<Colon>\<preceq>(G,L)"
 20.1190 +apply (drule (3) eval_type_sound)
 20.1191 +apply (unfold Let_def)
 20.1192 +apply clarsimp
 20.1193 +done
 20.1194 +
 20.1195 +(*
 20.1196 +theorem dyn_methods_understood: 
 20.1197 + "\<And>s. \<lbrakk>wf_prog G; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>{t,md,IntVir}e..mn({pTs'}ps)\<Colon>-rT;  
 20.1198 +  s\<Colon>\<preceq>(G,L); G\<turnstile>s \<midarrow>e-\<succ>a'\<rightarrow> Norm s'; a' \<noteq> Null\<rbrakk> \<Longrightarrow>  
 20.1199 +  \<exists>a obj. a'=Addr a \<and> heap s' a = Some obj \<and> 
 20.1200 +  cmethd G (obj_class obj) (mn, pTs') \<noteq> None"
 20.1201 +apply (erule wt_elim_cases)
 20.1202 +apply (drule max_spec2mheads)
 20.1203 +apply (drule (3) eval_ts)
 20.1204 +apply (clarsimp split del: split_if split_if_asm)
 20.1205 +apply (drule (2) DynT_propI)
 20.1206 +apply  (simp (no_asm_simp))
 20.1207 +apply (tactic *) (* {* exhaust_cmethd_tac "the (cmethd G (target (invmode m e) s' a' md) (mn, pTs'))" 1 *} *)(*)
 20.1208 +apply (drule (4) DynT_mheadsD [THEN conjunct1], rule HOL.refl)
 20.1209 +apply (drule conf_RefTD)
 20.1210 +apply clarsimp
 20.1211 +done 
 20.1212 +*)
 20.1213 +
 20.1214 +end
    21.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    21.2 +++ b/src/HOL/Bali/Value.thy	Mon Jan 28 17:00:19 2002 +0100
    21.3 @@ -0,0 +1,51 @@
    21.4 +(*  Title:      isabelle/Bali/Value.thy
    21.5 +    ID:         $Id$
    21.6 +    Author:     David von Oheimb
    21.7 +    Copyright   1997 Technische Universitaet Muenchen
    21.8 +*)
    21.9 +header {* Java values *}
   21.10 +
   21.11 +
   21.12 +
   21.13 +theory Value = Type:
   21.14 +
   21.15 +typedecl loc            (* locations, i.e. abstract references on objects *)
   21.16 +arities	 loc :: "type"
   21.17 +
   21.18 +datatype val
   21.19 +	= Unit          (* dummy result value of void methods *)
   21.20 +	| Bool bool     (* Boolean value *)
   21.21 +	| Intg int      (* integer value *)
   21.22 +	| Null          (* null reference *)
   21.23 +	| Addr loc      (* addresses, i.e. locations of objects *)
   21.24 +
   21.25 +
   21.26 +translations "val" <= (type) "Term.val"
   21.27 +             "loc" <= (type) "Term.loc"
   21.28 +
   21.29 +consts   the_Bool   :: "val \<Rightarrow> bool"  
   21.30 +primrec "the_Bool (Bool b) = b"
   21.31 +consts   the_Intg   :: "val \<Rightarrow> int"
   21.32 +primrec "the_Intg (Intg i) = i"
   21.33 +consts   the_Addr   :: "val \<Rightarrow> loc"
   21.34 +primrec "the_Addr (Addr a) = a"
   21.35 +
   21.36 +types	dyn_ty      = "loc \<Rightarrow> ty option"
   21.37 +consts
   21.38 +  typeof        :: "dyn_ty \<Rightarrow> val \<Rightarrow> ty option"
   21.39 +  defpval       :: "prim_ty \<Rightarrow> val"      (* default value for primitive types *)
   21.40 +  default_val   :: "     ty \<Rightarrow> val"      (* default value for all types *)
   21.41 +
   21.42 +primrec "typeof dt  Unit    = Some (PrimT Void)"
   21.43 +	"typeof dt (Bool b) = Some (PrimT Boolean)"
   21.44 +	"typeof dt (Intg i) = Some (PrimT Integer)"
   21.45 +	"typeof dt  Null    = Some NT"
   21.46 +	"typeof dt (Addr a) = dt a"
   21.47 +
   21.48 +primrec	"defpval Void    = Unit"
   21.49 +	"defpval Boolean = Bool False"
   21.50 +	"defpval Integer = Intg 0"
   21.51 +primrec	"default_val (PrimT pt) = defpval pt"
   21.52 +	"default_val (RefT  r ) = Null"
   21.53 +
   21.54 +end
    22.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    22.2 +++ b/src/HOL/Bali/WellForm.thy	Mon Jan 28 17:00:19 2002 +0100
    22.3 @@ -0,0 +1,3004 @@
    22.4 +(*  Title:      isabelle/Bali/WellForm.thy
    22.5 +    ID:         $Id$
    22.6 +    Author:     David von Oheimb
    22.7 +    Copyright   1997 Technische Universitaet Muenchen
    22.8 +*)
    22.9 +
   22.10 +header {* Well-formedness of Java programs *}
   22.11 +
   22.12 +theory WellForm = WellType:
   22.13 +
   22.14 +text {*
   22.15 +For static checks on expressions and statements, see WellType.thy
   22.16 +
   22.17 +improvements over Java Specification 1.0 (cf. 8.4.6.3, 8.4.6.4, 9.4.1):
   22.18 +\begin{itemize}
   22.19 +\item a method implementing or overwriting another method may have a result 
   22.20 +      type that widens to the result type of the other method 
   22.21 +      (instead of identical type)
   22.22 +\item if a method hides another method (both methods have to be static!)
   22.23 +  there are no restrictions to the result type 
   22.24 +  since the methods have to be static and there is no dynamic binding of 
   22.25 +  static methods
   22.26 +\item if an interface inherits more than one method with the same signature, the
   22.27 +  methods need not have identical return types
   22.28 +\end{itemize}
   22.29 +simplifications:
   22.30 +\begin{itemize}
   22.31 +\item Object and standard exceptions are assumed to be declared like normal 
   22.32 +      classes
   22.33 +\end{itemize}
   22.34 +*}
   22.35 +
   22.36 +section "well-formed field declarations"
   22.37 +  (* well-formed field declaration (common part for classes and interfaces),
   22.38 +     cf. 8.3 and (9.3) *)
   22.39 +
   22.40 +constdefs
   22.41 +  wf_fdecl :: "prog \<Rightarrow> pname \<Rightarrow> fdecl \<Rightarrow> bool"
   22.42 + "wf_fdecl G P \<equiv> \<lambda>(fn,f). is_acc_type G P (type f)"
   22.43 +
   22.44 +lemma wf_fdecl_def2: "\<And>fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))"
   22.45 +apply (unfold wf_fdecl_def)
   22.46 +apply simp
   22.47 +done
   22.48 +
   22.49 +
   22.50 +
   22.51 +section "well-formed method declarations"
   22.52 +  (*well-formed method declaration,cf. 8.4, 8.4.1, 8.4.3, 8.4.5, 14.3.2, (9.4)*)
   22.53 +  (* cf. 14.15, 15.7.2, for scope issues cf. 8.4.1 and 14.3.2 *)
   22.54 +
   22.55 +text {*
   22.56 +A method head is wellformed if:
   22.57 +\begin{itemize}
   22.58 +\item the signature and the method head agree in the number of parameters
   22.59 +\item all types of the parameters are visible
   22.60 +\item the result type is visible
   22.61 +\item the parameter names are unique
   22.62 +\end{itemize} 
   22.63 +*}
   22.64 +constdefs
   22.65 +  wf_mhead :: "prog \<Rightarrow> pname \<Rightarrow> sig \<Rightarrow> mhead \<Rightarrow> bool"
   22.66 + "wf_mhead G P \<equiv> \<lambda> sig mh. length (parTs sig) = length (pars mh) \<and>
   22.67 +			    \<spacespace> ( \<forall>T\<in>set (parTs sig). is_acc_type G P T) \<and> 
   22.68 +                            is_acc_type G P (resTy mh) \<and>
   22.69 +			    \<spacespace> nodups (pars mh)"
   22.70 +
   22.71 +
   22.72 +text {*
   22.73 +A method declaration is wellformed if:
   22.74 +\begin{itemize}
   22.75 +\item the method head is wellformed
   22.76 +\item the names of the local variables are unique
   22.77 +\item the types of the local variables must be accessible
   22.78 +\item the local variables don't shadow the parameters
   22.79 +\item the class of the method is defined
   22.80 +\item the body statement is welltyped with respect to the
   22.81 +      modified environment of local names, were the local variables, 
   22.82 +      the parameters the special result variable (Res) and This are assoziated
   22.83 +      with there types. 
   22.84 +\end{itemize}
   22.85 +*}
   22.86 +
   22.87 +constdefs
   22.88 +  wf_mdecl :: "prog \<Rightarrow> qtname \<Rightarrow> mdecl \<Rightarrow> bool"
   22.89 + "wf_mdecl G C \<equiv> 
   22.90 +      \<lambda>(sig,m).
   22.91 +	  wf_mhead G (pid C) sig (mhead m) \<and> 
   22.92 +          unique (lcls (mbody m)) \<and> 
   22.93 +          (\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T) \<and> 
   22.94 +	  (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and>
   22.95 +          is_class G C \<and>
   22.96 +          \<lparr>prg=G,cls=C,
   22.97 +           lcl=\<lambda> k. 
   22.98 +               (case k of
   22.99 +                  EName e 
  22.100 +                  \<Rightarrow> (case e of 
  22.101 +                        VNam v 
  22.102 +                        \<Rightarrow>(table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
  22.103 +                      | Res \<Rightarrow> Some (resTy m))
  22.104 +	        | This \<Rightarrow> if static m then None else Some (Class C))
  22.105 +          \<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>"
  22.106 +
  22.107 +lemma wf_mheadI: 
  22.108 +"\<lbrakk>length (parTs sig) = length (pars m); \<forall>T\<in>set (parTs sig). is_acc_type G P T;
  22.109 +  is_acc_type G P (resTy m); nodups (pars m)\<rbrakk> \<Longrightarrow>  
  22.110 +  wf_mhead G P sig m"
  22.111 +apply (unfold wf_mhead_def)
  22.112 +apply (simp (no_asm_simp))
  22.113 +done
  22.114 +
  22.115 +lemma wf_mdeclI: "\<lbrakk>  
  22.116 +  wf_mhead G (pid C) sig (mhead m); unique (lcls (mbody m));  
  22.117 +  (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None); 
  22.118 +  \<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T;
  22.119 +  is_class G C;
  22.120 +  \<lparr>prg=G,cls=C,
  22.121 +   lcl=\<lambda> k. 
  22.122 +       (case k of
  22.123 +          EName e 
  22.124 +          \<Rightarrow> (case e of 
  22.125 +                VNam v 
  22.126 +                \<Rightarrow> (table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
  22.127 +              | Res \<Rightarrow> Some (resTy m))
  22.128 +        | This \<Rightarrow> if static m then None else Some (Class C))
  22.129 +  \<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>
  22.130 +  \<rbrakk> \<Longrightarrow>  
  22.131 +  wf_mdecl G C (sig,m)"
  22.132 +apply (unfold wf_mdecl_def)
  22.133 +apply simp
  22.134 +done
  22.135 +
  22.136 +
  22.137 +lemma wf_mdeclD1: 
  22.138 +"wf_mdecl G C (sig,m) \<Longrightarrow>  
  22.139 +   wf_mhead G (pid C) sig (mhead m) \<and> unique (lcls (mbody m)) \<and>  
  22.140 +  (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and> 
  22.141 +  (\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T)"
  22.142 +apply (unfold wf_mdecl_def)
  22.143 +apply auto
  22.144 +done
  22.145 +
  22.146 +lemma wf_mdecl_bodyD: 
  22.147 +"wf_mdecl G C (sig,m) \<Longrightarrow>  
  22.148 + (\<exists>T. \<lparr>prg=G,cls=C,
  22.149 +       lcl = \<lambda> k. 
  22.150 +         (case k of
  22.151 +            EName e 
  22.152 +            \<Rightarrow> (case e of 
  22.153 +                VNam v \<Rightarrow> (table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
  22.154 +                | Res  \<Rightarrow> Some (resTy m))
  22.155 +          | This \<Rightarrow> if static m then None else Some (Class C))
  22.156 +       \<rparr>\<turnstile>Body C (stmt (mbody m))\<Colon>-T \<and> G\<turnstile>T\<preceq>(resTy m))"
  22.157 +apply (unfold wf_mdecl_def)
  22.158 +apply clarify
  22.159 +apply (rule_tac x="(resTy m)" in exI)
  22.160 +apply (unfold wf_mhead_def)
  22.161 +apply (auto simp add: wf_mhead_def is_acc_type_def intro: wt.Body )
  22.162 +done
  22.163 +
  22.164 +
  22.165 +(*
  22.166 +lemma static_Object_methodsE [elim!]: 
  22.167 + "\<lbrakk>wf_mdecl G Object (sig, m);static m\<rbrakk> \<Longrightarrow> R"
  22.168 +apply (unfold wf_mdecl_def)
  22.169 +apply auto
  22.170 +done
  22.171 +*)
  22.172 +
  22.173 +lemma rT_is_acc_type: 
  22.174 +  "wf_mhead G P sig m \<Longrightarrow> is_acc_type G P (resTy m)"
  22.175 +apply (unfold wf_mhead_def)
  22.176 +apply auto
  22.177 +done
  22.178 +
  22.179 +section "well-formed interface declarations"
  22.180 +  (* well-formed interface declaration, cf. 9.1, 9.1.2.1, 9.1.3, 9.4 *)
  22.181 +
  22.182 +text {*
  22.183 +A interface declaration is wellformed if:
  22.184 +\begin{itemize}
  22.185 +\item the interface hierarchy is wellstructured
  22.186 +\item there is no class with the same name
  22.187 +\item the method heads are wellformed and not static and have Public access
  22.188 +\item the methods are uniquely named
  22.189 +\item all superinterfaces are accessible
  22.190 +\item the result type of a method overriding a method of Object widens to the
  22.191 +      result type of the overridden method.
  22.192 +      Shadowing static methods is forbidden.
  22.193 +\item the result type of a method overriding a set of methods defined in the
  22.194 +      superinterfaces widens to each of the corresponding result types
  22.195 +\end{itemize}
  22.196 +*}
  22.197 +constdefs
  22.198 +  wf_idecl :: "prog  \<Rightarrow> idecl \<Rightarrow> bool"
  22.199 + "wf_idecl G \<equiv> 
  22.200 +    \<lambda>(I,i). 
  22.201 +        ws_idecl G I (isuperIfs i) \<and> 
  22.202 +	\<not>is_class G I \<and>
  22.203 +	(\<forall>(sig,mh)\<in>set (imethods i). wf_mhead G (pid I) sig mh \<and> 
  22.204 +                                     \<not>is_static mh \<and>
  22.205 +                                      accmodi mh = Public) \<and>
  22.206 +	unique (imethods i) \<and>
  22.207 +        (\<forall> J\<in>set (isuperIfs i). is_acc_iface G (pid I) J) \<and>
  22.208 +        (table_of (imethods i)
  22.209 +          hiding (methd G Object)
  22.210 +          under  (\<lambda> new old. accmodi old \<noteq> Private)
  22.211 +          entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old \<and> 
  22.212 +                             is_static new = is_static old)) \<and> 
  22.213 +        (o2s \<circ> table_of (imethods i) 
  22.214 +               hidings Un_tables((\<lambda>J.(imethds G J))`set (isuperIfs i))
  22.215 +	       entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old))"
  22.216 +
  22.217 +lemma wf_idecl_mhead: "\<lbrakk>wf_idecl G (I,i); (sig,mh)\<in>set (imethods i)\<rbrakk> \<Longrightarrow>  
  22.218 +  wf_mhead G (pid I) sig mh \<and> \<not>is_static mh \<and> accmodi mh = Public"
  22.219 +apply (unfold wf_idecl_def)
  22.220 +apply auto
  22.221 +done
  22.222 +
  22.223 +lemma wf_idecl_hidings: 
  22.224 +"wf_idecl G (I, i) \<Longrightarrow> 
  22.225 +  (\<lambda>s. o2s (table_of (imethods i) s)) 
  22.226 +  hidings Un_tables ((\<lambda>J. imethds G J) ` set (isuperIfs i))  
  22.227 +  entails \<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old"
  22.228 +apply (unfold wf_idecl_def o_def)
  22.229 +apply simp
  22.230 +done
  22.231 +
  22.232 +lemma wf_idecl_hiding:
  22.233 +"wf_idecl G (I, i) \<Longrightarrow> 
  22.234 + (table_of (imethods i)
  22.235 +           hiding (methd G Object)
  22.236 +           under  (\<lambda> new old. accmodi old \<noteq> Private)
  22.237 +           entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old \<and> 
  22.238 +                              is_static new = is_static old))"
  22.239 +apply (unfold wf_idecl_def)
  22.240 +apply simp
  22.241 +done
  22.242 +
  22.243 +lemma wf_idecl_supD: 
  22.244 +"\<lbrakk>wf_idecl G (I,i); J \<in> set (isuperIfs i)\<rbrakk> 
  22.245 + \<Longrightarrow> is_acc_iface G (pid I) J \<and> (J, I) \<notin> (subint1 G)^+"
  22.246 +apply (unfold wf_idecl_def ws_idecl_def)
  22.247 +apply auto
  22.248 +done
  22.249 +
  22.250 +section "well-formed class declarations"
  22.251 +  (* well-formed class declaration, cf. 8.1, 8.1.2.1, 8.1.2.2, 8.1.3, 8.1.4 and
  22.252 +   class method declaration, cf. 8.4.3.3, 8.4.6.1, 8.4.6.2, 8.4.6.3, 8.4.6.4 *)
  22.253 +
  22.254 +text {*
  22.255 +A class declaration is wellformed if:
  22.256 +\begin{itemize}
  22.257 +\item there is no interface with the same name
  22.258 +\item all superinterfaces are accessible and for all methods implementing 
  22.259 +      an interface method the result type widens to the result type of 
  22.260 +      the interface method, the method is not static and offers at least 
  22.261 +      as much access 
  22.262 +      (this actually means that the method has Public access, since all 
  22.263 +      interface methods have public access)
  22.264 +\item all field declarations are wellformed and the field names are unique
  22.265 +\item all method declarations are wellformed and the method names are unique
  22.266 +\item the initialization statement is welltyped
  22.267 +\item the classhierarchy is wellstructured
  22.268 +\item Unless the class is Object:
  22.269 +      \begin{itemize}
  22.270 +      \item the superclass is accessible
  22.271 +      \item for all methods overriding another method (of a superclass )the
  22.272 +            result type widens to the result type of the overridden method,
  22.273 +            the access modifier of the new method provides at least as much
  22.274 +            access as the overwritten one.
  22.275 +      \item for all methods hiding a method (of a superclass) the hidden 
  22.276 +            method must be static and offer at least as much access rights.
  22.277 +            Remark: In contrast to the Java Language Specification we don't
  22.278 +            restrict the result types of the method
  22.279 +            (as in case of overriding), because there seems to be no reason,
  22.280 +            since there is no dynamic binding of static methods.
  22.281 +            (cf. 8.4.6.3 vs. 15.12.1).
  22.282 +            Stricly speaking the restrictions on the access rights aren't 
  22.283 +            necessary to, since the static type and the access rights 
  22.284 +            together determine which method is to be called statically. 
  22.285 +            But if a class gains more then one static method with the 
  22.286 +            same signature due to inheritance, it is confusing when the 
  22.287 +            method selection depends on the access rights only: 
  22.288 +            e.g.
  22.289 +              Class C declares static public method foo().
  22.290 +              Class D is subclass of C and declares static method foo()
  22.291 +              with default package access.
  22.292 +              D.foo() ? if this call is in the same package as D then
  22.293 +                        foo of class D is called, otherwise foo of class C.
  22.294 +      \end{itemize}
  22.295 +
  22.296 +\end{itemize}
  22.297 +*}
  22.298 +(* to Table *)
  22.299 +constdefs entails:: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool"
  22.300 +                                 ("_ entails _" 20)
  22.301 +"t entails P \<equiv> \<forall>k. \<forall> x \<in> t k: P x"
  22.302 +
  22.303 +lemma entailsD:
  22.304 + "\<lbrakk>t entails P; t k = Some x\<rbrakk> \<Longrightarrow> P x"
  22.305 +by (simp add: entails_def)
  22.306 +
  22.307 +lemma empty_entails[simp]: "empty entails P"
  22.308 +by (simp add: entails_def)
  22.309 +
  22.310 +constdefs
  22.311 + wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool"
  22.312 +"wf_cdecl G \<equiv> 
  22.313 +   \<lambda>(C,c).
  22.314 +      \<not>is_iface G C \<and>
  22.315 +      (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
  22.316 +        (\<forall>s. \<forall> im \<in> imethds G I s.
  22.317 +      	    (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
  22.318 +      	                             \<not> is_static cm \<and>
  22.319 +                                     accmodi im \<le> accmodi cm))) \<and>
  22.320 +      (\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and> 
  22.321 +      (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and> 
  22.322 +      \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd> \<and> ws_cdecl G C (super c) \<and>
  22.323 +      (C \<noteq> Object \<longrightarrow> 
  22.324 +            (is_acc_class G (pid C) (super c) \<and>
  22.325 +            (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
  22.326 +             entails (\<lambda> new. \<forall> old sig. 
  22.327 +                       (G,sig\<turnstile>new overrides\<^sub>S old 
  22.328 +                        \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
  22.329 +                             accmodi old \<le> accmodi new \<and>
  22.330 +      	                     \<not>is_static old)) \<and>
  22.331 +                       (G,sig\<turnstile>new hides old 
  22.332 +                         \<longrightarrow> (accmodi old \<le> accmodi new \<and>
  22.333 +      	                      is_static old)))) 
  22.334 +            ))"
  22.335 +
  22.336 +(*
  22.337 +constdefs
  22.338 + wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool"
  22.339 +"wf_cdecl G \<equiv> 
  22.340 +   \<lambda>(C,c).
  22.341 +      \<not>is_iface G C \<and>
  22.342 +      (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
  22.343 +        (\<forall>s. \<forall> im \<in> imethds G I s.
  22.344 +      	    (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy (mthd cm)\<preceq>resTy (mthd im) \<and>
  22.345 +      	                             \<not> is_static cm \<and>
  22.346 +                                     accmodi im \<le> accmodi cm))) \<and>
  22.347 +      (\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and> 
  22.348 +      (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and> 
  22.349 +      \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd> \<and> ws_cdecl G C (super c) \<and>
  22.350 +      (C \<noteq> Object \<longrightarrow> 
  22.351 +            (is_acc_class G (pid C) (super c) \<and>
  22.352 +            (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
  22.353 +              hiding methd G (super c)
  22.354 +              under (\<lambda> new old. G\<turnstile>new overrides old)
  22.355 +              entails (\<lambda> new old. 
  22.356 +                           (G\<turnstile>resTy (mthd new)\<preceq>resTy (mthd old) \<and>
  22.357 +                            accmodi old \<le> accmodi new \<and>
  22.358 +      	                   \<not> is_static old)))  \<and>
  22.359 +            (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
  22.360 +              hiding methd G (super c)
  22.361 +              under (\<lambda> new old. G\<turnstile>new hides old)
  22.362 +              entails (\<lambda> new old. is_static old \<and> 
  22.363 +                                  accmodi old \<le> accmodi new))  \<and>
  22.364 +            (table_of (cfields c) hiding accfield G C (super c)
  22.365 +              entails (\<lambda> newF oldF. accmodi oldF \<le> access newF))))"
  22.366 +*)
  22.367 +
  22.368 +lemma wf_cdecl_unique: 
  22.369 +"wf_cdecl G (C,c) \<Longrightarrow> unique (cfields c) \<and> unique (methods c)"
  22.370 +apply (unfold wf_cdecl_def)
  22.371 +apply auto
  22.372 +done
  22.373 +
  22.374 +lemma wf_cdecl_fdecl: 
  22.375 +"\<lbrakk>wf_cdecl G (C,c); f\<in>set (cfields c)\<rbrakk> \<Longrightarrow> wf_fdecl G (pid C) f"
  22.376 +apply (unfold wf_cdecl_def)
  22.377 +apply auto
  22.378 +done
  22.379 +
  22.380 +lemma wf_cdecl_mdecl: 
  22.381 +"\<lbrakk>wf_cdecl G (C,c); m\<in>set (methods c)\<rbrakk> \<Longrightarrow> wf_mdecl G C m"
  22.382 +apply (unfold wf_cdecl_def)
  22.383 +apply auto
  22.384 +done
  22.385 +
  22.386 +lemma wf_cdecl_impD: 
  22.387 +"\<lbrakk>wf_cdecl G (C,c); I\<in>set (superIfs c)\<rbrakk> 
  22.388 +\<Longrightarrow> is_acc_iface G (pid C) I \<and>  
  22.389 +    (\<forall>s. \<forall>im \<in> imethds G I s.  
  22.390 +        (\<exists>cm \<in> methd G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and> \<not>is_static cm \<and>
  22.391 +                                   accmodi im \<le> accmodi cm))"
  22.392 +apply (unfold wf_cdecl_def)
  22.393 +apply auto
  22.394 +done
  22.395 +
  22.396 +lemma wf_cdecl_supD: 
  22.397 +"\<lbrakk>wf_cdecl G (C,c); C \<noteq> Object\<rbrakk> \<Longrightarrow>  
  22.398 +  is_acc_class G (pid C) (super c) \<and> (super c,C) \<notin> (subcls1 G)^+ \<and> 
  22.399 +   (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
  22.400 +    entails (\<lambda> new. \<forall> old sig. 
  22.401 +                 (G,sig\<turnstile>new overrides\<^sub>S old 
  22.402 +                  \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
  22.403 +                       accmodi old \<le> accmodi new \<and>
  22.404 +                       \<not>is_static old)) \<and>
  22.405 +                 (G,sig\<turnstile>new hides old 
  22.406 +                   \<longrightarrow> (accmodi old \<le> accmodi new \<and>
  22.407 +                        is_static old))))"
  22.408 +apply (unfold wf_cdecl_def ws_cdecl_def)
  22.409 +apply auto
  22.410 +done
  22.411 +
  22.412 +
  22.413 +lemma wf_cdecl_overrides_SomeD:
  22.414 +"\<lbrakk>wf_cdecl G (C,c); C \<noteq> Object; table_of (methods c) sig = Some newM;
  22.415 +  G,sig\<turnstile>(C,newM) overrides\<^sub>S old
  22.416 +\<rbrakk> \<Longrightarrow>  G\<turnstile>resTy newM\<preceq>resTy old \<and>
  22.417 +       accmodi old \<le> accmodi newM \<and>
  22.418 +       \<not> is_static old" 
  22.419 +apply (drule (1) wf_cdecl_supD)
  22.420 +apply (clarify)
  22.421 +apply (drule entailsD)
  22.422 +apply   (blast intro: table_of_map_SomeI)
  22.423 +apply (drule_tac x="old" in spec)
  22.424 +apply (auto dest: overrides_eq_sigD simp add: msig_def)
  22.425 +done
  22.426 +
  22.427 +lemma wf_cdecl_hides_SomeD:
  22.428 +"\<lbrakk>wf_cdecl G (C,c); C \<noteq> Object; table_of (methods c) sig = Some newM;
  22.429 +  G,sig\<turnstile>(C,newM) hides old
  22.430 +\<rbrakk> \<Longrightarrow>  accmodi old \<le> access newM \<and>
  22.431 +       is_static old" 
  22.432 +apply (drule (1) wf_cdecl_supD)
  22.433 +apply (clarify)
  22.434 +apply (drule entailsD)
  22.435 +apply   (blast intro: table_of_map_SomeI)
  22.436 +apply (drule_tac x="old" in spec)
  22.437 +apply (auto dest: hides_eq_sigD simp add: msig_def)
  22.438 +done
  22.439 +
  22.440 +lemma wf_cdecl_wt_init: 
  22.441 + "wf_cdecl G (C, c) \<Longrightarrow> \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>init c\<Colon>\<surd>"
  22.442 +apply (unfold wf_cdecl_def)
  22.443 +apply auto
  22.444 +done
  22.445 +
  22.446 +
  22.447 +section "well-formed programs"
  22.448 +  (* well-formed program, cf. 8.1, 9.1 *)
  22.449 +
  22.450 +text {*
  22.451 +A program declaration is wellformed if:
  22.452 +\begin{itemize}
  22.453 +\item the class ObjectC of Object is defined
  22.454 +\item every method of has an access modifier distinct from Package. This is
  22.455 +      necessary since every interface automatically inherits from Object.  
  22.456 +      We must know, that every time a Object method is "overriden" by an 
  22.457 +      interface method this is also overriden by the class implementing the
  22.458 +      the interface (see @{text "implement_dynmethd and class_mheadsD"})
  22.459 +\item all standard Exceptions are defined
  22.460 +\item all defined interfaces are wellformed
  22.461 +\item all defined classes are wellformed
  22.462 +\end{itemize}
  22.463 +*}
  22.464 +constdefs
  22.465 +  wf_prog  :: "prog \<Rightarrow> bool"
  22.466 + "wf_prog G \<equiv> let is = ifaces G; cs = classes G in
  22.467 +	         ObjectC \<in> set cs \<and> 
  22.468 +                (\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package) \<and>
  22.469 +                (\<forall>xn. SXcptC xn \<in> set cs) \<and>
  22.470 +		(\<forall>i\<in>set is. wf_idecl G i) \<and> unique is \<and>
  22.471 +		(\<forall>c\<in>set cs. wf_cdecl G c) \<and> unique cs"
  22.472 +
  22.473 +lemma wf_prog_idecl: "\<lbrakk>iface G I = Some i; wf_prog G\<rbrakk> \<Longrightarrow> wf_idecl G (I,i)"
  22.474 +apply (unfold wf_prog_def Let_def)
  22.475 +apply simp
  22.476 +apply (fast dest: map_of_SomeD)
  22.477 +done
  22.478 +
  22.479 +lemma wf_prog_cdecl: "\<lbrakk>class G C = Some c; wf_prog G\<rbrakk> \<Longrightarrow> wf_cdecl G (C,c)"
  22.480 +apply (unfold wf_prog_def Let_def)
  22.481 +apply simp
  22.482 +apply (fast dest: map_of_SomeD)
  22.483 +done
  22.484 +
  22.485 +lemma wf_prog_Object_mdecls:
  22.486 +"wf_prog G \<Longrightarrow> (\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package)"
  22.487 +apply (unfold wf_prog_def Let_def)
  22.488 +apply simp
  22.489 +done
  22.490 +
  22.491 +lemma wf_prog_acc_superD:
  22.492 + "\<lbrakk>wf_prog G; class G C = Some c; C \<noteq> Object \<rbrakk> 
  22.493 +  \<Longrightarrow> is_acc_class G (pid C) (super c)"
  22.494 +by (auto dest: wf_prog_cdecl wf_cdecl_supD)
  22.495 +
  22.496 +lemma wf_ws_prog [elim!,simp]: "wf_prog G \<Longrightarrow> ws_prog G"
  22.497 +apply (unfold wf_prog_def Let_def)
  22.498 +apply (rule ws_progI)
  22.499 +apply  (simp_all (no_asm))
  22.500 +apply  (auto simp add: is_acc_class_def is_acc_iface_def 
  22.501 +             dest!: wf_idecl_supD wf_cdecl_supD )+
  22.502 +done
  22.503 +
  22.504 +lemma class_Object [simp]: 
  22.505 +"wf_prog G \<Longrightarrow> 
  22.506 +  class G Object = Some \<lparr>access=Public,cfields=[],methods=Object_mdecls,
  22.507 +                                  init=Skip,super=arbitrary,superIfs=[]\<rparr>"
  22.508 +apply (unfold wf_prog_def Let_def ObjectC_def)
  22.509 +apply (fast dest!: map_of_SomeI)
  22.510 +done
  22.511 +
  22.512 +lemma methd_Object[simp]: "wf_prog G \<Longrightarrow> methd G Object =  
  22.513 +  table_of (map (\<lambda>(s,m). (s, Object, m)) Object_mdecls)"
  22.514 +apply (subst methd_rec)
  22.515 +apply (auto simp add: Let_def)
  22.516 +done
  22.517 +
  22.518 +lemma wf_prog_Object_methd:
  22.519 +"\<lbrakk>wf_prog G; methd G Object sig = Some m\<rbrakk> \<Longrightarrow> accmodi m \<noteq> Package"
  22.520 +by (auto dest!: wf_prog_Object_mdecls) (auto dest!: map_of_SomeD) 
  22.521 +
  22.522 +lemma wf_prog_Object_is_public[intro]:
  22.523 + "wf_prog G \<Longrightarrow> is_public G Object"
  22.524 +by (auto simp add: is_public_def dest: class_Object)
  22.525 +
  22.526 +lemma class_SXcpt [simp]: 
  22.527 +"wf_prog G \<Longrightarrow> 
  22.528 +  class G (SXcpt xn) = Some \<lparr>access=Public,cfields=[],methods=SXcpt_mdecls,
  22.529 +                                   init=Skip,
  22.530 +                                   super=if xn = Throwable then Object 
  22.531 +                                                           else SXcpt Throwable,
  22.532 +                                   superIfs=[]\<rparr>"
  22.533 +apply (unfold wf_prog_def Let_def SXcptC_def)
  22.534 +apply (fast dest!: map_of_SomeI)
  22.535 +done
  22.536 +
  22.537 +lemma wf_ObjectC [simp]: 
  22.538 +	"wf_cdecl G ObjectC = (\<not>is_iface G Object \<and> Ball (set Object_mdecls)
  22.539 +  (wf_mdecl G Object) \<and> unique Object_mdecls)"
  22.540 +apply (unfold wf_cdecl_def ws_cdecl_def ObjectC_def)
  22.541 +apply (simp (no_asm))
  22.542 +done
  22.543 +
  22.544 +lemma Object_is_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_class G Object"
  22.545 +apply (simp (no_asm_simp))
  22.546 +done
  22.547 + 
  22.548 +lemma Object_is_acc_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_acc_class G S Object"
  22.549 +apply (simp (no_asm_simp) add: is_acc_class_def is_public_def
  22.550 +                               accessible_in_RefT_simp)
  22.551 +done
  22.552 +
  22.553 +lemma SXcpt_is_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_class G (SXcpt xn)"
  22.554 +apply (simp (no_asm_simp))
  22.555 +done
  22.556 +
  22.557 +lemma SXcpt_is_acc_class [simp,elim!]: 
  22.558 +"wf_prog G \<Longrightarrow> is_acc_class G S (SXcpt xn)"
  22.559 +apply (simp (no_asm_simp) add: is_acc_class_def is_public_def
  22.560 +                               accessible_in_RefT_simp)
  22.561 +done
  22.562 +
  22.563 +lemma fields_Object [simp]: "wf_prog G \<Longrightarrow> DeclConcepts.fields G Object = []"
  22.564 +by (force intro: fields_emptyI)
  22.565 +
  22.566 +lemma accfield_Object [simp]: 
  22.567 + "wf_prog G \<Longrightarrow> accfield G S Object = empty"
  22.568 +apply (unfold accfield_def)
  22.569 +apply (simp (no_asm_simp) add: Let_def)
  22.570 +done
  22.571 +
  22.572 +lemma fields_Throwable [simp]: 
  22.573 + "wf_prog G \<Longrightarrow> DeclConcepts.fields G (SXcpt Throwable) = []"
  22.574 +by (force intro: fields_emptyI)
  22.575 +
  22.576 +lemma fields_SXcpt [simp]: "wf_prog G \<Longrightarrow> DeclConcepts.fields G (SXcpt xn) = []"
  22.577 +apply (case_tac "xn = Throwable")
  22.578 +apply  (simp (no_asm_simp))
  22.579 +by (force intro: fields_emptyI)
  22.580 +
  22.581 +lemmas widen_trans = ws_widen_trans [OF _ _ wf_ws_prog, elim]
  22.582 +lemma widen_trans2 [elim]: "\<lbrakk>G\<turnstile>U\<preceq>T; G\<turnstile>S\<preceq>U; wf_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>S\<preceq>T"
  22.583 +apply (erule (2) widen_trans)
  22.584 +done
  22.585 +
  22.586 +lemma Xcpt_subcls_Throwable [simp]: 
  22.587 +"wf_prog G \<Longrightarrow> G\<turnstile>SXcpt xn\<preceq>\<^sub>C SXcpt Throwable"
  22.588 +apply (rule SXcpt_subcls_Throwable_lemma)
  22.589 +apply auto
  22.590 +done
  22.591 +
  22.592 +lemma unique_fields: 
  22.593 + "\<lbrakk>is_class G C; wf_prog G\<rbrakk> \<Longrightarrow> unique (DeclConcepts.fields G C)"
  22.594 +apply (erule ws_unique_fields)
  22.595 +apply  (erule wf_ws_prog)
  22.596 +apply (erule (1) wf_prog_cdecl [THEN wf_cdecl_unique [THEN conjunct1]])
  22.597 +done
  22.598 +
  22.599 +lemma fields_mono: 
  22.600 +"\<lbrakk>table_of (DeclConcepts.fields G C) fn = Some f; G\<turnstile>D\<preceq>\<^sub>C C; 
  22.601 +  is_class G D; wf_prog G\<rbrakk> 
  22.602 +   \<Longrightarrow> table_of (DeclConcepts.fields G D) fn = Some f"
  22.603 +apply (rule map_of_SomeI)
  22.604 +apply  (erule (1) unique_fields)
  22.605 +apply (erule (1) map_of_SomeD [THEN fields_mono_lemma])
  22.606 +apply (erule wf_ws_prog)
  22.607 +done
  22.608 +
  22.609 +
  22.610 +lemma fields_is_type [elim]: 
  22.611 +"\<lbrakk>table_of (DeclConcepts.fields G C) m = Some f; wf_prog G; is_class G C\<rbrakk> \<Longrightarrow> 
  22.612 +      is_type G (type f)"
  22.613 +apply (frule wf_ws_prog)
  22.614 +apply (force dest: fields_declC [THEN conjunct1] 
  22.615 +                   wf_prog_cdecl [THEN wf_cdecl_fdecl]
  22.616 +             simp add: wf_fdecl_def2 is_acc_type_def)
  22.617 +done
  22.618 +
  22.619 +lemma imethds_wf_mhead [rule_format (no_asm)]: 
  22.620 +"\<lbrakk>m \<in> imethds G I sig; wf_prog G; is_iface G I\<rbrakk> \<Longrightarrow>  
  22.621 +  wf_mhead G (pid (decliface m)) sig (mthd m) \<and> 
  22.622 +  \<not> is_static m \<and> accmodi m = Public"
  22.623 +apply (frule wf_ws_prog)
  22.624 +apply (drule (2) imethds_declI [THEN conjunct1])
  22.625 +apply clarify
  22.626 +apply (frule_tac I="(decliface m)" in wf_prog_idecl,assumption)
  22.627 +apply (drule wf_idecl_mhead)
  22.628 +apply (erule map_of_SomeD)
  22.629 +apply (cases m, simp)
  22.630 +done
  22.631 +
  22.632 +lemma methd_wf_mdecl: 
  22.633 + "\<lbrakk>methd G C sig = Some m; wf_prog G; class G C = Some y\<rbrakk> \<Longrightarrow>  
  22.634 +  G\<turnstile>C\<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m) \<and> 
  22.635 +  wf_mdecl G (declclass m) (sig,(mthd m))"
  22.636 +apply (frule wf_ws_prog)
  22.637 +apply (drule (1) methd_declC)
  22.638 +apply  fast
  22.639 +apply clarsimp
  22.640 +apply (frule (1) wf_prog_cdecl, erule wf_cdecl_mdecl, erule map_of_SomeD)
  22.641 +done
  22.642 +
  22.643 +(*
  22.644 +This lemma doesn't hold!
  22.645 +lemma methd_rT_is_acc_type: 
  22.646 +"\<lbrakk>wf_prog G;methd G C C sig = Some (D,m);
  22.647 +    class G C = Some y\<rbrakk>
  22.648 +\<Longrightarrow> is_acc_type G (pid C) (resTy m)"
  22.649 +The result Type is only visible in the scope of defining class D 
  22.650 +"is_vis_type G (pid D) (resTy m)" but not necessarily in scope of class C!
  22.651 +(The same is true for the type of pramaters of a method)
  22.652 +*)
  22.653 +
  22.654 +
  22.655 +lemma methd_rT_is_type: 
  22.656 +"\<lbrakk>wf_prog G;methd G C sig = Some m;
  22.657 +    class G C = Some y\<rbrakk>
  22.658 +\<Longrightarrow> is_type G (resTy m)"
  22.659 +apply (drule (2) methd_wf_mdecl)
  22.660 +apply clarify
  22.661 +apply (drule wf_mdeclD1)
  22.662 +apply clarify
  22.663 +apply (drule rT_is_acc_type)
  22.664 +apply (cases m, simp add: is_acc_type_def)
  22.665 +done
  22.666 +
  22.667 +lemma accmethd_rT_is_type:
  22.668 +"\<lbrakk>wf_prog G;accmethd G S C sig = Some m;
  22.669 +    class G C = Some y\<rbrakk>
  22.670 +\<Longrightarrow> is_type G (resTy m)"
  22.671 +by (auto simp add: accmethd_def  
  22.672 +         intro: methd_rT_is_type)
  22.673 +
  22.674 +lemma methd_Object_SomeD:
  22.675 +"\<lbrakk>wf_prog G;methd G Object sig = Some m\<rbrakk> 
  22.676 + \<Longrightarrow> declclass m = Object"
  22.677 +by (auto dest: class_Object simp add: methd_rec )
  22.678 +
  22.679 +lemma wf_imethdsD: 
  22.680 + "\<lbrakk>im \<in> imethds G I sig;wf_prog G; is_iface G I\<rbrakk> 
  22.681 + \<Longrightarrow> \<not>is_static im \<and> accmodi im = Public"
  22.682 +proof -
  22.683 +  assume asm: "wf_prog G" "is_iface G I" "im \<in> imethds G I sig"
  22.684 +  have "wf_prog G \<longrightarrow> 
  22.685 +         (\<forall> i im. iface G I = Some i \<longrightarrow> im \<in> imethds G I sig
  22.686 +                  \<longrightarrow> \<not>is_static im \<and> accmodi im = Public)" (is "?P G I")
  22.687 +  proof (rule iface_rec.induct,intro allI impI)
  22.688 +    fix G I i im
  22.689 +    assume hyp: "\<forall> J i. J \<in> set (isuperIfs i) \<and> ws_prog G \<and> iface G I = Some i
  22.690 +                 \<longrightarrow> ?P G J"
  22.691 +    assume wf: "wf_prog G" and if_I: "iface G I = Some i" and 
  22.692 +           im: "im \<in> imethds G I sig" 
  22.693 +    show "\<not>is_static im \<and> accmodi im = Public" 
  22.694 +    proof -
  22.695 +      let ?inherited = "Un_tables (imethds G ` set (isuperIfs i))"
  22.696 +      let ?new = "(o2s \<circ> table_of (map (\<lambda>(s, mh). (s, I, mh)) (imethods i)))"
  22.697 +      from if_I wf im have imethds:"im \<in> (?inherited \<oplus>\<oplus> ?new) sig"
  22.698 +	by (simp add: imethds_rec)
  22.699 +      from wf if_I have 
  22.700 +	wf_supI: "\<forall> J. J \<in> set (isuperIfs i) \<longrightarrow> (\<exists> j. iface G J = Some j)"
  22.701 +	by (blast dest: wf_prog_idecl wf_idecl_supD is_acc_ifaceD)
  22.702 +      from wf if_I have
  22.703 +	"\<forall> im \<in> set (imethods i). \<not> is_static im \<and> accmodi im = Public"
  22.704 +	by (auto dest!: wf_prog_idecl wf_idecl_mhead)
  22.705 +      then have new_ok: "\<forall> im. table_of (imethods i) sig = Some im 
  22.706 +                         \<longrightarrow>  \<not> is_static im \<and> accmodi im = Public"
  22.707 +	by (auto dest!: table_of_Some_in_set)
  22.708 +      show ?thesis
  22.709 +	proof (cases "?new sig = {}")
  22.710 +	  case True
  22.711 +	  from True wf wf_supI if_I imethds hyp 
  22.712 +	  show ?thesis by (auto simp del:  split_paired_All)  
  22.713 +	next
  22.714 +	  case False
  22.715 +	  from False wf wf_supI if_I imethds new_ok hyp 
  22.716 +	  show ?thesis by (auto dest: wf_idecl_hidings hidings_entailsD)
  22.717 +	qed
  22.718 +      qed
  22.719 +    qed
  22.720 +  with asm show ?thesis by (auto simp del: split_paired_All)
  22.721 +qed
  22.722 +
  22.723 +lemma wf_prog_hidesD:
  22.724 + (assumes hides: "G \<turnstile>new hides old" and wf: "wf_prog G") 
  22.725 +   "accmodi old \<le> accmodi new \<and>
  22.726 +    is_static old"
  22.727 +proof -
  22.728 +  from hides 
  22.729 +  obtain c where 
  22.730 +    clsNew: "class G (declclass new) = Some c" and
  22.731 +    neqObj: "declclass new \<noteq> Object"
  22.732 +    by (auto dest: hidesD declared_in_classD)
  22.733 +  with hides obtain newM oldM where
  22.734 +    newM: "table_of (methods c) (msig new) = Some newM" and 
  22.735 +     new: "new = (declclass new,(msig new),newM)" and
  22.736 +     old: "old = (declclass old,(msig old),oldM)" and
  22.737 +          "msig new = msig old"
  22.738 +    by (cases new,cases old) 
  22.739 +       (auto dest: hidesD 
  22.740 +         simp add: cdeclaredmethd_def declared_in_def)
  22.741 +  with hides 
  22.742 +  have hides':
  22.743 +        "G,(msig new)\<turnstile>(declclass new,newM) hides (declclass old,oldM)"
  22.744 +    by auto
  22.745 +  from clsNew wf 
  22.746 +  have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl)
  22.747 +  note wf_cdecl_hides_SomeD [OF this neqObj newM hides']
  22.748 +  with new old 
  22.749 +  show ?thesis
  22.750 +    by (cases new, cases old) auto
  22.751 +qed
  22.752 +
  22.753 +text {* Compare this lemma about static  
  22.754 +overriding @{term "G \<turnstile>new overrides\<^sub>S old"} with the definition of 
  22.755 +dynamic overriding @{term "G \<turnstile>new overrides old"}. 
  22.756 +Conforming result types and restrictions on the access modifiers of the old 
  22.757 +and the new method are not part of the predicate for static overriding. But
  22.758 +they are enshured in a wellfromed program.  Dynamic overriding has 
  22.759 +no restrictions on the access modifiers but enforces confrom result types 
  22.760 +as precondition. But with some efford we can guarantee the access modifier
  22.761 +restriction for dynamic overriding, too. See lemma 
  22.762 +@{text wf_prog_dyn_override_prop}.
  22.763 +*}
  22.764 +lemma wf_prog_stat_overridesD:
  22.765 + (assumes stat_override: "G \<turnstile>new overrides\<^sub>S old" and wf: "wf_prog G") 
  22.766 +   "G\<turnstile>resTy new\<preceq>resTy old \<and>
  22.767 +    accmodi old \<le> accmodi new \<and>
  22.768 +    \<not> is_static old"
  22.769 +proof -
  22.770 +  from stat_override 
  22.771 +  obtain c where 
  22.772 +    clsNew: "class G (declclass new) = Some c" and
  22.773 +    neqObj: "declclass new \<noteq> Object"
  22.774 +    by (auto dest: stat_overrides_commonD declared_in_classD)
  22.775 +  with stat_override obtain newM oldM where
  22.776 +    newM: "table_of (methods c) (msig new) = Some newM" and 
  22.777 +     new: "new = (declclass new,(msig new),newM)" and
  22.778 +     old: "old = (declclass old,(msig old),oldM)" and
  22.779 +          "msig new = msig old"
  22.780 +    by (cases new,cases old) 
  22.781 +       (auto dest: stat_overrides_commonD 
  22.782 +         simp add: cdeclaredmethd_def declared_in_def)
  22.783 +  with stat_override 
  22.784 +  have stat_override':
  22.785 +        "G,(msig new)\<turnstile>(declclass new,newM) overrides\<^sub>S (declclass old,oldM)"
  22.786 +    by auto
  22.787 +  from clsNew wf 
  22.788 +  have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl)
  22.789 +  note wf_cdecl_overrides_SomeD [OF this neqObj newM stat_override']
  22.790 +  with new old 
  22.791 +  show ?thesis
  22.792 +    by (cases new, cases old) auto
  22.793 +qed
  22.794 +    
  22.795 +lemma static_to_dynamic_overriding: 
  22.796 +  (assumes stat_override: "G\<turnstile>new overrides\<^sub>S old" and wf : "wf_prog G")
  22.797 +   "G\<turnstile>new overrides old"
  22.798 +proof -
  22.799 +  from stat_override 
  22.800 +  show ?thesis (is "?Overrides new old")
  22.801 +  proof (induct)
  22.802 +    case (Direct new old superNew)
  22.803 +    then have stat_override:"G\<turnstile>new overrides\<^sub>S old" 
  22.804 +      by (rule stat_overridesR.Direct)
  22.805 +    from stat_override wf
  22.806 +    have resTy_widen: "G\<turnstile>resTy new\<preceq>resTy old" and
  22.807 +      not_static_old: "\<not> is_static old" 
  22.808 +      by (auto dest: wf_prog_stat_overridesD)  
  22.809 +    have not_private_new: "accmodi new \<noteq> Private"
  22.810 +    proof -
  22.811 +      from stat_override 
  22.812 +      have "accmodi old \<noteq> Private"
  22.813 +	by (rule no_Private_stat_override)
  22.814 +      moreover
  22.815 +      from stat_override wf
  22.816 +      have "accmodi old \<le> accmodi new"
  22.817 +	by (auto dest: wf_prog_stat_overridesD)
  22.818 +      ultimately
  22.819 +      show ?thesis
  22.820 +	by (auto dest: acc_modi_bottom)
  22.821 +    qed
  22.822 +    with Direct resTy_widen not_static_old 
  22.823 +    show "?Overrides new old" 
  22.824 +      by (auto intro: overridesR.Direct) 
  22.825 +  next
  22.826 +    case (Indirect inter new old)
  22.827 +    then show "?Overrides new old" 
  22.828 +      by (blast intro: overridesR.Indirect) 
  22.829 +  qed
  22.830 +qed
  22.831 +
  22.832 +lemma non_Package_instance_method_inheritance:
  22.833 +(assumes old_inheritable: "G\<turnstile>Method old inheritable_in (pid C)" and
  22.834 +             accmodi_old: "accmodi old \<noteq> Package" and 
  22.835 +         instance_method: "\<not> is_static old" and
  22.836 +                  subcls: "G\<turnstile>C \<prec>\<^sub>C declclass old" and
  22.837 +            old_declared: "G\<turnstile>Method old declared_in (declclass old)" and
  22.838 +                      wf: "wf_prog G"
  22.839 +) "G\<turnstile>Method old member_of C \<or>
  22.840 +   (\<exists> new. G\<turnstile> new overrides\<^sub>S old \<and>  G\<turnstile>Method new member_of C)"
  22.841 +proof -
  22.842 +  from wf have ws: "ws_prog G" by auto
  22.843 +  from old_declared have iscls_declC_old: "is_class G (declclass old)"
  22.844 +    by (auto simp add: declared_in_def cdeclaredmethd_def)
  22.845 +  from subcls have  iscls_C: "is_class G C"
  22.846 +    by (blast dest:  subcls_is_class)
  22.847 +  from iscls_C ws old_inheritable subcls 
  22.848 +  show ?thesis (is "?P C old")
  22.849 +  proof (induct rule: ws_class_induct')
  22.850 +    case Object
  22.851 +    assume "G\<turnstile>Object\<prec>\<^sub>C declclass old"
  22.852 +    then show "?P Object old"
  22.853 +      by blast
  22.854 +  next
  22.855 +    case (Subcls C c)
  22.856 +    assume cls_C: "class G C = Some c" and 
  22.857 +       neq_C_Obj: "C \<noteq> Object" and
  22.858 +             hyp: "\<lbrakk>G \<turnstile>Method old inheritable_in pid (super c); 
  22.859 +                   G\<turnstile>super c\<prec>\<^sub>C declclass old\<rbrakk> \<Longrightarrow> ?P (super c) old" and
  22.860 +     inheritable: "G \<turnstile>Method old inheritable_in pid C" and
  22.861 +         subclsC: "G\<turnstile>C\<prec>\<^sub>C declclass old"
  22.862 +    from cls_C neq_C_Obj  
  22.863 +    have super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c" 
  22.864 +      by (rule subcls1I)
  22.865 +    from wf cls_C neq_C_Obj
  22.866 +    have accessible_super: "G\<turnstile>(Class (super c)) accessible_in (pid C)" 
  22.867 +      by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
  22.868 +    {
  22.869 +      fix old
  22.870 +      assume    member_super: "G\<turnstile>Method old member_of (super c)"
  22.871 +      assume     inheritable: "G \<turnstile>Method old inheritable_in pid C"
  22.872 +      assume instance_method: "\<not> is_static old"
  22.873 +      from member_super
  22.874 +      have old_declared: "G\<turnstile>Method old declared_in (declclass old)"
  22.875 +       by (cases old) (auto dest: member_of_declC)
  22.876 +      have "?P C old"
  22.877 +      proof (cases "G\<turnstile>mid (msig old) undeclared_in C")
  22.878 +	case True
  22.879 +	with inheritable super accessible_super member_super
  22.880 +	have "G\<turnstile>Method old member_of C"
  22.881 +	  by (cases old) (auto intro: members.Inherited)
  22.882 +	then show ?thesis
  22.883 +	  by auto
  22.884 +      next
  22.885 +	case False
  22.886 +	then obtain new_member where
  22.887 +	     "G\<turnstile>new_member declared_in C" and
  22.888 +             "mid (msig old) = memberid new_member"
  22.889 +          by (auto dest: not_undeclared_declared)
  22.890 +	then obtain new where
  22.891 +	          new: "G\<turnstile>Method new declared_in C" and
  22.892 +               eq_sig: "msig old = msig new" and
  22.893 +	    declC_new: "declclass new = C" 
  22.894 +	  by (cases new_member) auto
  22.895 +	then have member_new: "G\<turnstile>Method new member_of C"
  22.896 +	  by (cases new) (auto intro: members.Immediate)
  22.897 +	from declC_new super member_super
  22.898 +	have subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old"
  22.899 +	  by (auto dest!: member_of_subclseq_declC
  22.900 +	            dest: r_into_trancl intro: trancl_rtrancl_trancl)
  22.901 +	show ?thesis
  22.902 +	proof (cases "is_static new")
  22.903 +	  case False
  22.904 +	  with eq_sig declC_new new old_declared inheritable
  22.905 +	       super member_super subcls_new_old
  22.906 +	  have "G\<turnstile>new overrides\<^sub>S old"
  22.907 +	    by (auto intro!: stat_overridesR.Direct)
  22.908 +	  with member_new show ?thesis
  22.909 +	    by blast
  22.910 +	next
  22.911 +	  case True
  22.912 +	  with eq_sig declC_new subcls_new_old new old_declared inheritable
  22.913 +	  have "G\<turnstile>new hides old"
  22.914 +	    by (auto intro: hidesI)    
  22.915 +	  with wf 
  22.916 +	  have "is_static old"
  22.917 +	    by (blast dest: wf_prog_hidesD)
  22.918 +	  with instance_method
  22.919 +	  show ?thesis
  22.920 +	    by (contradiction)
  22.921 +	qed
  22.922 +      qed
  22.923 +    } note hyp_member_super = this
  22.924 +    from subclsC cls_C 
  22.925 +    have "G\<turnstile>(super c)\<preceq>\<^sub>C declclass old"
  22.926 +      by (rule subcls_superD)
  22.927 +    then
  22.928 +    show "?P C old"
  22.929 +    proof (cases rule: subclseq_cases) 
  22.930 +      case Eq
  22.931 +      assume "super c = declclass old"
  22.932 +      with old_declared 
  22.933 +      have "G\<turnstile>Method old member_of (super c)" 
  22.934 +	by (cases old) (auto intro: members.Immediate)
  22.935 +      with inheritable instance_method 
  22.936 +      show ?thesis
  22.937 +	by (blast dest: hyp_member_super)
  22.938 +    next
  22.939 +      case Subcls
  22.940 +      assume "G\<turnstile>super c\<prec>\<^sub>C declclass old"
  22.941 +      moreover
  22.942 +      from inheritable accmodi_old
  22.943 +      have "G \<turnstile>Method old inheritable_in pid (super c)"
  22.944 +	by (cases "accmodi old") (auto simp add: inheritable_in_def)
  22.945 +      ultimately
  22.946 +      have "?P (super c) old"
  22.947 +	by (blast dest: hyp)
  22.948 +      then show ?thesis
  22.949 +      proof
  22.950 +	assume "G \<turnstile>Method old member_of super c"
  22.951 +	with inheritable instance_method
  22.952 +	show ?thesis
  22.953 +	  by (blast dest: hyp_member_super)
  22.954 +      next
  22.955 +	assume "\<exists>new. G \<turnstile> new overrides\<^sub>S old \<and> G \<turnstile>Method new member_of super c"
  22.956 +	then obtain super_new where
  22.957 +	  super_new_override:  "G \<turnstile> super_new overrides\<^sub>S old" and
  22.958 +            super_new_member:  "G \<turnstile>Method super_new member_of super c"
  22.959 +	  by blast
  22.960 +	from super_new_override wf
  22.961 +	have "accmodi old \<le> accmodi super_new"
  22.962 +	  by (auto dest: wf_prog_stat_overridesD)
  22.963 +	with inheritable accmodi_old
  22.964 +	have "G \<turnstile>Method super_new inheritable_in pid C"
  22.965 +	  by (auto simp add: inheritable_in_def 
  22.966 +	              split: acc_modi.splits
  22.967 +                       dest: acc_modi_le_Dests)
  22.968 +	moreover
  22.969 +	from super_new_override 
  22.970 +	have "\<not> is_static super_new"
  22.971 +	  by (auto dest: stat_overrides_commonD)
  22.972 +	moreover
  22.973 +	note super_new_member
  22.974 +	ultimately have "?P C super_new"
  22.975 +	  by (auto dest: hyp_member_super)
  22.976 +	then show ?thesis
  22.977 +	proof 
  22.978 +	  assume "G \<turnstile>Method super_new member_of C"
  22.979 +	  with super_new_override
  22.980 +	  show ?thesis
  22.981 +	    by blast
  22.982 +	next
  22.983 +	  assume "\<exists>new. G \<turnstile> new overrides\<^sub>S super_new \<and> 
  22.984 +                  G \<turnstile>Method new member_of C"
  22.985 +	  with super_new_override show ?thesis
  22.986 +	    by (blast intro: stat_overridesR.Indirect) 
  22.987 +	qed
  22.988 +      qed
  22.989 +    qed
  22.990 +  qed
  22.991 +qed
  22.992 +
  22.993 +lemma non_Package_instance_method_inheritance_cases [consumes 6,
  22.994 +         case_names Inheritance Overriding]:
  22.995 +(assumes old_inheritable: "G\<turnstile>Method old inheritable_in (pid C)" and
  22.996 +             accmodi_old: "accmodi old \<noteq> Package" and 
  22.997 +         instance_method: "\<not> is_static old" and
  22.998 +                  subcls: "G\<turnstile>C \<prec>\<^sub>C declclass old" and
  22.999 +            old_declared: "G\<turnstile>Method old declared_in (declclass old)" and
 22.1000 +                      wf: "wf_prog G" and
 22.1001 +             inheritance: "G\<turnstile>Method old member_of C \<Longrightarrow> P" and
 22.1002 +             overriding:  "\<And> new. 
 22.1003 +                           \<lbrakk>G\<turnstile> new overrides\<^sub>S old;G\<turnstile>Method new member_of C\<rbrakk>
 22.1004 +                           \<Longrightarrow> P"
 22.1005 +) "P"
 22.1006 +proof -
 22.1007 +  from old_inheritable accmodi_old instance_method subcls old_declared wf 
 22.1008 +       inheritance overriding
 22.1009 +  show ?thesis
 22.1010 +    by (auto dest: non_Package_instance_method_inheritance)
 22.1011 +qed
 22.1012 +
 22.1013 +lemma dynamic_to_static_overriding:
 22.1014 +(assumes dyn_override: "G\<turnstile> new overrides old" and
 22.1015 +          accmodi_old: "accmodi old \<noteq> Package" and
 22.1016 +                   wf: "wf_prog G"
 22.1017 +) "G\<turnstile> new overrides\<^sub>S old"  
 22.1018 +proof - 
 22.1019 +  from dyn_override accmodi_old
 22.1020 +  show ?thesis (is "?Overrides new old")
 22.1021 +  proof (induct rule: overridesR.induct)
 22.1022 +    case (Direct new old)
 22.1023 +    assume   new_declared: "G\<turnstile>Method new declared_in declclass new"
 22.1024 +    assume eq_sig_new_old: "msig new = msig old"
 22.1025 +    assume subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old"
 22.1026 +    assume "G \<turnstile>Method old inheritable_in pid (declclass new)" and
 22.1027 +           "accmodi old \<noteq> Package" and
 22.1028 +           "\<not> is_static old" and
 22.1029 +           "G\<turnstile>declclass new\<prec>\<^sub>C declclass old" and
 22.1030 +           "G\<turnstile>Method old declared_in declclass old" 
 22.1031 +    from this wf
 22.1032 +    show "?Overrides new old"
 22.1033 +    proof (cases rule: non_Package_instance_method_inheritance_cases)
 22.1034 +      case Inheritance
 22.1035 +      assume "G \<turnstile>Method old member_of declclass new"
 22.1036 +      then have "G\<turnstile>mid (msig old) undeclared_in declclass new"
 22.1037 +      proof cases
 22.1038 +	case Immediate 
 22.1039 +	with subcls_new_old wf show ?thesis 	
 22.1040 +	  by (auto dest: subcls_irrefl)
 22.1041 +      next
 22.1042 +	case Inherited
 22.1043 +	then show ?thesis
 22.1044 +	  by (cases old) auto
 22.1045 +      qed
 22.1046 +      with eq_sig_new_old new_declared
 22.1047 +      show ?thesis
 22.1048 +	by (cases old,cases new) (auto dest!: declared_not_undeclared)
 22.1049 +    next
 22.1050 +      case (Overriding new') 
 22.1051 +      assume stat_override_new': "G \<turnstile> new' overrides\<^sub>S old"
 22.1052 +      then have "msig new' = msig old"
 22.1053 +	by (auto dest: stat_overrides_commonD)
 22.1054 +      with eq_sig_new_old have eq_sig_new_new': "msig new=msig new'"
 22.1055 +	by simp
 22.1056 +      assume "G \<turnstile>Method new' member_of declclass new"
 22.1057 +      then show ?thesis
 22.1058 +      proof (cases)
 22.1059 +	case Immediate
 22.1060 +	then have declC_new: "declclass new' = declclass new" 
 22.1061 +	  by auto
 22.1062 +	from Immediate 
 22.1063 +	have "G\<turnstile>Method new' declared_in declclass new"
 22.1064 +	  by (cases new') auto
 22.1065 +	with new_declared eq_sig_new_new' declC_new 
 22.1066 +	have "new=new'"
 22.1067 +	  by (cases new, cases new') (auto dest: unique_declared_in) 
 22.1068 +	with stat_override_new'
 22.1069 +	show ?thesis
 22.1070 +	  by simp
 22.1071 +      next
 22.1072 +	case Inherited
 22.1073 +	then have "G\<turnstile>mid (msig new') undeclared_in declclass new"
 22.1074 +	  by (cases new') (auto)
 22.1075 +	with eq_sig_new_new' new_declared
 22.1076 +	show ?thesis
 22.1077 +	  by (cases new,cases new') (auto dest!: declared_not_undeclared)
 22.1078 +      qed
 22.1079 +    qed
 22.1080 +  next
 22.1081 +    case (Indirect inter new old)
 22.1082 +    assume accmodi_old: "accmodi old \<noteq> Package"
 22.1083 +    assume "accmodi old \<noteq> Package \<Longrightarrow> G \<turnstile> inter overrides\<^sub>S old"
 22.1084 +    with accmodi_old 
 22.1085 +    have stat_override_inter_old: "G \<turnstile> inter overrides\<^sub>S old"
 22.1086 +      by blast
 22.1087 +    moreover 
 22.1088 +    assume hyp_inter: "accmodi inter \<noteq> Package \<Longrightarrow> G \<turnstile> new overrides\<^sub>S inter"
 22.1089 +    moreover
 22.1090 +    have "accmodi inter \<noteq> Package"
 22.1091 +    proof -
 22.1092 +      from stat_override_inter_old wf 
 22.1093 +      have "accmodi old \<le> accmodi inter"
 22.1094 +	by (auto dest: wf_prog_stat_overridesD)
 22.1095 +      with stat_override_inter_old accmodi_old
 22.1096 +      show ?thesis
 22.1097 +	by (auto dest!: no_Private_stat_override
 22.1098 +                 split: acc_modi.splits 
 22.1099 +	         dest: acc_modi_le_Dests)
 22.1100 +    qed
 22.1101 +    ultimately show "?Overrides new old"
 22.1102 +      by (blast intro: stat_overridesR.Indirect)
 22.1103 +  qed
 22.1104 +qed
 22.1105 +
 22.1106 +lemma wf_prog_dyn_override_prop:
 22.1107 + (assumes dyn_override: "G \<turnstile> new overrides old" and
 22.1108 +                    wf: "wf_prog G"
 22.1109 + ) "accmodi old \<le> accmodi new"
 22.1110 +proof (cases "accmodi old = Package")
 22.1111 +  case True
 22.1112 +  note old_Package = this
 22.1113 +  show ?thesis
 22.1114 +  proof (cases "accmodi old \<le> accmodi new")
 22.1115 +    case True then show ?thesis .
 22.1116 +  next
 22.1117 +    case False
 22.1118 +    with old_Package 
 22.1119 +    have "accmodi new = Private"
 22.1120 +      by (cases "accmodi new") (auto simp add: le_acc_def less_acc_def)
 22.1121 +    with dyn_override 
 22.1122 +    show ?thesis
 22.1123 +      by (auto dest: overrides_commonD)
 22.1124 +  qed    
 22.1125 +next
 22.1126 +  case False
 22.1127 +  with dyn_override wf
 22.1128 +  have "G \<turnstile> new overrides\<^sub>S old"
 22.1129 +    by (blast intro: dynamic_to_static_overriding)
 22.1130 +  with wf 
 22.1131 +  show ?thesis
 22.1132 +   by (blast dest: wf_prog_stat_overridesD)
 22.1133 +qed 
 22.1134 +
 22.1135 +lemma overrides_Package_old: 
 22.1136 +(assumes dyn_override: "G \<turnstile> new overrides old" and 
 22.1137 +          accmodi_new: "accmodi new = Package" and
 22.1138 +                   wf: "wf_prog G "
 22.1139 +) "accmodi old = Package"
 22.1140 +proof (cases "accmodi old")
 22.1141 +  case Private
 22.1142 +  with dyn_override show ?thesis
 22.1143 +    by (simp add: no_Private_override)
 22.1144 +next
 22.1145 +  case Package
 22.1146 +  then show ?thesis .
 22.1147 +next
 22.1148 +  case Protected
 22.1149 +  with dyn_override wf
 22.1150 +  have "G \<turnstile> new overrides\<^sub>S old"
 22.1151 +    by (auto intro: dynamic_to_static_overriding)
 22.1152 +  with wf 
 22.1153 +  have "accmodi old \<le> accmodi new"
 22.1154 +    by (auto dest: wf_prog_stat_overridesD)
 22.1155 +  with Protected accmodi_new
 22.1156 +  show ?thesis
 22.1157 +    by (simp add: less_acc_def le_acc_def)
 22.1158 +next
 22.1159 +  case Public
 22.1160 +  with dyn_override wf
 22.1161 +  have "G \<turnstile> new overrides\<^sub>S old"
 22.1162 +    by (auto intro: dynamic_to_static_overriding)
 22.1163 +  with wf 
 22.1164 +  have "accmodi old \<le> accmodi new"
 22.1165 +    by (auto dest: wf_prog_stat_overridesD)
 22.1166 +  with Public accmodi_new
 22.1167 +  show ?thesis
 22.1168 +    by (simp add: less_acc_def le_acc_def)
 22.1169 +qed
 22.1170 +
 22.1171 +lemma dyn_override_Package:
 22.1172 + (assumes dyn_override: "G \<turnstile> new overrides old" and
 22.1173 +          accmodi_old: "accmodi old = Package" and 
 22.1174 +          accmodi_new: "accmodi new = Package" and
 22.1175 +                    wf: "wf_prog G"
 22.1176 + ) "pid (declclass old) = pid (declclass new)"
 22.1177 +proof - 
 22.1178 +  from dyn_override accmodi_old accmodi_new
 22.1179 +  show ?thesis (is "?EqPid old new")
 22.1180 +  proof (induct rule: overridesR.induct)
 22.1181 +    case (Direct new old)
 22.1182 +    assume "accmodi old = Package"
 22.1183 +           "G \<turnstile>Method old inheritable_in pid (declclass new)"
 22.1184 +    then show "pid (declclass old) =  pid (declclass new)"
 22.1185 +      by (auto simp add: inheritable_in_def)
 22.1186 +  next
 22.1187 +    case (Indirect inter new old)
 22.1188 +    assume accmodi_old: "accmodi old = Package" and
 22.1189 +           accmodi_new: "accmodi new = Package" 
 22.1190 +    assume "G \<turnstile> new overrides inter"
 22.1191 +    with accmodi_new wf
 22.1192 +    have "accmodi inter = Package"
 22.1193 +      by  (auto intro: overrides_Package_old)
 22.1194 +    with Indirect
 22.1195 +    show "pid (declclass old) =  pid (declclass new)"
 22.1196 +      by auto
 22.1197 +  qed
 22.1198 +qed
 22.1199 +
 22.1200 +lemma dyn_override_Package_escape:
 22.1201 + (assumes dyn_override: "G \<turnstile> new overrides old" and
 22.1202 +          accmodi_old: "accmodi old = Package" and 
 22.1203 +         outside_pack: "pid (declclass old) \<noteq> pid (declclass new)" and
 22.1204 +                    wf: "wf_prog G"
 22.1205 + ) "\<exists> inter. G \<turnstile> new overrides inter \<and> G \<turnstile> inter overrides old \<and>
 22.1206 +             pid (declclass old) = pid (declclass inter) \<and>
 22.1207 +             Protected \<le> accmodi inter"
 22.1208 +proof -
 22.1209 +  from dyn_override accmodi_old outside_pack
 22.1210 +  show ?thesis (is "?P new old")
 22.1211 +  proof (induct rule: overridesR.induct)
 22.1212 +    case (Direct new old)
 22.1213 +    assume accmodi_old: "accmodi old = Package"
 22.1214 +    assume outside_pack: "pid (declclass old) \<noteq> pid (declclass new)"
 22.1215 +    assume "G \<turnstile>Method old inheritable_in pid (declclass new)"
 22.1216 +    with accmodi_old 
 22.1217 +    have "pid (declclass old) = pid (declclass new)"
 22.1218 +      by (simp add: inheritable_in_def)
 22.1219 +    with outside_pack 
 22.1220 +    show "?P new old"
 22.1221 +      by (contradiction)
 22.1222 +  next
 22.1223 +    case (Indirect inter new old)
 22.1224 +    assume accmodi_old: "accmodi old = Package"
 22.1225 +    assume outside_pack: "pid (declclass old) \<noteq> pid (declclass new)"
 22.1226 +    assume override_new_inter: "G \<turnstile> new overrides inter"
 22.1227 +    assume override_inter_old: "G \<turnstile> inter overrides old"
 22.1228 +    assume hyp_new_inter: "\<lbrakk>accmodi inter = Package; 
 22.1229 +                           pid (declclass inter) \<noteq> pid (declclass new)\<rbrakk>
 22.1230 +                           \<Longrightarrow> ?P new inter"
 22.1231 +    assume hyp_inter_old: "\<lbrakk>accmodi old = Package; 
 22.1232 +                           pid (declclass old) \<noteq> pid (declclass inter)\<rbrakk>
 22.1233 +                           \<Longrightarrow> ?P inter old"
 22.1234 +    show "?P new old"
 22.1235 +    proof (cases "pid (declclass old) = pid (declclass inter)")
 22.1236 +      case True
 22.1237 +      note same_pack_old_inter = this
 22.1238 +      show ?thesis
 22.1239 +      proof (cases "pid (declclass inter) = pid (declclass new)")
 22.1240 +	case True
 22.1241 +	with same_pack_old_inter outside_pack
 22.1242 +	show ?thesis
 22.1243 +	  by auto
 22.1244 +      next
 22.1245 +	case False
 22.1246 +	note diff_pack_inter_new = this
 22.1247 +	show ?thesis
 22.1248 +	proof (cases "accmodi inter = Package")
 22.1249 +	  case True
 22.1250 +	  with diff_pack_inter_new hyp_new_inter  
 22.1251 +	  obtain newinter where
 22.1252 +	    over_new_newinter: "G \<turnstile> new overrides newinter" and
 22.1253 +            over_newinter_inter: "G \<turnstile> newinter overrides inter" and 
 22.1254 +            eq_pid: "pid (declclass inter) = pid (declclass newinter)" and
 22.1255 +            accmodi_newinter: "Protected \<le> accmodi newinter"
 22.1256 +	    by auto
 22.1257 +	  from over_newinter_inter override_inter_old
 22.1258 +	  have "G\<turnstile>newinter overrides old"
 22.1259 +	    by (rule overridesR.Indirect)
 22.1260 +	  moreover
 22.1261 +	  from eq_pid same_pack_old_inter 
 22.1262 +	  have "pid (declclass old) = pid (declclass newinter)"
 22.1263 +	    by simp
 22.1264 +	  moreover
 22.1265 +	  note over_new_newinter accmodi_newinter
 22.1266 +	  ultimately show ?thesis
 22.1267 +	    by blast
 22.1268 +	next
 22.1269 +	  case False
 22.1270 +	  with override_new_inter
 22.1271 +	  have "Protected \<le> accmodi inter"
 22.1272 +	    by (cases "accmodi inter") (auto dest: no_Private_override)
 22.1273 +	  with override_new_inter override_inter_old same_pack_old_inter
 22.1274 +	  show ?thesis
 22.1275 +	    by blast
 22.1276 +	qed
 22.1277 +      qed
 22.1278 +    next
 22.1279 +      case False
 22.1280 +      with accmodi_old hyp_inter_old
 22.1281 +      obtain newinter where
 22.1282 +	over_inter_newinter: "G \<turnstile> inter overrides newinter" and
 22.1283 +          over_newinter_old: "G \<turnstile> newinter overrides old" and 
 22.1284 +                eq_pid: "pid (declclass old) = pid (declclass newinter)" and
 22.1285 +	accmodi_newinter: "Protected \<le> accmodi newinter"
 22.1286 +	by auto
 22.1287 +      from override_new_inter over_inter_newinter 
 22.1288 +      have "G \<turnstile> new overrides newinter"
 22.1289 +	by (rule overridesR.Indirect)
 22.1290 +      with eq_pid over_newinter_old accmodi_newinter
 22.1291 +      show ?thesis
 22.1292 +	by blast
 22.1293 +    qed
 22.1294 +  qed
 22.1295 +qed
 22.1296 +
 22.1297 +declare split_paired_All [simp del] split_paired_Ex [simp del]
 22.1298 +ML_setup {*
 22.1299 +simpset_ref() := simpset() delloop "split_all_tac";
 22.1300 +claset_ref () := claset () delSWrapper "split_all_tac"
 22.1301 +*}
 22.1302 +
 22.1303 +lemma declclass_widen[rule_format]: 
 22.1304 + "wf_prog G 
 22.1305 + \<longrightarrow> (\<forall>c m. class G C = Some c \<longrightarrow> methd G C sig = Some m 
 22.1306 + \<longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m)" (is "?P G C")
 22.1307 +proof (rule class_rec.induct,intro allI impI)
 22.1308 +  fix G C c m
 22.1309 +  assume Hyp: "\<forall>c. C \<noteq> Object \<and> ws_prog G \<and> class G C = Some c 
 22.1310 +               \<longrightarrow> ?P G (super c)"
 22.1311 +  assume wf: "wf_prog G" and cls_C: "class G C = Some c" and
 22.1312 +         m:  "methd G C sig = Some m"
 22.1313 +  show "G\<turnstile>C\<preceq>\<^sub>C declclass m" 
 22.1314 +  proof (cases "C=Object")
 22.1315 +    case True 
 22.1316 +    with wf m show ?thesis by (simp add: methd_Object_SomeD)
 22.1317 +  next
 22.1318 +    let ?filter="filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)"
 22.1319 +    let ?table = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
 22.1320 +    case False 
 22.1321 +    with cls_C wf m
 22.1322 +    have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m "
 22.1323 +      by (simp add: methd_rec)
 22.1324 +    show ?thesis
 22.1325 +    proof (cases "?table sig")
 22.1326 +      case None
 22.1327 +      from this methd_C have "?filter (methd G (super c)) sig = Some m"
 22.1328 +	by simp
 22.1329 +      moreover
 22.1330 +      from wf cls_C False obtain sup where "class G (super c) = Some sup"
 22.1331 +	by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
 22.1332 +      moreover note wf False cls_C Hyp 
 22.1333 +      ultimately have "G\<turnstile>super c \<preceq>\<^sub>C declclass m"  by auto
 22.1334 +      moreover from cls_C False have  "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c" by (rule subcls1I)
 22.1335 +      ultimately show ?thesis by - (rule rtrancl_into_rtrancl2)
 22.1336 +    next
 22.1337 +      case Some
 22.1338 +      from this wf False cls_C methd_C show ?thesis by auto
 22.1339 +    qed
 22.1340 +  qed
 22.1341 +qed
 22.1342 +
 22.1343 +(*
 22.1344 +lemma declclass_widen[rule_format]: 
 22.1345 + "wf_prog G 
 22.1346 + \<longrightarrow> (\<forall>c m. class G C = Some c \<longrightarrow> methd G C sig = Some m 
 22.1347 + \<longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m)" (is "?P G C")
 22.1348 +apply (rule class_rec.induct)
 22.1349 +apply (rule impI)+
 22.1350 +apply (case_tac "C=Object")
 22.1351 +apply   (force simp add: methd_Object_SomeD)
 22.1352 +
 22.1353 +apply   (rule allI)+
 22.1354 +apply   (rule impI)
 22.1355 +apply   (simp (no_asm_simp) add: methd_rec)
 22.1356 +apply   (case_tac "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig")
 22.1357 +apply    (simp add: override_def)
 22.1358 +apply    (frule (1) subcls1I)
 22.1359 +apply    (drule (1) wf_prog_cdecl)
 22.1360 +apply    (drule (1) wf_cdecl_supD)
 22.1361 +apply    clarify
 22.1362 +apply    (drule is_acc_class_is_class)
 22.1363 +apply    clarify
 22.1364 +apply    (blast dest: rtrancl_into_rtrancl2)
 22.1365 +
 22.1366 +apply    auto
 22.1367 +done
 22.1368 +*)
 22.1369 +
 22.1370 +(*
 22.1371 +lemma accessible_public_inheritance_lemma1:
 22.1372 +  "\<lbrakk>wf_prog G; class G C = Some c; C \<noteq> Object; accmodi m = Public;
 22.1373 +    G\<turnstile>m accessible_through_inheritance_in (super c)\<rbrakk> 
 22.1374 +   \<Longrightarrow> G\<turnstile>m accessible_through_inheritance_in C"
 22.1375 +apply   (frule (1) subcls1I)
 22.1376 +apply   (rule accessible_through_inheritance.Indirect)
 22.1377 +apply     (blast)
 22.1378 +apply     (erule accessible_through_inheritance_subclsD)
 22.1379 +apply     (blast dest: wf_prog_acc_superD is_acc_classD)
 22.1380 +apply     assumption
 22.1381 +apply     (force dest: wf_prog_acc_superD is_acc_classD
 22.1382 +                 simp add: accessible_for_inheritance_in_def)
 22.1383 +done
 22.1384 +
 22.1385 +lemma accessible_public_inheritance_lemma[rule_format]:
 22.1386 +  "\<lbrakk>wf_prog G;C \<noteq> Object; class G C = Some c; 
 22.1387 +    accmodi m = Public
 22.1388 +   \<rbrakk> \<Longrightarrow> methd G (super c) sig = Some m 
 22.1389 +        \<longrightarrow> G\<turnstile>m accessible_through_inheritance_in C" 
 22.1390 +apply (frule (2) wf_prog_acc_superD [THEN is_acc_classD])
 22.1391 +apply (erule conjE)
 22.1392 +apply (simp only: not_None_eq)
 22.1393 +apply (erule exE)
 22.1394 +apply (case_tac "(super c) = Object")
 22.1395 +apply   (rule impI)
 22.1396 +apply   (rule accessible_through_inheritance.Direct)
 22.1397 +apply     force
 22.1398 +apply     (force simp add: accessible_for_inheritance_in_def)
 22.1399 +
 22.1400 +apply   (frule wf_ws_prog) 
 22.1401 +apply   (simp add: methd_rec)
 22.1402 +apply   (case_tac "table_of (map (\<lambda>(s, m). (s, super c, m)) (methods y)) sig")
 22.1403 +apply     simp
 22.1404 +apply     (clarify)
 22.1405 +apply     (rule_tac D="super c" in accessible_through_inheritance.Indirect)
 22.1406 +apply       (blast dest: subcls1I)
 22.1407 +apply       (blast)
 22.1408 +apply       simp
 22.1409 +apply       assumption
 22.1410 +apply       (simp add: accessible_for_inheritance_in_def)
 22.1411 +
 22.1412 +apply     clarsimp
 22.1413 +apply     (rule accessible_through_inheritance.Direct)
 22.1414 +apply     (auto dest: subcls1I simp add: accessible_for_inheritance_in_def)
 22.1415 +done
 22.1416 +
 22.1417 +lemma accessible_public_inheritance:
 22.1418 +  "\<lbrakk>wf_prog G; class G D = Some d; G\<turnstile>C \<prec>\<^sub>C D; methd G D sig = Some m; 
 22.1419 +    accmodi m = Public\<rbrakk> 
 22.1420 +   \<Longrightarrow> G \<turnstile> m accessible_through_inheritance_in C"
 22.1421 +apply (erule converse_trancl_induct)
 22.1422 +apply  (blast dest: subcls1D intro: accessible_public_inheritance_lemma)
 22.1423 +
 22.1424 +apply  (frule subcls1D)
 22.1425 +apply  clarify
 22.1426 +apply  (frule  (2) wf_prog_acc_superD [THEN is_acc_classD])
 22.1427 +apply  clarify
 22.1428 +apply  (rule_tac D="super c" in accessible_through_inheritance.Indirect)
 22.1429 +apply   (auto intro:trancl_into_trancl2 
 22.1430 +                    accessible_through_inheritance_subclsD
 22.1431 +              simp add: accessible_for_inheritance_in_def)
 22.1432 +done
 22.1433 +*)
 22.1434 +
 22.1435 +
 22.1436 +lemma declclass_methd_Object: 
 22.1437 + "\<lbrakk>wf_prog G; methd G Object sig = Some m\<rbrakk> \<Longrightarrow> declclass m = Object"
 22.1438 +by auto
 22.1439 +
 22.1440 +
 22.1441 +lemma methd_declaredD: 
 22.1442 + "\<lbrakk>wf_prog G; is_class G C;methd G C sig = Some m\<rbrakk> 
 22.1443 +  \<Longrightarrow> G\<turnstile>(mdecl (sig,mthd m)) declared_in (declclass m)"
 22.1444 +proof -
 22.1445 +  assume    wf: "wf_prog G"
 22.1446 +  then have ws: "ws_prog G" ..
 22.1447 +  assume  clsC: "is_class G C"
 22.1448 +  from clsC ws 
 22.1449 +  show "methd G C sig = Some m 
 22.1450 +        \<Longrightarrow> G\<turnstile>(mdecl (sig,mthd m)) declared_in (declclass m)"
 22.1451 +    (is "PROP ?P C") 
 22.1452 +  proof (induct ?P C rule: ws_class_induct')
 22.1453 +    case Object
 22.1454 +    assume "methd G Object sig = Some m" 
 22.1455 +    with wf show ?thesis
 22.1456 +      by - (rule method_declared_inI, auto) 
 22.1457 +  next
 22.1458 +    case Subcls
 22.1459 +    fix C c
 22.1460 +    assume clsC: "class G C = Some c"
 22.1461 +    and       m: "methd G C sig = Some m"
 22.1462 +    and     hyp: "methd G (super c) sig = Some m \<Longrightarrow> ?thesis" 
 22.1463 +    let ?newMethods = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
 22.1464 +    show ?thesis
 22.1465 +    proof (cases "?newMethods sig")
 22.1466 +      case None
 22.1467 +      from None ws clsC m hyp 
 22.1468 +      show ?thesis by (auto intro: method_declared_inI simp add: methd_rec)
 22.1469 +    next
 22.1470 +      case Some
 22.1471 +      from Some ws clsC m 
 22.1472 +      show ?thesis by (auto intro: method_declared_inI simp add: methd_rec) 
 22.1473 +    qed
 22.1474 +  qed
 22.1475 +qed
 22.1476 +
 22.1477 +
 22.1478 +lemma methd_rec_Some_cases [consumes 4, case_names NewMethod InheritedMethod]:
 22.1479 +(assumes methd_C: "methd G C sig = Some m" and
 22.1480 +                    ws: "ws_prog G" and
 22.1481 +                  clsC: "class G C = Some c" and
 22.1482 +             neq_C_Obj: "C\<noteq>Object" )  
 22.1483 +"\<lbrakk>table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = Some m \<Longrightarrow> P;
 22.1484 +  \<lbrakk>G\<turnstile>C inherits (method sig m); methd G (super c) sig = Some m\<rbrakk> \<Longrightarrow> P 
 22.1485 + \<rbrakk> \<Longrightarrow> P"
 22.1486 +proof -
 22.1487 +  let ?inherited   = "filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m) 
 22.1488 +                              (methd G (super c))"
 22.1489 +  let ?new = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
 22.1490 +  from ws clsC neq_C_Obj methd_C 
 22.1491 +  have methd_unfold: "(?inherited ++ ?new) sig = Some m"
 22.1492 +    by (simp add: methd_rec)
 22.1493 +  assume NewMethod: "?new sig = Some m \<Longrightarrow> P"
 22.1494 +  assume InheritedMethod: "\<lbrakk>G\<turnstile>C inherits (method sig m); 
 22.1495 +                            methd G (super c) sig = Some m\<rbrakk> \<Longrightarrow> P"
 22.1496 +  show P
 22.1497 +  proof (cases "?new sig")
 22.1498 +    case None
 22.1499 +    with methd_unfold have "?inherited sig = Some m"
 22.1500 +      by (auto)
 22.1501 +    with InheritedMethod show P by blast
 22.1502 +  next
 22.1503 +    case Some
 22.1504 +    with methd_unfold have "?new sig = Some m"
 22.1505 +      by auto
 22.1506 +    with NewMethod show P by blast
 22.1507 +  qed
 22.1508 +qed
 22.1509 +
 22.1510 +  
 22.1511 +lemma methd_member_of:
 22.1512 + (assumes wf: "wf_prog G") 
 22.1513 +  "\<lbrakk>is_class G C; methd G C sig = Some m\<rbrakk> \<Longrightarrow> G\<turnstile>Methd sig m member_of C" 
 22.1514 +  (is "?Class C \<Longrightarrow> ?Method C \<Longrightarrow> ?MemberOf C") 
 22.1515 +proof -
 22.1516 +  from wf   have   ws: "ws_prog G" ..
 22.1517 +  assume defC: "is_class G C"
 22.1518 +  from defC ws 
 22.1519 +  show "?Class C \<Longrightarrow> ?Method C \<Longrightarrow> ?MemberOf C"
 22.1520 +  proof (induct rule: ws_class_induct')  
 22.1521 +    case Object
 22.1522 +    with wf have declC: "declclass m = Object"
 22.1523 +      by (blast intro: declclass_methd_Object)
 22.1524 +    with Object wf have "G\<turnstile>Methd sig m declared_in Object"
 22.1525 +      by (auto dest: methd_declaredD simp del: methd_Object)
 22.1526 +    with declC 
 22.1527 +    show "?MemberOf Object"
 22.1528 +      by (auto intro!: members.Immediate
 22.1529 +                  simp del: methd_Object)
 22.1530 +  next
 22.1531 +    case (Subcls C c)
 22.1532 +    assume  clsC: "class G C = Some c" and
 22.1533 +       neq_C_Obj: "C \<noteq> Object"  
 22.1534 +    assume methd: "?Method C"
 22.1535 +    from methd ws clsC neq_C_Obj
 22.1536 +    show "?MemberOf C"
 22.1537 +    proof (cases rule: methd_rec_Some_cases)
 22.1538 +      case NewMethod
 22.1539 +      with clsC show ?thesis
 22.1540 +	by (auto dest: method_declared_inI intro!: members.Immediate)
 22.1541 +    next
 22.1542 +      case InheritedMethod
 22.1543 +      then show "?thesis"
 22.1544 +	by (blast dest: inherits_member)
 22.1545 +    qed
 22.1546 +  qed
 22.1547 +qed
 22.1548 +
 22.1549 +lemma current_methd: 
 22.1550 +      "\<lbrakk>table_of (methods c) sig = Some new;
 22.1551 +        ws_prog G; class G C = Some c; C \<noteq> Object; 
 22.1552 +        methd G (super c) sig = Some old\<rbrakk> 
 22.1553 +    \<Longrightarrow> methd G C sig = Some (C,new)"
 22.1554 +by (auto simp add: methd_rec
 22.1555 +            intro: filter_tab_SomeI override_find_right table_of_map_SomeI)
 22.1556 +
 22.1557 +lemma wf_prog_staticD:
 22.1558 + (assumes     wf: "wf_prog G" and 
 22.1559 +            clsC: "class G C = Some c" and
 22.1560 +       neq_C_Obj: "C \<noteq> Object" and 
 22.1561 +             old: "methd G (super c) sig = Some old" and 
 22.1562 +     accmodi_old: "Protected \<le> accmodi old" and
 22.1563 +             new: "table_of (methods c) sig = Some new"
 22.1564 + ) "is_static new = is_static old"
 22.1565 +proof -
 22.1566 +  from clsC wf 
 22.1567 +  have wf_cdecl: "wf_cdecl G (C,c)" by (rule wf_prog_cdecl)
 22.1568 +  from wf clsC neq_C_Obj
 22.1569 +  have is_cls_super: "is_class G (super c)" 
 22.1570 +    by (blast dest: wf_prog_acc_superD is_acc_classD)
 22.1571 +  from wf is_cls_super old 
 22.1572 +  have old_member_of: "G\<turnstile>Methd sig old member_of (super c)"  
 22.1573 +    by (rule methd_member_of)
 22.1574 +  from old wf is_cls_super 
 22.1575 +  have old_declared: "G\<turnstile>Methd sig old declared_in (declclass old)"
 22.1576 +    by (auto dest: methd_declared_in_declclass)
 22.1577 +  from new clsC 
 22.1578 +  have new_declared: "G\<turnstile>Methd sig (C,new) declared_in C"
 22.1579 +    by (auto intro: method_declared_inI)
 22.1580 +  note trancl_rtrancl_tranc = trancl_rtrancl_trancl [trans] (* ### in Basis *)
 22.1581 +  from clsC neq_C_Obj
 22.1582 +  have subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
 22.1583 +    by (rule subcls1I)
 22.1584 +  then have "G\<turnstile>C \<prec>\<^sub>C super c" ..
 22.1585 +  also from old wf is_cls_super
 22.1586 +  have "G\<turnstile>super c \<preceq>\<^sub>C (declclass old)" by (auto dest: methd_declC)
 22.1587 +  finally have subcls_C_old:  "G\<turnstile>C \<prec>\<^sub>C (declclass old)" .
 22.1588 +  from accmodi_old 
 22.1589 +  have inheritable: "G\<turnstile>Methd sig old inheritable_in pid C"
 22.1590 +    by (auto simp add: inheritable_in_def
 22.1591 +                 dest: acc_modi_le_Dests)
 22.1592 +  show ?thesis
 22.1593 +  proof (cases "is_static new")
 22.1594 +    case True
 22.1595 +    with subcls_C_old new_declared old_declared inheritable
 22.1596 +    have "G,sig\<turnstile>(C,new) hides old"
 22.1597 +      by (auto intro: hidesI)
 22.1598 +    with True wf_cdecl neq_C_Obj new 
 22.1599 +    show ?thesis
 22.1600 +      by (auto dest: wf_cdecl_hides_SomeD)
 22.1601 +  next
 22.1602 +    case False
 22.1603 +    with subcls_C_old new_declared old_declared inheritable subcls1_C_super
 22.1604 +         old_member_of
 22.1605 +    have "G,sig\<turnstile>(C,new) overrides\<^sub>S old"
 22.1606 +      by (auto intro: stat_overridesR.Direct)
 22.1607 +    with False wf_cdecl neq_C_Obj new 
 22.1608 +    show ?thesis
 22.1609 +      by (auto dest: wf_cdecl_overrides_SomeD)
 22.1610 +  qed
 22.1611 +qed
 22.1612 +
 22.1613 +lemma inheritable_instance_methd: 
 22.1614 + (assumes subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" and
 22.1615 +              is_cls_D: "is_class G D" and
 22.1616 +                    wf: "wf_prog G" and 
 22.1617 +                   old: "methd G D sig = Some old" and
 22.1618 +           accmodi_old: "Protected \<le> accmodi old" and  
 22.1619 +        not_static_old: "\<not> is_static old"
 22.1620 + )  
 22.1621 +  "\<exists>new. methd G C sig = Some new \<and>
 22.1622 +         (new = old \<or> G,sig\<turnstile>new overrides\<^sub>S old)"
 22.1623 + (is "(\<exists>new. (?Constraint C new old))")
 22.1624 +proof -
 22.1625 +  from subclseq_C_D is_cls_D 
 22.1626 +  have is_cls_C: "is_class G C" by (rule subcls_is_class2) 
 22.1627 +  from wf 
 22.1628 +  have ws: "ws_prog G" ..
 22.1629 +  from is_cls_C ws subclseq_C_D 
 22.1630 +  show "\<exists>new. ?Constraint C new old"
 22.1631 +  proof (induct rule: ws_class_induct')
 22.1632 +    case (Object co)
 22.1633 +    then have eq_D_Obj: "D=Object" by auto
 22.1634 +    with old 
 22.1635 +    have "?Constraint Object old old"
 22.1636 +      by auto
 22.1637 +    with eq_D_Obj 
 22.1638 +    show "\<exists> new. ?Constraint Object new old" by auto
 22.1639 +  next
 22.1640 +    case (Subcls C c)
 22.1641 +    assume hyp: "G\<turnstile>super c\<preceq>\<^sub>C D \<Longrightarrow> \<exists>new. ?Constraint (super c) new old"
 22.1642 +    assume clsC: "class G C = Some c"
 22.1643 +    assume neq_C_Obj: "C\<noteq>Object"
 22.1644 +    from clsC wf 
 22.1645 +    have wf_cdecl: "wf_cdecl G (C,c)" 
 22.1646 +      by (rule wf_prog_cdecl)
 22.1647 +    from ws clsC neq_C_Obj
 22.1648 +    have is_cls_super: "is_class G (super c)"
 22.1649 +      by (auto dest: ws_prog_cdeclD)
 22.1650 +    from clsC wf neq_C_Obj 
 22.1651 +    have superAccessible: "G\<turnstile>(Class (super c)) accessible_in (pid C)" and
 22.1652 +	 subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
 22.1653 +      by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD
 22.1654 +              intro: subcls1I)
 22.1655 +    show "\<exists>new. ?Constraint C new old"
 22.1656 +    proof (cases "G\<turnstile>super c\<preceq>\<^sub>C D")
 22.1657 +      case False
 22.1658 +      from False Subcls 
 22.1659 +      have eq_C_D: "C=D"
 22.1660 +	by (auto dest: subclseq_superD)
 22.1661 +      with old 
 22.1662 +      have "?Constraint C old old"
 22.1663 +	by auto
 22.1664 +      with eq_C_D 
 22.1665 +      show "\<exists> new. ?Constraint C new old" by auto
 22.1666 +    next
 22.1667 +      case True
 22.1668 +      with hyp obtain super_method
 22.1669 +	where super: "?Constraint (super c) super_method old" by blast
 22.1670 +      from super not_static_old
 22.1671 +      have not_static_super: "\<not> is_static super_method"
 22.1672 +	by (auto dest!: stat_overrides_commonD)
 22.1673 +      from super old wf accmodi_old
 22.1674 +      have accmodi_super_method: "Protected \<le> accmodi super_method"
 22.1675 +	by (auto dest!: wf_prog_stat_overridesD
 22.1676 +                 intro: order_trans)
 22.1677 +      from super accmodi_old wf
 22.1678 +      have inheritable: "G\<turnstile>Methd sig super_method inheritable_in (pid C)"
 22.1679 +	by (auto dest!: wf_prog_stat_overridesD
 22.1680 +                        acc_modi_le_Dests
 22.1681 +              simp add: inheritable_in_def)	           
 22.1682 +      from super wf is_cls_super
 22.1683 +      have member: "G\<turnstile>Methd sig super_method member_of (super c)"
 22.1684 +	by (auto intro: methd_member_of) 
 22.1685 +      from member
 22.1686 +      have decl_super_method:
 22.1687 +        "G\<turnstile>Methd sig super_method declared_in (declclass super_method)"
 22.1688 +	by (auto dest: member_of_declC)
 22.1689 +      from super subcls1_C_super ws is_cls_super 
 22.1690 +      have subcls_C_super: "G\<turnstile>C \<prec>\<^sub>C (declclass super_method)"
 22.1691 +	by (auto intro: rtrancl_into_trancl2 dest: methd_declC) 
 22.1692 +      show "\<exists> new. ?Constraint C new old"
 22.1693 +      proof (cases "methd G C sig")
 22.1694 +	case None
 22.1695 +	have "methd G (super c) sig = None"
 22.1696 +	proof -
 22.1697 +	  from clsC ws None 
 22.1698 +	  have no_new: "table_of (methods c) sig = None" 
 22.1699 +	    by (auto simp add: methd_rec)
 22.1700 +	  with clsC 
 22.1701 +	  have undeclared: "G\<turnstile>mid sig undeclared_in C"
 22.1702 +	    by (auto simp add: undeclared_in_def cdeclaredmethd_def)
 22.1703 +	  with inheritable member superAccessible subcls1_C_super
 22.1704 +	  have inherits: "G\<turnstile>C inherits (method sig super_method)"
 22.1705 +	    by (auto simp add: inherits_def)
 22.1706 +	  with clsC ws no_new super neq_C_Obj
 22.1707 +	  have "methd G C sig = Some super_method"
 22.1708 +	    by (auto simp add: methd_rec override_def
 22.1709 +	                intro: filter_tab_SomeI)
 22.1710 +          with None show ?thesis
 22.1711 +	    by simp
 22.1712 +	qed
 22.1713 +	with super show ?thesis by auto
 22.1714 +      next
 22.1715 +	case (Some new)
 22.1716 +	from this ws clsC neq_C_Obj
 22.1717 +	show ?thesis
 22.1718 +	proof (cases rule: methd_rec_Some_cases)
 22.1719 +	  case InheritedMethod
 22.1720 +	  with super Some show ?thesis 
 22.1721 +	    by auto
 22.1722 +	next
 22.1723 +	  case NewMethod
 22.1724 +	  assume new: "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig 
 22.1725 +                       = Some new"
 22.1726 +	  from new 
 22.1727 +	  have declcls_new: "declclass new = C" 
 22.1728 +	    by auto
 22.1729 +	  from wf clsC neq_C_Obj super new not_static_super accmodi_super_method
 22.1730 +	  have not_static_new: "\<not> is_static new" 
 22.1731 +	    by (auto dest: wf_prog_staticD) 
 22.1732 +	  from clsC new
 22.1733 +	  have decl_new: "G\<turnstile>Methd sig new declared_in C"
 22.1734 +	    by (auto simp add: declared_in_def cdeclaredmethd_def)
 22.1735 +	  from not_static_new decl_new decl_super_method
 22.1736 +	       member subcls1_C_super inheritable declcls_new subcls_C_super 
 22.1737 +	  have "G,sig\<turnstile> new overrides\<^sub>S super_method"
 22.1738 +	    by (auto intro: stat_overridesR.Direct) 
 22.1739 +	  with super Some
 22.1740 +	  show ?thesis
 22.1741 +	    by (auto intro: stat_overridesR.Indirect)
 22.1742 +	qed
 22.1743 +      qed
 22.1744 +    qed
 22.1745 +  qed
 22.1746 +qed
 22.1747 +
 22.1748 +lemma inheritable_instance_methd_cases [consumes 6
 22.1749 +                                       , case_names Inheritance Overriding]: 
 22.1750 + (assumes subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" and
 22.1751 +              is_cls_D: "is_class G D" and
 22.1752 +                    wf: "wf_prog G" and 
 22.1753 +                   old: "methd G D sig = Some old" and
 22.1754 +           accmodi_old: "Protected \<le> accmodi old" and  
 22.1755 +        not_static_old: "\<not> is_static old" and
 22.1756 +           inheritance:  "methd G C sig = Some old \<Longrightarrow> P" and
 22.1757 +            overriding:  "\<And> new. \<lbrakk>methd G C sig = Some new;
 22.1758 +                                   G,sig\<turnstile>new overrides\<^sub>S old\<rbrakk> \<Longrightarrow> P"
 22.1759 +        
 22.1760 + ) "P"
 22.1761 +proof -
 22.1762 +from subclseq_C_D is_cls_D wf old accmodi_old not_static_old 
 22.1763 +     inheritance overriding
 22.1764 +show ?thesis
 22.1765 +  by (auto dest: inheritable_instance_methd)
 22.1766 +qed
 22.1767 +
 22.1768 +lemma inheritable_instance_methd_props: 
 22.1769 + (assumes subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" and
 22.1770 +              is_cls_D: "is_class G D" and
 22.1771 +                    wf: "wf_prog G" and 
 22.1772 +                   old: "methd G D sig = Some old" and
 22.1773 +           accmodi_old: "Protected \<le> accmodi old" and  
 22.1774 +        not_static_old: "\<not> is_static old"
 22.1775 + )  
 22.1776 +  "\<exists>new. methd G C sig = Some new \<and>
 22.1777 +          \<not> is_static new \<and> G\<turnstile>resTy new\<preceq>resTy old \<and> accmodi old \<le>accmodi new"
 22.1778 + (is "(\<exists>new. (?Constraint C new old))")
 22.1779 +proof -
 22.1780 +  from subclseq_C_D is_cls_D wf old accmodi_old not_static_old 
 22.1781 +  show ?thesis
 22.1782 +  proof (cases rule: inheritable_instance_methd_cases)
 22.1783 +    case Inheritance
 22.1784 +    with not_static_old accmodi_old show ?thesis by auto
 22.1785 +  next
 22.1786 +    case (Overriding new)
 22.1787 +    then have "\<not> is_static new" by (auto dest: stat_overrides_commonD)
 22.1788 +    with Overriding not_static_old accmodi_old wf 
 22.1789 +    show ?thesis 
 22.1790 +      by (auto dest!: wf_prog_stat_overridesD
 22.1791 +               intro: order_trans)
 22.1792 +  qed
 22.1793 +qed
 22.1794 + 	  
 22.1795 +(* ### Probleme: Die tollen methd_subcls_cases Lemma wird warscheinlich
 22.1796 +  kaum gebraucht: 
 22.1797 +Redundanz: stat_overrides.Direct old declared in declclass old folgt aus
 22.1798 +           member of 
 22.1799 +   Problem: Predikate wie overrides, sind global üper die Hierarchie hinweg
 22.1800 +            definiert, aber oft barucht man eben zusätlich Induktion
 22.1801 +            : von super c auf C; Dann ist aber auss dem Kontext
 22.1802 +            die Unterscheidung in die 5 fälle overkill,
 22.1803 +            da man dann warscheinlich meistens eh in einem speziellen
 22.1804 +            Fall kommt (durch die Hypothesen)
 22.1805 +*)
 22.1806 +    
 22.1807 +(* local lemma *)
 22.1808 +ML {* bind_thm("bexI'",permute_prems 0 1 bexI) *}
 22.1809 +ML {* bind_thm("ballE'",permute_prems 1 1 ballE) *}
 22.1810 +lemma subint_widen_imethds: 
 22.1811 + "\<lbrakk>G\<turnstile>I\<preceq>I J; wf_prog G; is_iface G J; jm \<in> imethds G J sig\<rbrakk> \<Longrightarrow>   
 22.1812 +  \<exists> im \<in> imethds G I sig. is_static im = is_static jm \<and> 
 22.1813 +                          accmodi im = accmodi jm \<and>
 22.1814 +                          G\<turnstile>resTy im\<preceq>resTy jm"
 22.1815 +proof -
 22.1816 +  assume irel: "G\<turnstile>I\<preceq>I J" and
 22.1817 +           wf: "wf_prog G" and
 22.1818 +     is_iface: "is_iface G J"
 22.1819 +  from irel show "jm \<in> imethds G J sig \<Longrightarrow> ?thesis" 
 22.1820 +               (is "PROP ?P I" is "PROP ?Prem J \<Longrightarrow> ?Concl I")
 22.1821 +  proof (induct ?P I rule: converse_rtrancl_induct) 
 22.1822 +    case Id
 22.1823 +    assume "jm \<in> imethds G J sig"
 22.1824 +    then show "?Concl J" by  (blast elim: bexI')
 22.1825 +  next
 22.1826 +    case Step
 22.1827 +    fix I SI
 22.1828 +    assume subint1_I_SI: "G\<turnstile>I \<prec>I1 SI" and 
 22.1829 +            subint_SI_J: "G\<turnstile>SI \<preceq>I J" and
 22.1830 +                    hyp: "PROP ?P SI" and
 22.1831 +                     jm: "jm \<in> imethds G J sig"
 22.1832 +    from subint1_I_SI 
 22.1833 +    obtain i where
 22.1834 +      ifI: "iface G I = Some i" and
 22.1835 +       SI: "SI \<in> set (isuperIfs i)"
 22.1836 +      by (blast dest: subint1D)
 22.1837 +
 22.1838 +    let ?newMethods 
 22.1839 +          = "(o2s \<circ> table_of (map (\<lambda>(sig, mh). (sig, I, mh)) (imethods i)))"
 22.1840 +    show "?Concl I"
 22.1841 +    proof (cases "?newMethods sig = {}")
 22.1842 +      case True
 22.1843 +      with ifI SI hyp wf jm 
 22.1844 +      show "?thesis" 
 22.1845 +	by (auto simp add: imethds_rec) 
 22.1846 +    next
 22.1847 +      case False
 22.1848 +      from ifI wf False
 22.1849 +      have imethds: "imethds G I sig = ?newMethods sig"
 22.1850 +	by (simp add: imethds_rec)
 22.1851 +      from False
 22.1852 +      obtain im where
 22.1853 +        imdef: "im \<in> ?newMethods sig" 
 22.1854 +	by (blast)
 22.1855 +      with imethds 
 22.1856 +      have im: "im \<in> imethds G I sig"
 22.1857 +	by (blast)
 22.1858 +      with im wf ifI 
 22.1859 +      obtain
 22.1860 +	 imStatic: "\<not> is_static im" and
 22.1861 +         imPublic: "accmodi im = Public"
 22.1862 +	by (auto dest!: imethds_wf_mhead)	
 22.1863 +      from ifI wf 
 22.1864 +      have wf_I: "wf_idecl G (I,i)" 
 22.1865 +	by (rule wf_prog_idecl)
 22.1866 +      with SI wf  
 22.1867 +      obtain si where
 22.1868 +	 ifSI: "iface G SI = Some si" and
 22.1869 +	wf_SI: "wf_idecl G (SI,si)" 
 22.1870 +	by (auto dest!: wf_idecl_supD is_acc_ifaceD
 22.1871 +                  dest: wf_prog_idecl)
 22.1872 +      from jm hyp 
 22.1873 +      obtain sim::"qtname \<times> mhead"  where
 22.1874 +                      sim: "sim \<in> imethds G SI sig" and
 22.1875 +         eq_static_sim_jm: "is_static sim = is_static jm" and 
 22.1876 +         eq_access_sim_jm: "accmodi sim = accmodi jm" and 
 22.1877 +        resTy_widen_sim_jm: "G\<turnstile>resTy sim\<preceq>resTy jm"
 22.1878 +	by blast
 22.1879 +      with wf_I SI imdef sim 
 22.1880 +      have "G\<turnstile>resTy im\<preceq>resTy sim"   
 22.1881 +	by (auto dest!: wf_idecl_hidings hidings_entailsD)
 22.1882 +      with wf resTy_widen_sim_jm 
 22.1883 +      have resTy_widen_im_jm: "G\<turnstile>resTy im\<preceq>resTy jm"
 22.1884 +	by (blast intro: widen_trans)
 22.1885 +      from sim wf ifSI  
 22.1886 +      obtain
 22.1887 +	simStatic: "\<not> is_static sim" and
 22.1888 +        simPublic: "accmodi sim = Public"
 22.1889 +	by (auto dest!: imethds_wf_mhead)
 22.1890 +      from im 
 22.1891 +           imStatic simStatic eq_static_sim_jm
 22.1892 +           imPublic simPublic eq_access_sim_jm
 22.1893 +           resTy_widen_im_jm
 22.1894 +      show ?thesis 
 22.1895 +	by auto 
 22.1896 +    qed
 22.1897 +  qed
 22.1898 +qed
 22.1899 +     
 22.1900 +(* Tactical version *)
 22.1901 +(* 
 22.1902 +lemma subint_widen_imethds: "\<lbrakk>G\<turnstile>I\<preceq>I J; wf_prog G; is_iface G J\<rbrakk> \<Longrightarrow>  
 22.1903 +  \<forall> jm \<in> imethds G J sig.  
 22.1904 +  \<exists> im \<in> imethds G I sig. static (mthd im)=static (mthd jm) \<and> 
 22.1905 +                          access (mthd im)= access (mthd jm) \<and>
 22.1906 +                          G\<turnstile>resTy (mthd im)\<preceq>resTy (mthd jm)"
 22.1907 +apply (erule converse_rtrancl_induct)
 22.1908 +apply  (clarsimp elim!: bexI')
 22.1909 +apply (frule subint1D)
 22.1910 +apply clarify
 22.1911 +apply (erule ballE')
 22.1912 +apply  fast
 22.1913 +apply (erule_tac V = "?x \<in> imethds G J sig" in thin_rl)
 22.1914 +apply clarsimp
 22.1915 +apply (subst imethds_rec, assumption, erule wf_ws_prog)
 22.1916 +apply (unfold overrides_t_def)
 22.1917 +apply (drule (1) wf_prog_idecl)
 22.1918 +apply (frule (3) imethds_wf_mhead [OF _ _ wf_idecl_supD [THEN conjunct1 
 22.1919 +                                       [THEN is_acc_ifaceD [THEN conjunct1]]]])
 22.1920 +apply (case_tac "(o2s \<circ> table_of (map (\<lambda>(s, mh). (s, y, mh)) (imethods i)))
 22.1921 +                  sig ={}")
 22.1922 +apply   force
 22.1923 +
 22.1924 +apply   (simp only:)
 22.1925 +apply   (simp)
 22.1926 +apply   clarify
 22.1927 +apply   (frule wf_idecl_hidings [THEN hidings_entailsD])
 22.1928 +apply     blast
 22.1929 +apply     blast
 22.1930 +apply   (rule bexI')
 22.1931 +apply     simp
 22.1932 +apply     (drule table_of_map_SomeI [of _ "sig"])
 22.1933 +apply     simp
 22.1934 +
 22.1935 +apply     (frule wf_idecl_mhead [of _ _ _ "sig"])
 22.1936 +apply       (rule table_of_Some_in_set)
 22.1937 +apply       assumption
 22.1938 +apply     auto
 22.1939 +done
 22.1940 +*)
 22.1941 +    
 22.1942 +
 22.1943 +(* local lemma *)
 22.1944 +lemma implmt1_methd: 
 22.1945 + "\<And>sig. \<lbrakk>G\<turnstile>C\<leadsto>1I; wf_prog G; im \<in> imethds G I sig\<rbrakk> \<Longrightarrow>  
 22.1946 +  \<exists>cm \<in>methd G C sig: \<not> is_static cm \<and> \<not> is_static im \<and> 
 22.1947 +                       G\<turnstile>resTy cm\<preceq>resTy im \<and>
 22.1948 +                       accmodi im = Public \<and> accmodi cm = Public"
 22.1949 +apply (drule implmt1D)
 22.1950 +apply clarify
 22.1951 +apply (drule (2) wf_prog_cdecl [THEN wf_cdecl_impD])
 22.1952 +apply (frule (1) imethds_wf_mhead)
 22.1953 +apply  (simp add: is_acc_iface_def)
 22.1954 +apply (force)
 22.1955 +done
 22.1956 +
 22.1957 +
 22.1958 +(* local lemma *)
 22.1959 +lemma implmt_methd [rule_format (no_asm)]: 
 22.1960 +"\<lbrakk>wf_prog G; G\<turnstile>C\<leadsto>I\<rbrakk> \<Longrightarrow> is_iface G I \<longrightarrow>  
 22.1961 + (\<forall> im    \<in>imethds G I   sig.  
 22.1962 +  \<exists> cm\<in>methd G C sig: \<not>is_static cm \<and> \<not> is_static im \<and> 
 22.1963 +                      G\<turnstile>resTy cm\<preceq>resTy im \<and>
 22.1964 +                      accmodi im = Public \<and> accmodi cm = Public)"
 22.1965 +apply (frule implmt_is_class)
 22.1966 +apply (erule implmt.induct)
 22.1967 +apply   safe
 22.1968 +apply   (drule (2) implmt1_methd)
 22.1969 +apply   fast
 22.1970 +apply  (drule (1) subint_widen_imethds)
 22.1971 +apply   simp
 22.1972 +apply   assumption
 22.1973 +apply  clarify
 22.1974 +apply  (drule (2) implmt1_methd)
 22.1975 +apply  (force)
 22.1976 +apply (frule subcls1D)
 22.1977 +apply (drule (1) bspec)
 22.1978 +apply clarify
 22.1979 +apply (drule (3) r_into_rtrancl [THEN inheritable_instance_methd_props, 
 22.1980 +                                 OF _ implmt_is_class])
 22.1981 +apply auto 
 22.1982 +done
 22.1983 +
 22.1984 +
 22.1985 +declare split_paired_All [simp] split_paired_Ex [simp]
 22.1986 +ML_setup {*
 22.1987 +claset_ref()  := claset() addSbefore ("split_all_tac", split_all_tac);
 22.1988 +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
 22.1989 +*}
 22.1990 +lemma mheadsD [rule_format (no_asm)]: 
 22.1991 +"emh \<in> mheads G S t sig \<longrightarrow> wf_prog G \<longrightarrow>
 22.1992 + (\<exists>C D m. t = ClassT C \<and> declrefT emh = ClassT D \<and> 
 22.1993 +          accmethd G S C sig = Some m \<and>
 22.1994 +          (declclass m = D) \<and> mhead (mthd m) = (mhd emh)) \<or>
 22.1995 + (\<exists>I. t = IfaceT I \<and> ((\<exists>im. im  \<in> accimethds G (pid S) I sig \<and> 
 22.1996 +          mthd im = mhd emh) \<or> 
 22.1997 +  (\<exists>m. G\<turnstile>Iface I accessible_in (pid S) \<and> accmethd G S Object sig = Some m \<and> 
 22.1998 +       accmodi m \<noteq> Private \<and> 
 22.1999 +       declrefT emh = ClassT Object \<and> mhead (mthd m) = mhd emh))) \<or>
 22.2000 + (\<exists>T m. t = ArrayT T \<and> G\<turnstile>Array T accessible_in (pid S) \<and>
 22.2001 +        accmethd G S Object sig = Some m \<and> accmodi m \<noteq> Private \<and> 
 22.2002 +        declrefT emh = ClassT Object \<and> mhead (mthd m) = mhd emh)"
 22.2003 +apply (rule_tac "ref_ty1"="t" in ref_ty_ty.induct [THEN conjunct1])
 22.2004 +apply auto
 22.2005 +apply (auto simp add: cmheads_def accObjectmheads_def Objectmheads_def)
 22.2006 +apply (auto  dest!: accmethd_SomeD)
 22.2007 +done
 22.2008 +
 22.2009 +lemma mheads_cases [consumes 2, case_names Class_methd 
 22.2010 +                    Iface_methd Iface_Object_methd Array_Object_methd]: 
 22.2011 +"\<lbrakk>emh \<in> mheads G S t sig; wf_prog G;
 22.2012 + \<And> C D m. \<lbrakk>t = ClassT C;declrefT emh = ClassT D; accmethd G S C sig = Some m;
 22.2013 +           (declclass m = D); mhead (mthd m) = (mhd emh)\<rbrakk> \<Longrightarrow> P emh; 
 22.2014 + \<And> I im. \<lbrakk>t = IfaceT I; im  \<in> accimethds G (pid S) I sig; mthd im = mhd emh\<rbrakk>
 22.2015 +          \<Longrightarrow> P emh;
 22.2016 + \<And> I m. \<lbrakk>t = IfaceT I; G\<turnstile>Iface I accessible_in (pid S);
 22.2017 +          accmethd G S Object sig = Some m; accmodi m \<noteq> Private;
 22.2018 +         declrefT emh = ClassT Object; mhead (mthd m) = mhd emh\<rbrakk> \<Longrightarrow> P emh;
 22.2019 + \<And> T m. \<lbrakk>t = ArrayT T;G\<turnstile>Array T accessible_in (pid S);
 22.2020 +          accmethd G S Object sig = Some m; accmodi m \<noteq> Private; 
 22.2021 +          declrefT emh = ClassT Object; mhead (mthd m) = mhd emh\<rbrakk> \<Longrightarrow>  P emh 
 22.2022 +\<rbrakk> \<Longrightarrow> P emh"
 22.2023 +by (blast dest!: mheadsD)
 22.2024 +
 22.2025 +lemma declclassD[rule_format]:
 22.2026 + "\<lbrakk>wf_prog G;class G C = Some c; methd G C sig = Some m; 
 22.2027 +   class G (declclass m) = Some d\<rbrakk>
 22.2028 +  \<Longrightarrow> table_of (methods d) sig  = Some (mthd m)"
 22.2029 +proof -
 22.2030 +  assume    wf: "wf_prog G"
 22.2031 +  then have ws: "ws_prog G" ..
 22.2032 +  assume  clsC: "class G C = Some c"
 22.2033 +  from clsC ws 
 22.2034 +  show "\<And> m d. \<lbrakk>methd G C sig = Some m; class G (declclass m) = Some d\<rbrakk>
 22.2035 +        \<Longrightarrow> table_of (methods d) sig  = Some (mthd m)" 
 22.2036 +         (is "PROP ?P C") 
 22.2037 +  proof (induct ?P C rule: ws_class_induct)
 22.2038 +    case Object
 22.2039 +    fix m d
 22.2040 +    assume "methd G Object sig = Some m" 
 22.2041 +           "class G (declclass m) = Some d"
 22.2042 +    with wf show "?thesis m d" by auto
 22.2043 +  next
 22.2044 +    case Subcls
 22.2045 +    fix C c m d
 22.2046 +    assume hyp: "PROP ?P (super c)"
 22.2047 +    and      m: "methd G C sig = Some m"
 22.2048 +    and  declC: "class G (declclass m) = Some d"
 22.2049 +    and   clsC: "class G C = Some c"
 22.2050 +    and   nObj: "C \<noteq> Object"
 22.2051 +    let ?newMethods = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig"
 22.2052 +    show "?thesis m d" 
 22.2053 +    proof (cases "?newMethods")
 22.2054 +      case None
 22.2055 +      from None clsC nObj ws m declC hyp  
 22.2056 +      show "?thesis" by (auto simp add: methd_rec)
 22.2057 +    next
 22.2058 +      case Some
 22.2059 +      from Some clsC nObj ws m declC hyp  
 22.2060 +      show "?thesis" 
 22.2061 +	by (auto simp add: methd_rec
 22.2062 +                     dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
 22.2063 +    qed
 22.2064 +  qed
 22.2065 +qed
 22.2066 +
 22.2067 +(* Tactical version *)
 22.2068 +(*
 22.2069 +lemma declclassD[rule_format]:
 22.2070 + "wf_prog G \<longrightarrow>  
 22.2071 + (\<forall> c d m. class G C = Some c \<longrightarrow> methd G C sig = Some m \<longrightarrow> 
 22.2072 +  class G (declclass m) = Some d
 22.2073 + \<longrightarrow> table_of (methods d) sig  = Some (mthd m))"
 22.2074 +apply (rule class_rec.induct)
 22.2075 +apply (rule impI)
 22.2076 +apply (rule allI)+
 22.2077 +apply (rule impI)
 22.2078 +apply (case_tac "C=Object")
 22.2079 +apply   (force simp add: methd_rec)
 22.2080 +
 22.2081 +apply   (subst methd_rec)
 22.2082 +apply     (blast dest: wf_ws_prog)+
 22.2083 +apply   (case_tac "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig")
 22.2084 +apply     (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
 22.2085 +done
 22.2086 +*)
 22.2087 +
 22.2088 +lemma dynmethd_Object:
 22.2089 + (assumes statM: "methd G Object sig = Some statM" and
 22.2090 +        private: "accmodi statM = Private" and 
 22.2091 +       is_cls_C: "is_class G C" and
 22.2092 +             wf: "wf_prog G"
 22.2093 + ) "dynmethd G Object C sig = Some statM"
 22.2094 +proof -
 22.2095 +  from is_cls_C wf 
 22.2096 +  have subclseq: "G\<turnstile>C \<preceq>\<^sub>C Object" 
 22.2097 +    by (auto intro: subcls_ObjectI)
 22.2098 +  from wf have ws: "ws_prog G" 
 22.2099 +    by simp
 22.2100 +  from wf 
 22.2101 +  have is_cls_Obj: "is_class G Object" 
 22.2102 +    by simp
 22.2103 +  from statM subclseq is_cls_Obj ws private
 22.2104 +  show ?thesis
 22.2105 +  proof (cases rule: dynmethd_cases)
 22.2106 +    case Static then show ?thesis .
 22.2107 +  next
 22.2108 +    case Overrides 
 22.2109 +    with private show ?thesis 
 22.2110 +      by (auto dest: no_Private_override)
 22.2111 +  qed
 22.2112 +qed
 22.2113 +
 22.2114 +lemma wf_imethds_hiding_objmethdsD: 
 22.2115 +  (assumes     old: "methd G Object sig = Some old" and
 22.2116 +           is_if_I: "is_iface G I" and
 22.2117 +                wf: "wf_prog G" and    
 22.2118 +       not_private: "accmodi old \<noteq> Private" and
 22.2119 +               new: "new \<in> imethds G I sig" 
 22.2120 +  ) "G\<turnstile>resTy new\<preceq>resTy old \<and> is_static new = is_static old" (is "?P new")
 22.2121 +proof -
 22.2122 +  from wf have ws: "ws_prog G" by simp
 22.2123 +  {
 22.2124 +    fix I i new
 22.2125 +    assume ifI: "iface G I = Some i"
 22.2126 +    assume new: "table_of (imethods i) sig = Some new" 
 22.2127 +    from ifI new not_private wf old  
 22.2128 +    have "?P (I,new)"
 22.2129 +      by (auto dest!: wf_prog_idecl wf_idecl_hiding cond_hiding_entailsD
 22.2130 +            simp del: methd_Object)
 22.2131 +  } note hyp_newmethod = this  
 22.2132 +  from is_if_I ws new 
 22.2133 +  show ?thesis
 22.2134 +  proof (induct rule: ws_interface_induct)
 22.2135 +    case (Step I i)
 22.2136 +    assume ifI: "iface G I = Some i" 
 22.2137 +    assume new: "new \<in> imethds G I sig" 
 22.2138 +    from Step
 22.2139 +    have hyp: "\<forall> J \<in> set (isuperIfs i). (new \<in> imethds G J sig \<longrightarrow> ?P new)"
 22.2140 +      by auto 
 22.2141 +    from new ifI ws
 22.2142 +    show "?P new"
 22.2143 +    proof (cases rule: imethds_cases)
 22.2144 +      case NewMethod
 22.2145 +      with ifI hyp_newmethod
 22.2146 +      show ?thesis
 22.2147 +	by auto
 22.2148 +    next
 22.2149 +      case (InheritedMethod J)
 22.2150 +      assume "J \<in> set (isuperIfs i)" 
 22.2151 +             "new \<in> imethds G J sig"
 22.2152 +      with hyp 
 22.2153 +      show "?thesis"
 22.2154 +	by auto
 22.2155 +    qed
 22.2156 +  qed
 22.2157 +qed
 22.2158 +
 22.2159 +text {*
 22.2160 +Which dynamic classes are valid to look up a member of a distinct static type?
 22.2161 +We have to distinct class members (named static members in Java) 
 22.2162 +from instance members. Class members are global to all Objects of a class,
 22.2163 +instance members are local to a single Object instance. If a member is
 22.2164 +equipped with the static modifier it is a class member, else it is an 
 22.2165 +instance member.
 22.2166 +The following table gives an overview of the current framework. We assume
 22.2167 +to have a reference with static type statT and a dynamic class dynC. Between
 22.2168 +both of these types the widening relation holds 
 22.2169 +@{term "G\<turnstile>Class dynC\<preceq> statT"}. Unfortunately this ordinary widening relation 
 22.2170 +isn't enough to describe the valid lookup classes, since we must cope the
 22.2171 +special cases of arrays and interfaces,too. If we statically expect an array or
 22.2172 +inteface we may lookup a field or a method in Object which isn't covered in 
 22.2173 +the widening relation.
 22.2174 +\begin{verbatim}
 22.2175 +statT      field         instance method       static (class) method
 22.2176 +------------------------------------------------------------------------
 22.2177 + NullT      /                  /                   /
 22.2178 + Iface      /                dynC                Object
 22.2179 + Class    dynC               dynC                 dynC
 22.2180 + Array      /                Object              Object
 22.2181 +\end{verbatim}
 22.2182 +In most cases we con lookup the member in the dynamic class. But as an
 22.2183 +interface can't declare new static methods, nor an array can define new
 22.2184 +methods at all, we have to lookup methods in the base class Object.
 22.2185 +
 22.2186 +The limitation to classes in the field column is artificial  and comes out
 22.2187 +of the typing rule for the field access (see rule @{text "FVar"} in the 
 22.2188 +welltyping relation @{term "wt"} in theory WellType). 
 22.2189 +I stems out of the fact, that Object
 22.2190 +indeed has no non private fields. So interfaces and arrays can actually
 22.2191 +have no fields at all and a field access would be senseless. (In Java
 22.2192 +interfaces are allowed to declare new fields but in current Bali not!).
 22.2193 +So there is no principal reason why we should not allow Objects to declare
 22.2194 +non private fields. Then we would get the following column:
 22.2195 +\begin{verbatim}
 22.2196 + statT    field
 22.2197 +----------------- 
 22.2198 + NullT      /  
 22.2199 + Iface    Object 
 22.2200 + Class    dynC 
 22.2201 + Array    Object
 22.2202 +\end{verbatim}
 22.2203 +*}
 22.2204 +consts valid_lookup_cls:: "prog \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> bool \<Rightarrow> bool"
 22.2205 +                        ("_,_ \<turnstile> _ valid'_lookup'_cls'_for _" [61,61,61,61] 60)
 22.2206 +primrec
 22.2207 +"G,NullT    \<turnstile> dynC valid_lookup_cls_for static_membr = False"
 22.2208 +"G,IfaceT I \<turnstile> dynC valid_lookup_cls_for static_membr 
 22.2209 +              = (if static_membr 
 22.2210 +                    then dynC=Object 
 22.2211 +                    else G\<turnstile>Class dynC\<preceq> Iface I)"
 22.2212 +"G,ClassT C \<turnstile> dynC valid_lookup_cls_for static_membr = G\<turnstile>Class dynC\<preceq> Class C"
 22.2213 +"G,ArrayT T \<turnstile> dynC valid_lookup_cls_for static_membr = (dynC=Object)"
 22.2214 +
 22.2215 +lemma valid_lookup_cls_is_class:
 22.2216 + (assumes dynC: "G,statT \<turnstile> dynC valid_lookup_cls_for static_membr" and
 22.2217 +      ty_statT: "isrtype G statT" and
 22.2218 +            wf: "wf_prog G"
 22.2219 + ) "is_class G dynC"
 22.2220 +proof (cases statT)
 22.2221 +  case NullT
 22.2222 +  with dynC ty_statT show ?thesis
 22.2223 +    by (auto dest: widen_NT2)
 22.2224 +next
 22.2225 +  case (IfaceT I)
 22.2226 +  with dynC wf show ?thesis
 22.2227 +    by (auto dest: implmt_is_class)
 22.2228 +next
 22.2229 +  case (ClassT C)
 22.2230 +  with dynC ty_statT show ?thesis
 22.2231 +    by (auto dest: subcls_is_class2)
 22.2232 +next
 22.2233 +  case (ArrayT T)
 22.2234 +  with dynC wf show ?thesis
 22.2235 +    by (auto)
 22.2236 +qed
 22.2237 +
 22.2238 +declare split_paired_All [simp del] split_paired_Ex [simp del]
 22.2239 +ML_setup {*
 22.2240 +simpset_ref() := simpset() delloop "split_all_tac";
 22.2241 +claset_ref () := claset () delSWrapper "split_all_tac"
 22.2242 +*}
 22.2243 +
 22.2244 +lemma dynamic_mheadsD:   
 22.2245 +"\<lbrakk>emh \<in> mheads G S statT sig;    
 22.2246 +  G,statT \<turnstile> dynC valid_lookup_cls_for (is_static emh);
 22.2247 +  isrtype G statT; wf_prog G
 22.2248 + \<rbrakk> \<Longrightarrow> \<exists>m \<in> dynlookup G statT dynC sig: 
 22.2249 +          is_static m=is_static emh \<and> G\<turnstile>resTy m\<preceq>resTy emh"
 22.2250 +proof - 
 22.2251 +  assume      emh: "emh \<in> mheads G S statT sig"
 22.2252 +  and          wf: "wf_prog G"
 22.2253 +  and   dynC_Prop: "G,statT \<turnstile> dynC valid_lookup_cls_for (is_static emh)"
 22.2254 +  and      istype: "isrtype G statT"
 22.2255 +  from dynC_Prop istype wf 
 22.2256 +  obtain y where
 22.2257 +    dynC: "class G dynC = Some y" 
 22.2258 +    by (auto dest: valid_lookup_cls_is_class)
 22.2259 +  from emh wf show ?thesis
 22.2260 +  proof (cases rule: mheads_cases)
 22.2261 +    case Class_methd
 22.2262 +    fix statC statDeclC sm
 22.2263 +    assume     statC: "statT = ClassT statC"
 22.2264 +    assume            "accmethd G S statC sig = Some sm"
 22.2265 +    then have     sm: "methd G statC sig = Some sm" 
 22.2266 +      by (blast dest: accmethd_SomeD)  
 22.2267 +    assume eq_mheads: "mhead (mthd sm) = mhd emh"
 22.2268 +    from statC 
 22.2269 +    have dynlookup: "dynlookup G statT dynC sig = dynmethd G statC dynC sig"
 22.2270 +      by (simp add: dynlookup_def)
 22.2271 +    from wf statC istype dynC_Prop sm 
 22.2272 +    obtain dm where
 22.2273 +      "dynmethd G statC dynC sig = Some dm"
 22.2274 +      "is_static dm = is_static sm" 
 22.2275 +      "G\<turnstile>resTy dm\<preceq>resTy sm"  
 22.2276 +      by (auto dest!: ws_dynmethd accmethd_SomeD)
 22.2277 +    with dynlookup eq_mheads 
 22.2278 +    show ?thesis 
 22.2279 +      by (cases emh type: *) (auto)
 22.2280 +  next
 22.2281 +    case Iface_methd
 22.2282 +    fix I im
 22.2283 +    assume    statI: "statT = IfaceT I" and
 22.2284 +          eq_mheads: "mthd im = mhd emh" and
 22.2285 +                     "im \<in> accimethds G (pid S) I sig" 
 22.2286 +    then have im: "im \<in> imethds G I sig" 
 22.2287 +      by (blast dest: accimethdsD)
 22.2288 +    with istype statI eq_mheads wf 
 22.2289 +    have not_static_emh: "\<not> is_static emh"
 22.2290 +      by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead)
 22.2291 +    from statI im
 22.2292 +    have dynlookup: "dynlookup G statT dynC sig = methd G dynC sig"
 22.2293 +      by (auto simp add: dynlookup_def dynimethd_def) 
 22.2294 +    from wf dynC_Prop statI istype im not_static_emh 
 22.2295 +    obtain dm where
 22.2296 +      "methd G dynC sig = Some dm"
 22.2297 +      "is_static dm = is_static im" 
 22.2298 +      "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd im)" 
 22.2299 +      by (auto dest: implmt_methd)
 22.2300 +    with dynlookup eq_mheads
 22.2301 +    show ?thesis 
 22.2302 +      by (cases emh type: *) (auto)
 22.2303 +  next
 22.2304 +    case Iface_Object_methd
 22.2305 +    fix I sm
 22.2306 +    assume   statI: "statT = IfaceT I" and
 22.2307 +                sm: "accmethd G S Object sig = Some sm" and 
 22.2308 +         eq_mheads: "mhead (mthd sm) = mhd emh" and
 22.2309 +             nPriv: "accmodi sm \<noteq> Private"
 22.2310 +     show ?thesis 
 22.2311 +     proof (cases "imethds G I sig = {}")
 22.2312 +       case True
 22.2313 +       with statI 
 22.2314 +       have dynlookup: "dynlookup G statT dynC sig = dynmethd G Object dynC sig"
 22.2315 +	 by (simp add: dynlookup_def dynimethd_def)
 22.2316 +       from wf dynC 
 22.2317 +       have subclsObj: "G\<turnstile>dynC \<preceq>\<^sub>C Object"
 22.2318 +	 by (auto intro: subcls_ObjectI)
 22.2319 +       from wf dynC dynC_Prop istype sm subclsObj 
 22.2320 +       obtain dm where
 22.2321 +	 "dynmethd G Object dynC sig = Some dm"
 22.2322 +	 "is_static dm = is_static sm" 
 22.2323 +	 "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd sm)"  
 22.2324 +	 by (auto dest!: ws_dynmethd accmethd_SomeD 
 22.2325 +                  intro: class_Object [OF wf])
 22.2326 +       with dynlookup eq_mheads
 22.2327 +       show ?thesis 
 22.2328 +	 by (cases emh type: *) (auto)
 22.2329 +     next
 22.2330 +       case False
 22.2331 +       with statI
 22.2332 +       have dynlookup: "dynlookup G statT dynC sig = methd G dynC sig"
 22.2333 +	 by (simp add: dynlookup_def dynimethd_def)
 22.2334 +       from istype statI
 22.2335 +       have "is_iface G I"
 22.2336 +	 by auto
 22.2337 +       with wf sm nPriv False 
 22.2338 +       obtain im where
 22.2339 +	      im: "im \<in> imethds G I sig" and
 22.2340 +	 eq_stat: "is_static im = is_static sm" and
 22.2341 +         resProp: "G\<turnstile>resTy (mthd im)\<preceq>resTy (mthd sm)"
 22.2342 +	 by (auto dest: wf_imethds_hiding_objmethdsD accmethd_SomeD)
 22.2343 +       from im wf statI istype eq_stat eq_mheads
 22.2344 +       have not_static_sm: "\<not> is_static emh"
 22.2345 +	 by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead)
 22.2346 +       from im wf dynC_Prop dynC istype statI not_static_sm
 22.2347 +       obtain dm where
 22.2348 +	 "methd G dynC sig = Some dm"
 22.2349 +	 "is_static dm = is_static im" 
 22.2350 +	 "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd im)" 
 22.2351 +	 by (auto dest: implmt_methd)
 22.2352 +       with wf eq_stat resProp dynlookup eq_mheads
 22.2353 +       show ?thesis 
 22.2354 +	 by (cases emh type: *) (auto intro: widen_trans)
 22.2355 +     qed
 22.2356 +  next
 22.2357 +    case Array_Object_methd
 22.2358 +    fix T sm
 22.2359 +    assume statArr: "statT = ArrayT T" and
 22.2360 +                sm: "accmethd G S Object sig = Some sm" and 
 22.2361 +         eq_mheads: "mhead (mthd sm) = mhd emh" 
 22.2362 +    from statArr dynC_Prop wf
 22.2363 +    have dynlookup: "dynlookup G statT dynC sig = methd G Object sig"
 22.2364 +      by (auto simp add: dynlookup_def dynmethd_C_C)
 22.2365 +    with sm eq_mheads sm 
 22.2366 +    show ?thesis 
 22.2367 +      by (cases emh type: *) (auto dest: accmethd_SomeD)
 22.2368 +  qed
 22.2369 +qed
 22.2370 +
 22.2371 +(* Tactical version *)
 22.2372 +(*
 22.2373 +lemma dynamic_mheadsD: "  
 22.2374 + \<lbrakk>emh \<in> mheads G S statT sig; wf_prog G; class G dynC = Some y;  
 22.2375 +   if (\<exists>T. statT=ArrayT T) then dynC=Object else G\<turnstile>Class dynC\<preceq>RefT statT; 
 22.2376 +   isrtype G statT\<rbrakk> \<Longrightarrow>  
 22.2377 +  \<exists>m \<in> dynlookup G statT dynC sig: 
 22.2378 +     static (mthd m)=static (mhd emh) \<and> G\<turnstile>resTy (mthd m)\<preceq>resTy (mhd emh)"
 22.2379 +apply (drule mheadsD)
 22.2380 +apply safe
 22.2381 +       -- reftype statT is a class  
 22.2382 +apply  (case_tac "\<exists>T. ClassT C = ArrayT T")
 22.2383 +apply    (simp)
 22.2384 +
 22.2385 +apply    (clarsimp simp add: dynlookup_def )
 22.2386 +apply    (frule_tac statC="C" and dynC="dynC"  and sig="sig"  
 22.2387 +         in ws_dynmethd)
 22.2388 +apply      assumption+
 22.2389 +apply    (case_tac "emh")  
 22.2390 +apply    (force dest: accmethd_SomeD)
 22.2391 +
 22.2392 +       -- reftype statT is a interface, method defined in interface 
 22.2393 +apply    (clarsimp simp add: dynlookup_def)
 22.2394 +apply    (drule (1) implmt_methd)
 22.2395 +apply      blast
 22.2396 +apply      blast
 22.2397 +apply    (clarify)  
 22.2398 +apply    (unfold dynimethd_def)
 22.2399 +apply    (rule_tac x="cm" in bexI)
 22.2400 +apply      (case_tac "emh")
 22.2401 +apply      force
 22.2402 +
 22.2403 +apply      force
 22.2404 +
 22.2405 +        -- reftype statT is a interface, method defined in Object 
 22.2406 +apply    (simp add: dynlookup_def)
 22.2407 +apply    (simp only: dynimethd_def)
 22.2408 +apply    (case_tac "imethds G I sig = {}")
 22.2409 +apply       simp
 22.2410 +apply       (frule_tac statC="Object" and dynC="dynC"  and sig="sig"  
 22.2411 +             in ws_dynmethd)
 22.2412 +apply          (blast intro: subcls_ObjectI wf_ws_prog) 
 22.2413 +apply          (blast dest: class_Object)
 22.2414 +apply       (case_tac "emh") 
 22.2415 +apply       (force dest: accmethd_SomeD)
 22.2416 +
 22.2417 +apply       simp
 22.2418 +apply       (subgoal_tac "\<exists> im. im \<in> imethds G I sig") 
 22.2419 +prefer 2      apply blast
 22.2420 +apply       clarify
 22.2421 +apply       (frule (1) implmt_methd)
 22.2422 +apply         simp
 22.2423 +apply         blast  
 22.2424 +apply       (clarify dest!: accmethd_SomeD)
 22.2425 +apply       (frule (4) iface_overrides_Object)
 22.2426 +apply       clarify
 22.2427 +apply       (case_tac emh)
 22.2428 +apply       force
 22.2429 +
 22.2430 +        -- reftype statT is a array
 22.2431 +apply    (simp add: dynlookup_def)
 22.2432 +apply    (case_tac emh)
 22.2433 +apply    (force dest: accmethd_SomeD simp add: dynmethd_def)
 22.2434 +done
 22.2435 +*)
 22.2436 +
 22.2437 +(* ### auf ws_class_induct umstellen *)
 22.2438 +lemma methd_declclass:
 22.2439 +"\<lbrakk>class G C = Some c; wf_prog G; methd G C sig = Some m\<rbrakk> 
 22.2440 + \<Longrightarrow> methd G (declclass m) sig = Some m"
 22.2441 +proof -
 22.2442 +  assume asm: "class G C = Some c" "wf_prog G" "methd G C sig = Some m"
 22.2443 +  have "wf_prog G  \<longrightarrow> 
 22.2444 +	   (\<forall> c m. class G C = Some c \<longrightarrow>  methd G C sig = Some m 
 22.2445 +                   \<longrightarrow>  methd G (declclass m) sig = Some m)"      (is "?P G C") 
 22.2446 +  proof (rule class_rec.induct,intro allI impI)
 22.2447 +    fix G C c m
 22.2448 +    assume hyp: "\<forall>c. C \<noteq> Object \<and> ws_prog G \<and> class G C = Some c \<longrightarrow>
 22.2449 +                     ?P G (super c)"
 22.2450 +    assume wf: "wf_prog G" and cls_C: "class G C = Some c" and
 22.2451 +            m: "methd G C sig = Some m"
 22.2452 +    show "methd G (declclass m) sig = Some m"
 22.2453 +    proof (cases "C=Object")
 22.2454 +      case True
 22.2455 +      with wf m show ?thesis by (auto intro: table_of_map_SomeI)
 22.2456 +    next
 22.2457 +      let ?filter="filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)"
 22.2458 +      let ?table = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
 22.2459 +      case False
 22.2460 +      with cls_C wf m
 22.2461 +      have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m "
 22.2462 +	by (simp add: methd_rec)
 22.2463 +      show ?thesis
 22.2464 +      proof (cases "?table sig")
 22.2465 +	case None
 22.2466 +	from this methd_C have "?filter (methd G (super c)) sig = Some m"
 22.2467 +	  by simp
 22.2468 +	moreover
 22.2469 +	from wf cls_C False obtain sup where "class G (super c) = Some sup"
 22.2470 +	  by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
 22.2471 +	moreover note wf False cls_C hyp
 22.2472 +	ultimately show ?thesis by auto
 22.2473 +      next
 22.2474 +	case Some
 22.2475 +	from this methd_C m show ?thesis by auto 
 22.2476 +      qed
 22.2477 +    qed
 22.2478 +  qed	
 22.2479 +  with asm show ?thesis by auto
 22.2480 +qed
 22.2481 +
 22.2482 +lemma dynmethd_declclass:
 22.2483 + "\<lbrakk>dynmethd G statC dynC sig = Some m;
 22.2484 +   wf_prog G; is_class G statC
 22.2485 +  \<rbrakk> \<Longrightarrow> methd G (declclass m) sig = Some m"
 22.2486 +by (auto dest: dynmethd_declC)
 22.2487 +
 22.2488 +lemma dynlookup_declC:
 22.2489 + "\<lbrakk>dynlookup G statT dynC sig = Some m; wf_prog G;
 22.2490 +   is_class G dynC;isrtype G statT
 22.2491 +  \<rbrakk> \<Longrightarrow> G\<turnstile>dynC \<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m)"
 22.2492 +by (cases "statT")
 22.2493 +   (auto simp add: dynlookup_def dynimethd_def 
 22.2494 +             dest: methd_declC dynmethd_declC)
 22.2495 +
 22.2496 +lemma dynlookup_Array_declclassD [simp]:
 22.2497 +"\<lbrakk>dynlookup G (ArrayT T) Object sig = Some dm;wf_prog G\<rbrakk> 
 22.2498 + \<Longrightarrow> declclass dm = Object"
 22.2499 +proof -
 22.2500 +  assume dynL: "dynlookup G (ArrayT T) Object sig = Some dm"
 22.2501 +  assume wf: "wf_prog G"
 22.2502 +  from wf have ws: "ws_prog G" by auto
 22.2503 +  from wf have is_cls_Obj: "is_class G Object" by auto
 22.2504 +  from dynL wf
 22.2505 +  show ?thesis
 22.2506 +    by (auto simp add: dynlookup_def dynmethd_C_C [OF is_cls_Obj ws]
 22.2507 +                 dest: methd_Object_SomeD)
 22.2508 +qed   
 22.2509 +  
 22.2510 +declare split_paired_All [simp del] split_paired_Ex [simp del]
 22.2511 +ML_setup {*
 22.2512 +simpset_ref() := simpset() delloop "split_all_tac";
 22.2513 +claset_ref () := claset () delSWrapper "split_all_tac"
 22.2514 +*}
 22.2515 +
 22.2516 +lemma wt_is_type: "E,dt\<Turnstile>v\<Colon>T \<Longrightarrow>  wf_prog (prg E) \<longrightarrow> 
 22.2517 +  dt=empty_dt \<longrightarrow> (case T of 
 22.2518 +                     Inl T \<Rightarrow> is_type (prg E) T 
 22.2519 +                   | Inr Ts \<Rightarrow> Ball (set Ts) (is_type (prg E)))"
 22.2520 +apply (unfold empty_dt_def)
 22.2521 +apply (erule wt.induct)
 22.2522 +apply (auto split del: split_if_asm simp del: snd_conv 
 22.2523 +            simp add: is_acc_class_def is_acc_type_def)
 22.2524 +apply    (erule typeof_empty_is_type)
 22.2525 +apply   (frule (1) wf_prog_cdecl [THEN wf_cdecl_supD], rotate_tac -1, 
 22.2526 +        force simp del: snd_conv, clarsimp simp add: is_acc_class_def)
 22.2527 +apply  (drule (1) max_spec2mheads [THEN conjunct1, THEN mheadsD])
 22.2528 +apply  (drule_tac [2] accfield_fields) 
 22.2529 +apply  (frule class_Object)
 22.2530 +apply  (auto dest: accmethd_rT_is_type 
 22.2531 +                   imethds_wf_mhead [THEN conjunct1, THEN rT_is_acc_type]
 22.2532 +             dest!:accimethdsD
 22.2533 +             simp del: class_Object
 22.2534 +             simp add: is_acc_type_def
 22.2535 +    )
 22.2536 +done
 22.2537 +declare split_paired_All [simp] split_paired_Ex [simp]
 22.2538 +ML_setup {*
 22.2539 +claset_ref()  := claset() addSbefore ("split_all_tac", split_all_tac);
 22.2540 +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
 22.2541 +*}
 22.2542 +
 22.2543 +lemma ty_expr_is_type: 
 22.2544 +"\<lbrakk>E\<turnstile>e\<Colon>-T; wf_prog (prg E)\<rbrakk> \<Longrightarrow> is_type (prg E) T"
 22.2545 +by (auto dest!: wt_is_type)
 22.2546 +lemma ty_var_is_type: 
 22.2547 +"\<lbrakk>E\<turnstile>v\<Colon>=T; wf_prog (prg E)\<rbrakk> \<Longrightarrow> is_type (prg E) T"
 22.2548 +by (auto dest!: wt_is_type)
 22.2549 +lemma ty_exprs_is_type: 
 22.2550 +"\<lbrakk>E\<turnstile>es\<Colon>\<doteq>Ts; wf_prog (prg E)\<rbrakk> \<Longrightarrow> Ball (set Ts) (is_type (prg E))"
 22.2551 +by (auto dest!: wt_is_type)
 22.2552 +
 22.2553 +
 22.2554 +lemma static_mheadsD: 
 22.2555 + "\<lbrakk> emh \<in> mheads G S t sig; wf_prog G; E\<turnstile>e\<Colon>-RefT t; prg E=G ; 
 22.2556 +   invmode (mhd emh) e \<noteq> IntVir 
 22.2557 +  \<rbrakk> \<Longrightarrow> \<exists>m. (   (\<exists> C. t = ClassT C \<and> accmethd G S C sig = Some m)
 22.2558 +               \<or> (\<forall> C. t \<noteq> ClassT C \<and> accmethd G S Object sig = Some m )) \<and> 
 22.2559 +          declrefT emh = ClassT (declclass m) \<and>  mhead (mthd m) = (mhd emh)"
 22.2560 +apply (subgoal_tac "is_static emh \<or> e = Super")
 22.2561 +defer apply (force simp add: invmode_def)
 22.2562 +apply (frule  ty_expr_is_type)
 22.2563 +apply   simp
 22.2564 +apply (case_tac "is_static emh")
 22.2565 +apply  (frule (1) mheadsD)
 22.2566 +apply  clarsimp
 22.2567 +apply  safe
 22.2568 +apply    blast
 22.2569 +apply   (auto dest!: imethds_wf_mhead
 22.2570 +                     accmethd_SomeD 
 22.2571 +                     accimethdsD
 22.2572 +              simp add: accObjectmheads_def Objectmheads_def)
 22.2573 +
 22.2574 +apply  (erule wt_elim_cases)
 22.2575 +apply  (force simp add: cmheads_def)
 22.2576 +done
 22.2577 +
 22.2578 +lemma wt_MethdI:  
 22.2579 +"\<lbrakk>methd G C sig = Some m; wf_prog G;  
 22.2580 +  class G C = Some c\<rbrakk> \<Longrightarrow>  
 22.2581 + \<exists>T. \<lparr>prg=G,cls=(declclass m),
 22.2582 +      lcl=\<lambda> k. 
 22.2583 +          (case k of
 22.2584 +             EName e 
 22.2585 +             \<Rightarrow> (case e of 
 22.2586 +                   VNam v 
 22.2587 +                   \<Rightarrow> (table_of (lcls (mbody (mthd m)))
 22.2588 +                                ((pars (mthd m))[\<mapsto>](parTs sig))) v
 22.2589 +                 | Res \<Rightarrow> Some (resTy (mthd m)))
 22.2590 +           | This \<Rightarrow> if is_static m then None else Some (Class (declclass m)))
 22.2591 +     \<rparr>\<turnstile> Methd C sig\<Colon>-T \<and> G\<turnstile>T\<preceq>resTy m"
 22.2592 +apply (frule (2) methd_wf_mdecl, clarify)
 22.2593 +apply (force dest!: wf_mdecl_bodyD intro!: wt.Methd)
 22.2594 +done
 22.2595 +
 22.2596 +subsection "accessibility concerns"
 22.2597 +
 22.2598 +lemma mheads_type_accessible:
 22.2599 + "\<lbrakk>emh \<in> mheads G S T sig; wf_prog G\<rbrakk>
 22.2600 + \<Longrightarrow> G\<turnstile>RefT T accessible_in (pid S)"
 22.2601 +by (erule mheads_cases)
 22.2602 +   (auto dest: accmethd_SomeD accessible_from_commonD accimethdsD)
 22.2603 +
 22.2604 +lemma static_to_dynamic_accessible_from:
 22.2605 +"\<lbrakk>G\<turnstile>m of C accessible_from accC;wf_prog G\<rbrakk> 
 22.2606 + \<Longrightarrow> G\<turnstile>m in C dyn_accessible_from accC"
 22.2607 +proof (induct rule: accessible_fromR.induct)
 22.2608 +qed (auto intro: dyn_accessible_fromR.intros 
 22.2609 +                 member_of_to_member_in
 22.2610 +                 static_to_dynamic_overriding)
 22.2611 +
 22.2612 +lemma static_to_dynamic_accessible_from:
 22.2613 + (assumes stat_acc: "G\<turnstile>m of statC accessible_from accC" and
 22.2614 +          subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
 22.2615 +                wf: "wf_prog G"
 22.2616 + ) "G\<turnstile>m in dynC dyn_accessible_from accC"
 22.2617 +proof - 
 22.2618 +  from stat_acc subclseq 
 22.2619 +  show ?thesis (is "?Dyn_accessible m")
 22.2620 +  proof (induct rule: accessible_fromR.induct)
 22.2621 +    case (immediate statC m)
 22.2622 +    then show "?Dyn_accessible m"
 22.2623 +      by (blast intro: dyn_accessible_fromR.immediate
 22.2624 +                       member_inI
 22.2625 +                       permits_acc_inheritance)
 22.2626 +  next
 22.2627 +    case (overriding _ _ m)
 22.2628 +    with wf show "?Dyn_accessible m"
 22.2629 +      by (blast intro: dyn_accessible_fromR.overriding
 22.2630 +                       member_inI
 22.2631 +                       static_to_dynamic_overriding  
 22.2632 +                       rtrancl_trancl_trancl 
 22.2633 +                       static_to_dynamic_accessible_from)
 22.2634 +  qed
 22.2635 +qed
 22.2636 +
 22.2637 +lemma dynmethd_member_in:
 22.2638 + (assumes    m: "dynmethd G statC dynC sig = Some m" and
 22.2639 +   iscls_statC: "is_class G statC" and
 22.2640 +            wf: "wf_prog G"
 22.2641 + ) "G\<turnstile>Methd sig m member_in dynC"
 22.2642 +proof -
 22.2643 +  from m 
 22.2644 +  have subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC"
 22.2645 +    by (auto simp add: dynmethd_def)
 22.2646 +  from subclseq iscls_statC 
 22.2647 +  have iscls_dynC: "is_class G dynC"
 22.2648 +    by (rule subcls_is_class2)
 22.2649 +  from  iscls_dynC iscls_statC wf m
 22.2650 +  have "G\<turnstile>dynC \<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m) \<and>
 22.2651 +        methd G (declclass m) sig = Some m" 
 22.2652 +    by - (drule dynmethd_declC, auto)
 22.2653 +  with wf 
 22.2654 +  show ?thesis
 22.2655 +    by (auto intro: member_inI dest: methd_member_of)
 22.2656 +qed
 22.2657 +
 22.2658 +lemma dynmethd_access_prop:
 22.2659 +  (assumes statM: "methd G statC sig = Some statM" and
 22.2660 +        stat_acc: "G\<turnstile>Methd sig statM of statC accessible_from accC" and
 22.2661 +            dynM: "dynmethd G statC dynC sig = Some dynM" and
 22.2662 +              wf: "wf_prog G" 
 22.2663 +   ) "G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
 22.2664 +proof -
 22.2665 +  from wf have ws: "ws_prog G" ..
 22.2666 +  from dynM 
 22.2667 +  have subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC"
 22.2668 +    by (auto simp add: dynmethd_def)
 22.2669 +  from stat_acc 
 22.2670 +  have is_cls_statC: "is_class G statC"
 22.2671 +    by (auto dest: accessible_from_commonD member_of_is_classD)
 22.2672 +  with subclseq 
 22.2673 +  have is_cls_dynC: "is_class G dynC"
 22.2674 +    by (rule subcls_is_class2)
 22.2675 +  from is_cls_statC statM wf 
 22.2676 +  have member_statC: "G\<turnstile>Methd sig statM member_of statC"
 22.2677 +    by (auto intro: methd_member_of)
 22.2678 +  from stat_acc 
 22.2679 +  have statC_acc: "G\<turnstile>Class statC accessible_in (pid accC)"
 22.2680 +    by (auto dest: accessible_from_commonD)
 22.2681 +  from statM subclseq is_cls_statC ws 
 22.2682 +  show ?thesis
 22.2683 +  proof (cases rule: dynmethd_cases)
 22.2684 +    case Static
 22.2685 +    assume dynmethd: "dynmethd G statC dynC sig = Some statM"
 22.2686 +    with dynM have eq_dynM_statM: "dynM=statM" 
 22.2687 +      by simp
 22.2688 +    with stat_acc subclseq wf 
 22.2689 +    show ?thesis
 22.2690 +      by (auto intro: static_to_dynamic_accessible_from)
 22.2691 +  next
 22.2692 +    case (Overrides newM)
 22.2693 +    assume dynmethd: "dynmethd G statC dynC sig = Some newM"
 22.2694 +    assume override: "G,sig\<turnstile>newM overrides statM"
 22.2695 +    assume      neq: "newM\<noteq>statM"
 22.2696 +    from dynmethd dynM 
 22.2697 +    have eq_dynM_newM: "dynM=newM" 
 22.2698 +      by simp
 22.2699 +    from dynmethd eq_dynM_newM wf is_cls_statC
 22.2700 +    have "G\<turnstile>Methd sig dynM member_in dynC"
 22.2701 +      by (auto intro: dynmethd_member_in)
 22.2702 +    moreover
 22.2703 +    from subclseq
 22.2704 +    have "G\<turnstile>dynC\<prec>\<^sub>C statC"
 22.2705 +    proof (cases rule: subclseq_cases)
 22.2706 +      case Eq
 22.2707 +      assume "dynC=statC"
 22.2708 +      moreover
 22.2709 +      from is_cls_statC obtain c
 22.2710 +	where "class G statC = Some c"
 22.2711 +	by auto
 22.2712 +      moreover 
 22.2713 +      note statM ws dynmethd 
 22.2714 +      ultimately
 22.2715 +      have "newM=statM" 
 22.2716 +	by (auto simp add: dynmethd_C_C)
 22.2717 +      with neq show ?thesis 
 22.2718 +	by (contradiction)
 22.2719 +    next
 22.2720 +      case Subcls show ?thesis .
 22.2721 +    qed 
 22.2722 +    moreover
 22.2723 +    from stat_acc wf 
 22.2724 +    have "G\<turnstile>Methd sig statM in statC dyn_accessible_from accC"
 22.2725 +      by (blast intro: static_to_dynamic_accessible_from)
 22.2726 +    moreover
 22.2727 +    note override eq_dynM_newM
 22.2728 +    ultimately show ?thesis
 22.2729 +      by (cases dynM,cases statM) (auto intro: dyn_accessible_fromR.overriding)
 22.2730 +  qed
 22.2731 +qed
 22.2732 +
 22.2733 +lemma implmt_methd_access:
 22.2734 +(fixes accC::qtname
 22.2735 + assumes iface_methd: "imethds G I sig \<noteq> {}" and
 22.2736 +           implements: "G\<turnstile>dynC\<leadsto>I"  and
 22.2737 +               isif_I: "is_iface G I" and
 22.2738 +                   wf: "wf_prog G" 
 22.2739 + ) "\<exists> dynM. methd G dynC sig = Some dynM \<and> 
 22.2740 +            G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
 22.2741 +proof -
 22.2742 +  from implements 
 22.2743 +  have iscls_dynC: "is_class G dynC" by (rule implmt_is_class)
 22.2744 +  from iface_methd
 22.2745 +  obtain im
 22.2746 +    where "im \<in> imethds G I sig"
 22.2747 +    by auto
 22.2748 +  with wf implements isif_I 
 22.2749 +  obtain dynM 
 22.2750 +    where dynM: "methd G dynC sig = Some dynM" and
 22.2751 +           pub: "accmodi dynM = Public"
 22.2752 +    by (blast dest: implmt_methd)
 22.2753 +  with iscls_dynC wf
 22.2754 +  have "G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
 22.2755 +    by (auto intro!: dyn_accessible_fromR.immediate 
 22.2756 +              intro: methd_member_of member_of_to_member_in
 22.2757 +                     simp add: permits_acc_def)
 22.2758 +  with dynM    
 22.2759 +  show ?thesis
 22.2760 +    by blast
 22.2761 +qed
 22.2762 +
 22.2763 +corollary implmt_dynimethd_access:
 22.2764 +(fixes accC::qtname
 22.2765 + assumes iface_methd: "imethds G I sig \<noteq> {}" and
 22.2766 +           implements: "G\<turnstile>dynC\<leadsto>I"  and
 22.2767 +               isif_I: "is_iface G I" and
 22.2768 +                   wf: "wf_prog G" 
 22.2769 + ) "\<exists> dynM. dynimethd G I dynC sig = Some dynM \<and> 
 22.2770 +            G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
 22.2771 +proof -
 22.2772 +  from iface_methd
 22.2773 +  have "dynimethd G I dynC sig = methd G dynC sig"
 22.2774 +    by (simp add: dynimethd_def)
 22.2775 +  with iface_methd implements isif_I wf 
 22.2776 +  show ?thesis
 22.2777 +    by (simp only:)
 22.2778 +       (blast intro: implmt_methd_access)
 22.2779 +qed
 22.2780 +
 22.2781 +lemma dynlookup_access_prop:
 22.2782 + (assumes emh: "emh \<in> mheads G accC statT sig" and
 22.2783 +         dynM: "dynlookup G statT dynC sig = Some dynM" and
 22.2784 +    dynC_prop: "G,statT \<turnstile> dynC valid_lookup_cls_for is_static emh" and
 22.2785 +    isT_statT: "isrtype G statT" and
 22.2786 +           wf: "wf_prog G"
 22.2787 + ) "G \<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
 22.2788 +proof -
 22.2789 +  from emh wf
 22.2790 +  have statT_acc: "G\<turnstile>RefT statT accessible_in (pid accC)"
 22.2791 +    by (rule mheads_type_accessible)
 22.2792 +  from dynC_prop isT_statT wf
 22.2793 +  have iscls_dynC: "is_class G dynC"
 22.2794 +    by (rule valid_lookup_cls_is_class)
 22.2795 +  from emh dynC_prop isT_statT wf dynM
 22.2796 +  have eq_static: "is_static emh = is_static dynM"
 22.2797 +    by (auto dest: dynamic_mheadsD)
 22.2798 +  from emh wf show ?thesis
 22.2799 +  proof (cases rule: mheads_cases)
 22.2800 +    case (Class_methd statC _ statM)
 22.2801 +    assume statT: "statT = ClassT statC"
 22.2802 +    assume "accmethd G accC statC sig = Some statM"
 22.2803 +    then have    statM: "methd G statC sig = Some statM" and
 22.2804 +              stat_acc: "G\<turnstile>Methd sig statM of statC accessible_from accC"
 22.2805 +      by (auto dest: accmethd_SomeD)
 22.2806 +    from dynM statT
 22.2807 +    have dynM': "dynmethd G statC dynC sig = Some dynM"
 22.2808 +      by (simp add: dynlookup_def) 
 22.2809 +    from statM stat_acc wf dynM'
 22.2810 +    show ?thesis
 22.2811 +      by (auto dest!: dynmethd_access_prop)
 22.2812 +  next
 22.2813 +    case (Iface_methd I im)
 22.2814 +    then have iface_methd: "imethds G I sig \<noteq> {}" and
 22.2815 +                 statT_acc: "G\<turnstile>RefT statT accessible_in (pid accC)" 
 22.2816 +      by (auto dest: accimethdsD)
 22.2817 +    assume   statT: "statT = IfaceT I"
 22.2818 +    assume      im: "im \<in>  accimethds G (pid accC) I sig"
 22.2819 +    assume eq_mhds: "mthd im = mhd emh"
 22.2820 +    from dynM statT
 22.2821 +    have dynM': "dynimethd G I dynC sig = Some dynM"
 22.2822 +      by (simp add: dynlookup_def)
 22.2823 +    from isT_statT statT 
 22.2824 +    have isif_I: "is_iface G I"
 22.2825 +      by simp
 22.2826 +    show ?thesis
 22.2827 +    proof (cases "is_static emh")
 22.2828 +      case False
 22.2829 +      with statT dynC_prop 
 22.2830 +      have widen_dynC: "G\<turnstile>Class dynC \<preceq> RefT statT"
 22.2831 +	by simp
 22.2832 +      from statT widen_dynC
 22.2833 +      have implmnt: "G\<turnstile>dynC\<leadsto>I"
 22.2834 +	by auto    
 22.2835 +      from eq_static False 
 22.2836 +      have not_static_dynM: "\<not> is_static dynM" 
 22.2837 +	by simp
 22.2838 +      from iface_methd implmnt isif_I wf dynM'
 22.2839 +      show ?thesis
 22.2840 +	by - (drule implmt_dynimethd_access, auto)
 22.2841 +    next
 22.2842 +      case True
 22.2843 +      assume "is_static emh"
 22.2844 +      moreover
 22.2845 +      from wf isT_statT statT im 
 22.2846 +      have "\<not> is_static im"
 22.2847 +	by (auto dest: accimethdsD wf_prog_idecl imethds_wf_mhead)
 22.2848 +      moreover note eq_mhds
 22.2849 +      ultimately show ?thesis
 22.2850 +	by (cases emh) auto
 22.2851 +    qed
 22.2852 +  next
 22.2853 +    case (Iface_Object_methd I statM)
 22.2854 +    assume statT: "statT = IfaceT I"
 22.2855 +    assume "accmethd G accC Object sig = Some statM"
 22.2856 +    then have    statM: "methd G Object sig = Some statM" and
 22.2857 +              stat_acc: "G\<turnstile>Methd sig statM of Object accessible_from accC"
 22.2858 +      by (auto dest: accmethd_SomeD)
 22.2859 +    assume not_Private_statM: "accmodi statM \<noteq> Private"
 22.2860 +    assume eq_mhds: "mhead (mthd statM) = mhd emh"
 22.2861 +    from iscls_dynC wf
 22.2862 +    have widen_dynC_Obj: "G\<turnstile>dynC \<preceq>\<^sub>C Object"
 22.2863 +      by (auto intro: subcls_ObjectI)
 22.2864 +    show ?thesis
 22.2865 +    proof (cases "imethds G I sig = {}")
 22.2866 +      case True
 22.2867 +      from dynM statT True
 22.2868 +      have dynM': "dynmethd G Object dynC sig = Some dynM"
 22.2869 +	by (simp add: dynlookup_def dynimethd_def)
 22.2870 +      from statT  
 22.2871 +      have "G\<turnstile>RefT statT \<preceq>Class Object"
 22.2872 +	by auto
 22.2873 +      with statM statT_acc stat_acc widen_dynC_Obj statT isT_statT 
 22.2874 +        wf dynM' eq_static dynC_prop  
 22.2875 +      show ?thesis
 22.2876 +	by - (drule dynmethd_access_prop,force+) 
 22.2877 +    next
 22.2878 +      case False
 22.2879 +      then obtain im where
 22.2880 +	im: "im \<in>  imethds G I sig"
 22.2881 +	by auto
 22.2882 +      have not_static_emh: "\<not> is_static emh"
 22.2883 +      proof -
 22.2884 +	from im statM statT isT_statT wf not_Private_statM 
 22.2885 +	have "is_static im = is_static statM"
 22.2886 +	  by (auto dest: wf_imethds_hiding_objmethdsD)
 22.2887 +	with wf isT_statT statT im 
 22.2888 +	have "\<not> is_static statM"
 22.2889 +	  by (auto dest: wf_prog_idecl imethds_wf_mhead)
 22.2890 +	with eq_mhds
 22.2891 +	show ?thesis  
 22.2892 +	  by (cases emh) auto
 22.2893 +      qed
 22.2894 +      with statT dynC_prop
 22.2895 +      have implmnt: "G\<turnstile>dynC\<leadsto>I"
 22.2896 +	by simp
 22.2897 +      with isT_statT statT
 22.2898 +      have isif_I: "is_iface G I"
 22.2899 +	by simp
 22.2900 +      from dynM statT
 22.2901 +      have dynM': "dynimethd G I dynC sig = Some dynM"
 22.2902 +	by (simp add: dynlookup_def) 
 22.2903 +      from False implmnt isif_I wf dynM'
 22.2904 +      show ?thesis
 22.2905 +	by - (drule implmt_dynimethd_access, auto)
 22.2906 +    qed
 22.2907 +  next
 22.2908 +    case (Array_Object_methd T statM)
 22.2909 +    assume statT: "statT = ArrayT T"
 22.2910 +    assume "accmethd G accC Object sig = Some statM"
 22.2911 +    then have    statM: "methd G Object sig = Some statM" and
 22.2912 +              stat_acc: "G\<turnstile>Methd sig statM of Object accessible_from accC"
 22.2913 +      by (auto dest: accmethd_SomeD)
 22.2914 +    from statT dynC_prop
 22.2915 +    have dynC_Obj: "dynC = Object" 
 22.2916 +      by simp
 22.2917 +    then
 22.2918 +    have widen_dynC_Obj: "G\<turnstile>Class dynC \<preceq> Class Object"
 22.2919 +      by simp
 22.2920 +    from dynM statT    
 22.2921 +    have dynM': "dynmethd G Object dynC sig = Some dynM"
 22.2922 +      by (simp add: dynlookup_def)
 22.2923 +    from statM statT_acc stat_acc dynM' wf widen_dynC_Obj  
 22.2924 +         statT isT_statT  
 22.2925 +    show ?thesis   
 22.2926 +      by - (drule dynmethd_access_prop, simp+) 
 22.2927 +  qed
 22.2928 +qed
 22.2929 +
 22.2930 +lemma dynlookup_access:
 22.2931 + (assumes emh: "emh \<in> mheads G accC statT sig" and
 22.2932 +    dynC_prop: "G,statT \<turnstile> dynC valid_lookup_cls_for (is_static emh) " and
 22.2933 +    isT_statT: "isrtype G statT" and
 22.2934 +           wf: "wf_prog G"
 22.2935 + ) "\<exists> dynM. dynlookup G statT dynC sig = Some dynM \<and> 
 22.2936 +            G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
 22.2937 +proof - 
 22.2938 +  from dynC_prop isT_statT wf
 22.2939 +  have is_cls_dynC: "is_class G dynC"
 22.2940 +    by (auto dest: valid_lookup_cls_is_class)
 22.2941 +  with emh wf dynC_prop isT_statT
 22.2942 +  obtain dynM where 
 22.2943 +    "dynlookup G statT dynC sig = Some dynM"
 22.2944 +    by - (drule dynamic_mheadsD,auto)
 22.2945 +  with  emh dynC_prop isT_statT wf
 22.2946 +  show ?thesis 
 22.2947 +    by (blast intro: dynlookup_access_prop)
 22.2948 +qed
 22.2949 +
 22.2950 +lemma stat_overrides_Package_old: 
 22.2951 +(assumes stat_override: "G \<turnstile> new overrides\<^sub>S old" and 
 22.2952 +          accmodi_new: "accmodi new = Package" and
 22.2953 +                   wf: "wf_prog G "
 22.2954 +) "accmodi old = Package"
 22.2955 +proof -
 22.2956 +  from stat_override wf 
 22.2957 +  have "accmodi old \<le> accmodi new"
 22.2958 +    by (auto dest: wf_prog_stat_overridesD)
 22.2959 +  with stat_override accmodi_new show ?thesis
 22.2960 +    by (cases "accmodi old") (auto dest: no_Private_stat_override 
 22.2961 +                                   dest: acc_modi_le_Dests)
 22.2962 +qed
 22.2963 +
 22.2964 +text {* @{text dyn_accessible_Package} only works with the @{text wf_prog} assumption. 
 22.2965 +Without it. it is easy to leaf the Package!
 22.2966 +*}
 22.2967 +lemma dyn_accessible_Package:
 22.2968 + "\<lbrakk>G \<turnstile> m in C dyn_accessible_from accC; accmodi m = Package;
 22.2969 +   wf_prog G\<rbrakk>
 22.2970 +  \<Longrightarrow> pid accC = pid (declclass m)"
 22.2971 +proof -
 22.2972 +  assume wf: "wf_prog G "
 22.2973 +  assume accessible: "G \<turnstile> m in C dyn_accessible_from accC"
 22.2974 +  then show "accmodi m = Package 
 22.2975 +            \<Longrightarrow> pid accC = pid (declclass m)"
 22.2976 +    (is "?Pack m \<Longrightarrow> ?P m")
 22.2977 +  proof (induct rule: dyn_accessible_fromR.induct)
 22.2978 +    case (immediate C m)
 22.2979 +    assume "G\<turnstile>m member_in C"
 22.2980 +           "G \<turnstile> m in C permits_acc_to accC"
 22.2981 +           "accmodi m = Package"      
 22.2982 +    then show "?P m"
 22.2983 +      by (auto simp add: permits_acc_def)
 22.2984 +  next
 22.2985 +    case (overriding declC C new newm old Sup)
 22.2986 +    assume member_new: "G \<turnstile> new member_in C" and
 22.2987 +                  new: "new = (declC, mdecl newm)" and
 22.2988 +             override: "G \<turnstile> (declC, newm) overrides old" and
 22.2989 +         subcls_C_Sup: "G\<turnstile>C \<prec>\<^sub>C Sup" and
 22.2990 +              acc_old: "G \<turnstile> methdMembr old in Sup dyn_accessible_from accC" and
 22.2991 +                  hyp: "?Pack (methdMembr old) \<Longrightarrow> ?P (methdMembr old)" and
 22.2992 +          accmodi_new: "accmodi new = Package"
 22.2993 +    from override accmodi_new new wf 
 22.2994 +    have accmodi_old: "accmodi old = Package"  
 22.2995 +      by (auto dest: overrides_Package_old)
 22.2996 +    with hyp 
 22.2997 +    have P_sup: "?P (methdMembr old)"
 22.2998 +      by (simp)
 22.2999 +    from wf override new accmodi_old accmodi_new
 22.3000 +    have eq_pid_new_old: "pid (declclass new) = pid (declclass old)"
 22.3001 +      by (auto dest: dyn_override_Package)
 22.3002 +    with eq_pid_new_old P_sup show "?P new"
 22.3003 +      by auto
 22.3004 +  qed
 22.3005 +qed
 22.3006 +
 22.3007 +end
 22.3008 \ No newline at end of file
    23.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    23.2 +++ b/src/HOL/Bali/WellType.thy	Mon Jan 28 17:00:19 2002 +0100
    23.3 @@ -0,0 +1,598 @@
    23.4 +(*  Title:      isabelle/Bali/WellType.thy
    23.5 +    ID:         $Id$
    23.6 +    Author:     David von Oheimb
    23.7 +    Copyright   1997 Technische Universitaet Muenchen
    23.8 +*)
    23.9 +header {* Well-typedness of Java programs *}
   23.10 +
   23.11 +theory WellType = DeclConcepts:
   23.12 +
   23.13 +text {*
   23.14 +improvements over Java Specification 1.0:
   23.15 +\begin{itemize}
   23.16 +\item methods of Object can be called upon references of interface or array type
   23.17 +\end{itemize}
   23.18 +simplifications:
   23.19 +\begin{itemize}
   23.20 +\item the type rules include all static checks on statements and expressions, 
   23.21 +      e.g. definedness of names (of parameters, locals, fields, methods)
   23.22 +\end{itemize}
   23.23 +design issues:
   23.24 +\begin{itemize}
   23.25 +\item unified type judgment for statements, variables, expressions, 
   23.26 +      expression lists
   23.27 +\item statements are typed like expressions with dummy type Void
   23.28 +\item the typing rules take an extra argument that is capable of determining 
   23.29 +  the dynamic type of objects. Therefore, they can be used for both 
   23.30 +  checking static types and determining runtime types in transition semantics.
   23.31 +\end{itemize}
   23.32 +*}
   23.33 +
   23.34 +types	lenv
   23.35 +	= "(lname, ty) table"   (* local variables, including This and Result *)
   23.36 +
   23.37 +record env = 
   23.38 +         prg:: "prog"    (* program *)
   23.39 +         cls:: "qtname"  (* current package and class name *)
   23.40 +         lcl:: "lenv"    (* local environment *)     
   23.41 +  
   23.42 +translations
   23.43 +  "lenv" <= (type) "(lname, ty) table"
   23.44 +  "lenv" <= (type) "lname \<Rightarrow> ty option"
   23.45 +  "env" <= (type) "\<lparr>prg::prog,cls::qtname,lcl::lenv\<rparr>"
   23.46 +  "env" <= (type) "\<lparr>prg::prog,cls::qtname,lcl::lenv,\<dots>::'a\<rparr>"
   23.47 +
   23.48 +
   23.49 +
   23.50 +syntax
   23.51 +  pkg :: "env \<Rightarrow> pname" (* select the current package from an environment *)
   23.52 +translations
   23.53 +  "pkg e" == "pid (cls e)"
   23.54 +
   23.55 +section "Static overloading: maximally specific methods "
   23.56 +
   23.57 +types
   23.58 +  emhead = "ref_ty \<times> mhead"
   23.59 +
   23.60 +(* Some mnemotic selectors for emhead *)
   23.61 +constdefs 
   23.62 +  "declrefT" :: "emhead \<Rightarrow> ref_ty"
   23.63 +  "declrefT \<equiv> fst"
   23.64 +
   23.65 +  "mhd"     :: "emhead \<Rightarrow> mhead"
   23.66 +  "mhd \<equiv> snd"
   23.67 +
   23.68 +lemma declrefT_simp[simp]:"declrefT (r,m) = r"
   23.69 +by (simp add: declrefT_def)
   23.70 +
   23.71 +lemma mhd_simp[simp]:"mhd (r,m) = m"
   23.72 +by (simp add: mhd_def)
   23.73 +
   23.74 +lemma static_mhd_simp[simp]: "static (mhd m) = is_static m"
   23.75 +by (cases m) (simp add: member_is_static_simp mhd_def)
   23.76 +
   23.77 +lemma mhd_resTy_simp [simp]: "resTy (mhd m) = resTy m"
   23.78 +by (cases m) simp
   23.79 +
   23.80 +lemma mhd_is_static_simp [simp]: "is_static (mhd m) = is_static m"
   23.81 +by (cases m) simp
   23.82 +
   23.83 +lemma mhd_accmodi_simp [simp]: "accmodi (mhd m) = accmodi m"
   23.84 +by (cases m) simp
   23.85 +
   23.86 +consts
   23.87 +  cmheads        :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> emhead set"
   23.88 +  Objectmheads   :: "prog \<Rightarrow> qtname           \<Rightarrow> sig \<Rightarrow> emhead set"
   23.89 +  accObjectmheads:: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> emhead set"
   23.90 +  mheads         :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> emhead set"
   23.91 +defs
   23.92 + cmheads_def:
   23.93 +"cmheads G S C 
   23.94 +  \<equiv> \<lambda>sig. (\<lambda>(Cls,mthd). (ClassT Cls,(mhead mthd))) ` o2s (accmethd G S C sig)"
   23.95 +  Objectmheads_def:
   23.96 +"Objectmheads G S  
   23.97 +  \<equiv> \<lambda>sig. (\<lambda>(Cls,mthd). (ClassT Cls,(mhead mthd))) 
   23.98 +    ` o2s (filter_tab (\<lambda>sig m. accmodi m \<noteq> Private) (accmethd G S Object) sig)"
   23.99 +  accObjectmheads_def:
  23.100 +"accObjectmheads G S T
  23.101 +   \<equiv> if G\<turnstile>RefT T accessible_in (pid S)
  23.102 +        then Objectmheads G S
  23.103 +        else \<lambda>sig. {}"
  23.104 +primrec
  23.105 +"mheads G S  NullT     = (\<lambda>sig. {})"
  23.106 +"mheads G S (IfaceT I) = (\<lambda>sig. (\<lambda>(I,h).(IfaceT I,h)) 
  23.107 +                         ` accimethds G (pid S) I sig \<union> 
  23.108 +                           accObjectmheads G S (IfaceT I) sig)"
  23.109 +"mheads G S (ClassT C) = cmheads G S C"
  23.110 +"mheads G S (ArrayT T) = accObjectmheads G S (ArrayT T)"
  23.111 +
  23.112 +
  23.113 +(* more detailed than necessary for type-safety, see below. *)
  23.114 +constdefs
  23.115 +  (* applicable methods, cf. 15.11.2.1 *)
  23.116 +  appl_methds   :: "prog \<Rightarrow> qtname \<Rightarrow>  ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
  23.117 + "appl_methds G S rt \<equiv> \<lambda> sig. 
  23.118 +      {(mh,pTs') |mh pTs'. mh \<in> mheads G S rt \<lparr>name=name sig,parTs=pTs'\<rparr> \<and> 
  23.119 +                           G\<turnstile>(parTs sig)[\<preceq>]pTs'}"
  23.120 +
  23.121 +  (* more specific methods, cf. 15.11.2.2 *)
  23.122 +  more_spec     :: "prog \<Rightarrow> emhead \<times> ty list \<Rightarrow> emhead \<times> ty list \<Rightarrow> bool"
  23.123 + "more_spec G \<equiv> \<lambda>(mh,pTs). \<lambda>(mh',pTs'). G\<turnstile>pTs[\<preceq>]pTs'"
  23.124 +(*more_spec G \<equiv>\<lambda>((d,h),pTs). \<lambda>((d',h'),pTs'). G\<turnstile>RefT d\<preceq>RefT d'\<and>G\<turnstile>pTs[\<preceq>]pTs'*)
  23.125 +
  23.126 +  (* maximally specific methods, cf. 15.11.2.2 *)
  23.127 +   max_spec     :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
  23.128 +
  23.129 + "max_spec G S rt sig \<equiv>{m. m \<in>appl_methds G S rt sig \<and>
  23.130 +		      (\<forall>m'\<in>appl_methds G S rt sig. more_spec G m' m \<longrightarrow> m'=m)}"
  23.131 +(*
  23.132 +rules (* all properties of more_spec, appl_methods and max_spec we actually need
  23.133 +         these can easily be proven from the above definitions *)
  23.134 +
  23.135 +max_spec2mheads "max_spec G S rt (mn, pTs) = insert (mh, pTs') A \<Longrightarrow>
  23.136 +      mh\<in>mheads G S rt (mn, pTs') \<and> G\<turnstile>pTs[\<preceq>]pTs'"
  23.137 +*)
  23.138 +
  23.139 +lemma max_spec2appl_meths: 
  23.140 +  "x \<in> max_spec G S T sig \<Longrightarrow> x \<in> appl_methds G S T sig" 
  23.141 +by (auto simp: max_spec_def)
  23.142 +
  23.143 +lemma appl_methsD: "(mh,pTs')\<in>appl_methds G S T \<lparr>name=mn,parTs=pTs\<rparr> \<Longrightarrow>  
  23.144 +   mh \<in> mheads G S T \<lparr>name=mn,parTs=pTs'\<rparr> \<and> G\<turnstile>pTs[\<preceq>]pTs'"
  23.145 +by (auto simp: appl_methds_def)
  23.146 +
  23.147 +lemma max_spec2mheads: 
  23.148 +"max_spec G S rt \<lparr>name=mn,parTs=pTs\<rparr> = insert (mh, pTs') A 
  23.149 + \<Longrightarrow> mh \<in> mheads G S rt \<lparr>name=mn,parTs=pTs'\<rparr> \<and> G\<turnstile>pTs[\<preceq>]pTs'"
  23.150 +apply (auto dest: equalityD2 subsetD max_spec2appl_meths appl_methsD)
  23.151 +done
  23.152 +
  23.153 +
  23.154 +constdefs
  23.155 +  empty_dt :: "dyn_ty"
  23.156 + "empty_dt \<equiv> \<lambda>a. None"
  23.157 +
  23.158 +  invmode :: "('a::type)member_scheme \<Rightarrow> expr \<Rightarrow> inv_mode"
  23.159 +"invmode m e \<equiv> if static m then Static else if e=Super then SuperM else IntVir"
  23.160 +
  23.161 +lemma invmode_nonstatic [simp]: 
  23.162 +  "invmode \<lparr>access=a,static=False,\<dots>=x\<rparr> (Acc (LVar e)) = IntVir"
  23.163 +apply (unfold invmode_def)
  23.164 +apply (simp (no_asm))
  23.165 +done
  23.166 +
  23.167 +
  23.168 +lemma invmode_Static_eq [simp]: "(invmode m e = Static) = static m"
  23.169 +apply (unfold invmode_def)
  23.170 +apply (simp (no_asm))
  23.171 +done
  23.172 +
  23.173 +
  23.174 +lemma invmode_IntVir_eq: "(invmode m e = IntVir) = (\<not>(static m) \<and> e\<noteq>Super)"
  23.175 +apply (unfold invmode_def)
  23.176 +apply (simp (no_asm))
  23.177 +done
  23.178 +
  23.179 +lemma Null_staticD: 
  23.180 +  "a'=Null \<longrightarrow> (static m) \<Longrightarrow> invmode m e = IntVir \<longrightarrow> a' \<noteq> Null"
  23.181 +apply (clarsimp simp add: invmode_IntVir_eq)
  23.182 +done
  23.183 +
  23.184 +
  23.185 +types tys  =        "ty + ty list"
  23.186 +translations
  23.187 +  "tys"   <= (type) "ty + ty list"
  23.188 +
  23.189 +consts
  23.190 +  wt    :: "(env\<spacespace> \<times> dyn_ty\<spacespace> \<times>  term \<times> tys) set"
  23.191 +(*wt    :: " env \<Rightarrow> dyn_ty \<Rightarrow> (term \<times> tys) set" not feasible because of 
  23.192 +					      \<spacespace>  changing env in Try stmt *)
  23.193 +
  23.194 +syntax
  23.195 +wt      :: "env \<Rightarrow> dyn_ty \<Rightarrow> [term,tys]  \<Rightarrow> bool" ("_,_|=_::_" [51,51,51,51] 50)
  23.196 +wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow>  stmt       \<Rightarrow> bool" ("_,_|=_:<>" [51,51,51   ] 50)
  23.197 +ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_,_|=_:-_" [51,51,51,51] 50)
  23.198 +ty_var  :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_,_|=_:=_" [51,51,51,51] 50)
  23.199 +ty_exprs:: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr list,
  23.200 +	         \<spacespace>        \<spacespace>  ty   list] \<Rightarrow> bool" ("_,_|=_:#_" [51,51,51,51] 50)
  23.201 +
  23.202 +syntax (xsymbols)
  23.203 +wt      :: "env \<Rightarrow> dyn_ty \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>_"  [51,51,51,51] 50)
  23.204 +wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow>  stmt       \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>\<surd>"  [51,51,51   ] 50)
  23.205 +ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>-_" [51,51,51,51] 50)
  23.206 +ty_var  :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>=_" [51,51,51,51] 50)
  23.207 +ty_exprs:: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr list,
  23.208 +		  \<spacespace>        \<spacespace>  ty   list] \<Rightarrow> bool"("_,_\<Turnstile>_\<Colon>\<doteq>_" [51,51,51,51] 50)
  23.209 +
  23.210 +translations
  23.211 +	"E,dt\<Turnstile>t\<Colon>T" == "(E,dt,t,T) \<in> wt"
  23.212 +	"E,dt\<Turnstile>s\<Colon>\<surd>"  == "E,dt\<Turnstile>In1r s\<Colon>Inl (PrimT Void)"
  23.213 +	"E,dt\<Turnstile>e\<Colon>-T" == "E,dt\<Turnstile>In1l e\<Colon>Inl T"
  23.214 +	"E,dt\<Turnstile>e\<Colon>=T" == "E,dt\<Turnstile>In2  e\<Colon>Inl T"
  23.215 +	"E,dt\<Turnstile>e\<Colon>\<doteq>T" == "E,dt\<Turnstile>In3  e\<Colon>Inr T"
  23.216 +
  23.217 +syntax (* for purely static typing *)
  23.218 +  wt_      :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_|-_::_" [51,51,51] 50)
  23.219 +  wt_stmt_ :: "env \<Rightarrow>  stmt       \<Rightarrow> bool" ("_|-_:<>" [51,51   ] 50)
  23.220 +  ty_expr_ :: "env \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_|-_:-_" [51,51,51] 50)
  23.221 +  ty_var_  :: "env \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_|-_:=_" [51,51,51] 50)
  23.222 +  ty_exprs_:: "env \<Rightarrow> [expr list,
  23.223 +		     \<spacespace> ty   list] \<Rightarrow> bool" ("_|-_:#_" [51,51,51] 50)
  23.224 +
  23.225 +syntax (xsymbols)
  23.226 +  wt_      :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>_"  [51,51,51] 50)
  23.227 +  wt_stmt_ ::  "env \<Rightarrow>  stmt       \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<surd>"  [51,51   ] 50)
  23.228 +  ty_expr_ :: "env \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>-_" [51,51,51] 50)
  23.229 +  ty_var_  :: "env \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>=_" [51,51,51] 50)
  23.230 +  ty_exprs_ :: "env \<Rightarrow> [expr list,
  23.231 +		    \<spacespace>  ty   list] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<doteq>_" [51,51,51] 50)
  23.232 +
  23.233 +translations
  23.234 +	"E\<turnstile>t\<Colon> T" == "E,empty_dt\<Turnstile>t\<Colon> T"
  23.235 +	"E\<turnstile>s\<Colon>\<surd>" == "E,empty_dt\<Turnstile>s\<Colon>\<surd>"
  23.236 +	"E\<turnstile>e\<Colon>-T" == "E,empty_dt\<Turnstile>e\<Colon>-T"
  23.237 +	"E\<turnstile>e\<Colon>=T" == "E,empty_dt\<Turnstile>e\<Colon>=T"
  23.238 +	"E\<turnstile>e\<Colon>\<doteq>T" == "E,empty_dt\<Turnstile>e\<Colon>\<doteq>T"
  23.239 +
  23.240 +
  23.241 +inductive wt intros 
  23.242 +
  23.243 +
  23.244 +(* well-typed statements *)
  23.245 +
  23.246 +  Skip:					"E,dt\<Turnstile>Skip\<Colon>\<surd>"
  23.247 +
  23.248 +  Expr:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T\<rbrakk> \<Longrightarrow>
  23.249 +					 E,dt\<Turnstile>Expr e\<Colon>\<surd>"
  23.250 +  (* cf. 14.6 *)
  23.251 +  Lab:  "E,dt\<Turnstile>c\<Colon>\<surd> \<Longrightarrow>                   
  23.252 +                                         E,dt\<Turnstile>l\<bullet> c\<Colon>\<surd>" 
  23.253 +
  23.254 +  Comp:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; 
  23.255 +	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
  23.256 +					 E,dt\<Turnstile>c1;; c2\<Colon>\<surd>"
  23.257 +
  23.258 +  (* cf. 14.8 *)
  23.259 +  If:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
  23.260 +	  E,dt\<Turnstile>c1\<Colon>\<surd>;
  23.261 +	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
  23.262 +					 E,dt\<Turnstile>If(e) c1 Else c2\<Colon>\<surd>"
  23.263 +
  23.264 +  (* cf. 14.10 *)
  23.265 +  Loop:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
  23.266 +	  E,dt\<Turnstile>c\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
  23.267 +					 E,dt\<Turnstile>l\<bullet> While(e) c\<Colon>\<surd>"
  23.268 +  (* cf. 14.13, 14.15, 14.16 *)
  23.269 +  Do:                                   "E,dt\<Turnstile>Do jump\<Colon>\<surd>"
  23.270 +
  23.271 +  (* cf. 14.16 *)
  23.272 +  Throw: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class tn;
  23.273 +	  prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable\<rbrakk> \<Longrightarrow>
  23.274 +					 E,dt\<Turnstile>Throw e\<Colon>\<surd>"
  23.275 +  (* cf. 14.18 *)
  23.276 +  Try:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable;
  23.277 +	  lcl E (VName vn)=None; E \<lparr>lcl := lcl E(VName vn\<mapsto>Class tn)\<rparr>,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk>
  23.278 +          \<Longrightarrow>
  23.279 +					 E,dt\<Turnstile>Try c1 Catch(tn vn) c2\<Colon>\<surd>"
  23.280 +
  23.281 +  (* cf. 14.18 *)
  23.282 +  Fin:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
  23.283 +					 E,dt\<Turnstile>c1 Finally c2\<Colon>\<surd>"
  23.284 +
  23.285 +  Init:	"\<lbrakk>is_class (prg E) C\<rbrakk> \<Longrightarrow>
  23.286 +					 E,dt\<Turnstile>Init C\<Colon>\<surd>"
  23.287 +  (* Init is created on the fly during evaluation (see Eval.thy). The class
  23.288 +   * isn't necessarily accessible from the points Init is called. Therefor
  23.289 +   * we only demand is_class and not is_acc_class here 
  23.290 +   *)
  23.291 +
  23.292 +(* well-typed expressions *)
  23.293 +
  23.294 +  (* cf. 15.8 *)
  23.295 +  NewC:	"\<lbrakk>is_acc_class (prg E) (pkg E) C\<rbrakk> \<Longrightarrow>
  23.296 +					 E,dt\<Turnstile>NewC C\<Colon>-Class C"
  23.297 +  (* cf. 15.9 *)
  23.298 +  NewA:	"\<lbrakk>is_acc_type (prg E) (pkg E) T;
  23.299 +	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
  23.300 +					 E,dt\<Turnstile>New T[i]\<Colon>-T.[]"
  23.301 +
  23.302 +  (* cf. 15.15 *)
  23.303 +  Cast:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T; is_acc_type (prg E) (pkg E) T';
  23.304 +	  prg E\<turnstile>T\<preceq>? T'\<rbrakk> \<Longrightarrow>
  23.305 +					 E,dt\<Turnstile>Cast T' e\<Colon>-T'"
  23.306 +
  23.307 +  (* cf. 15.19.2 *)
  23.308 +  Inst:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT T; is_acc_type (prg E) (pkg E) (RefT T');
  23.309 +	  prg E\<turnstile>RefT T\<preceq>? RefT T'\<rbrakk> \<Longrightarrow>
  23.310 +					 E,dt\<Turnstile>e InstOf T'\<Colon>-PrimT Boolean"
  23.311 +
  23.312 +  (* cf. 15.7.1 *)
  23.313 +  Lit:	"\<lbrakk>typeof dt x = Some T\<rbrakk> \<Longrightarrow>
  23.314 +					 E,dt\<Turnstile>Lit x\<Colon>-T"
  23.315 +
  23.316 +  (* cf. 15.10.2, 15.11.1 *)
  23.317 +  Super: "\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object;
  23.318 +	  class (prg E) C = Some c\<rbrakk> \<Longrightarrow>
  23.319 +					 E,dt\<Turnstile>Super\<Colon>-Class (super c)"
  23.320 +
  23.321 +  (* cf. 15.13.1, 15.10.1, 15.12 *)
  23.322 +  Acc:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T\<rbrakk> \<Longrightarrow>
  23.323 +					 E,dt\<Turnstile>Acc va\<Colon>-T"
  23.324 +
  23.325 +  (* cf. 15.25, 15.25.1 *)
  23.326 +  Ass:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T; va \<noteq> LVar This;
  23.327 +	  E,dt\<Turnstile>v \<Colon>-T';
  23.328 +	  prg E\<turnstile>T'\<preceq>T\<rbrakk> \<Longrightarrow>
  23.329 +					 E,dt\<Turnstile>va:=v\<Colon>-T'"
  23.330 +
  23.331 +  (* cf. 15.24 *)
  23.332 +  Cond:	"\<lbrakk>E,dt\<Turnstile>e0\<Colon>-PrimT Boolean;
  23.333 +	  E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2;
  23.334 +	  prg E\<turnstile>T1\<preceq>T2 \<and> T = T2  \<or>  prg E\<turnstile>T2\<preceq>T1 \<and> T = T1\<rbrakk> \<Longrightarrow>
  23.335 +					 E,dt\<Turnstile>e0 ? e1 : e2\<Colon>-T"
  23.336 +
  23.337 +  (* cf. 15.11.1, 15.11.2, 15.11.3 *)
  23.338 +  Call:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT;
  23.339 +	  E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;
  23.340 +	  max_spec (prg E) (cls E) statT \<lparr>name=mn,parTs=pTs\<rparr> 
  23.341 +            = {((statDeclT,m),pTs')}
  23.342 +         \<rbrakk> \<Longrightarrow>
  23.343 +		   E,dt\<Turnstile>{statT,invmode m e}e\<cdot>mn({pTs'}ps)\<Colon>-(resTy m)"
  23.344 +
  23.345 +  Methd: "\<lbrakk>is_class (prg E) C;
  23.346 +	  methd (prg E) C sig = Some m;
  23.347 +	  E,dt\<Turnstile>Body (declclass m) (stmt (mbody (mthd m)))\<Colon>-T\<rbrakk> \<Longrightarrow>
  23.348 +					 E,dt\<Turnstile>Methd C sig\<Colon>-T"
  23.349 + (* The class C is the dynamic class of the method call (cf. Eval.thy). 
  23.350 +  * It hasn't got to be directly accessible from the current package (pkg E). 
  23.351 +  * Only the static class must be accessible (enshured indirectly by Call). 
  23.352 +  *)
  23.353 +
  23.354 +  Body:	"\<lbrakk>is_class (prg E) D;
  23.355 +	  E,dt\<Turnstile>blk\<Colon>\<surd>;
  23.356 +	  (lcl E) Result = Some T;
  23.357 +          is_type (prg E) T\<rbrakk> \<Longrightarrow>
  23.358 +   					 E,dt\<Turnstile>Body D blk\<Colon>-T"
  23.359 +  (* the class D implementing the method must not directly be accessible 
  23.360 +   * from the current package (pkg E), but can also be indirectly accessible 
  23.361 +   * due to inheritance (enshured in Call)
  23.362 +   * The result type hasn't got to be accessible in Java! (If it is not 
  23.363 +   * accessible you can only assign it to Object) 
  23.364 +   *)
  23.365 +
  23.366 +(* well-typed variables *)
  23.367 +
  23.368 +  (* cf. 15.13.1 *)
  23.369 +  LVar:	"\<lbrakk>lcl E vn = Some T; is_acc_type (prg E) (pkg E) T\<rbrakk> \<Longrightarrow>
  23.370 +					 E,dt\<Turnstile>LVar vn\<Colon>=T"
  23.371 +  (* cf. 15.10.1 *)
  23.372 +  FVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; 
  23.373 +	  accfield (prg E) (cls E) C fn = Some (fd,f)\<rbrakk> \<Longrightarrow>
  23.374 +			                 E,dt\<Turnstile>{fd,static f}e..fn\<Colon>=(type f)"
  23.375 +  (* cf. 15.12 *)
  23.376 +  AVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T.[]; 
  23.377 +	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
  23.378 +					 E,dt\<Turnstile>e.[i]\<Colon>=T"
  23.379 +
  23.380 +
  23.381 +(* well-typed expression lists *)
  23.382 +
  23.383 +  (* cf. 15.11.??? *)
  23.384 +  Nil:					"E,dt\<Turnstile>[]\<Colon>\<doteq>[]"
  23.385 +
  23.386 +  (* cf. 15.11.??? *)
  23.387 +  Cons:	"\<lbrakk>E,dt\<Turnstile>e \<Colon>-T;
  23.388 +	  E,dt\<Turnstile>es\<Colon>\<doteq>Ts\<rbrakk> \<Longrightarrow>
  23.389 +					 E,dt\<Turnstile>e#es\<Colon>\<doteq>T#Ts"
  23.390 +
  23.391 +
  23.392 +declare not_None_eq [simp del] 
  23.393 +declare split_if [split del] split_if_asm [split del]
  23.394 +declare split_paired_All [simp del] split_paired_Ex [simp del]
  23.395 +ML_setup {*
  23.396 +simpset_ref() := simpset() delloop "split_all_tac"
  23.397 +*}
  23.398 +inductive_cases wt_stmt_cases: "E,dt\<Turnstile>c\<Colon>\<surd>"
  23.399 +inductive_cases wt_elim_cases:
  23.400 +	"E,dt\<Turnstile>In2  (LVar vn)               \<Colon>T"
  23.401 +	"E,dt\<Turnstile>In2  ({fd,s}e..fn)           \<Colon>T"
  23.402 +	"E,dt\<Turnstile>In2  (e.[i])                 \<Colon>T"
  23.403 +	"E,dt\<Turnstile>In1l (NewC C)                \<Colon>T"
  23.404 +	"E,dt\<Turnstile>In1l (New T'[i])             \<Colon>T"
  23.405 +	"E,dt\<Turnstile>In1l (Cast T' e)             \<Colon>T"
  23.406 +	"E,dt\<Turnstile>In1l (e InstOf T')           \<Colon>T"
  23.407 +	"E,dt\<Turnstile>In1l (Lit x)                 \<Colon>T"
  23.408 +	"E,dt\<Turnstile>In1l (Super)                 \<Colon>T"
  23.409 +	"E,dt\<Turnstile>In1l (Acc va)                \<Colon>T"
  23.410 +	"E,dt\<Turnstile>In1l (Ass va v)              \<Colon>T"
  23.411 +	"E,dt\<Turnstile>In1l (e0 ? e1 : e2)          \<Colon>T"
  23.412 +	"E,dt\<Turnstile>In1l ({statT,mode}e\<cdot>mn({pT'}p))\<Colon>T"
  23.413 +	"E,dt\<Turnstile>In1l (Methd C sig)           \<Colon>T"
  23.414 +	"E,dt\<Turnstile>In1l (Body D blk)            \<Colon>T"
  23.415 +	"E,dt\<Turnstile>In3  ([])                    \<Colon>Ts"
  23.416 +	"E,dt\<Turnstile>In3  (e#es)                  \<Colon>Ts"
  23.417 +	"E,dt\<Turnstile>In1r  Skip                   \<Colon>x"
  23.418 +	"E,dt\<Turnstile>In1r (Expr e)                \<Colon>x"
  23.419 +	"E,dt\<Turnstile>In1r (c1;; c2)               \<Colon>x"
  23.420 +        "E,dt\<Turnstile>In1r (l\<bullet> c)                  \<Colon>x" 
  23.421 +	"E,dt\<Turnstile>In1r (If(e) c1 Else c2)      \<Colon>x"
  23.422 +	"E,dt\<Turnstile>In1r (l\<bullet> While(e) c)         \<Colon>x"
  23.423 +        "E,dt\<Turnstile>In1r (Do jump)               \<Colon>x"
  23.424 +	"E,dt\<Turnstile>In1r (Throw e)               \<Colon>x"
  23.425 +	"E,dt\<Turnstile>In1r (Try c1 Catch(tn vn) c2)\<Colon>x"
  23.426 +	"E,dt\<Turnstile>In1r (c1 Finally c2)         \<Colon>x"
  23.427 +	"E,dt\<Turnstile>In1r (Init C)                \<Colon>x"
  23.428 +declare not_None_eq [simp] 
  23.429 +declare split_if [split] split_if_asm [split]
  23.430 +declare split_paired_All [simp] split_paired_Ex [simp]
  23.431 +ML_setup {*
  23.432 +simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
  23.433 +*}
  23.434 +
  23.435 +lemma is_acc_class_is_accessible: 
  23.436 +  "is_acc_class G P C \<Longrightarrow> G\<turnstile>(Class C) accessible_in P"
  23.437 +by (auto simp add: is_acc_class_def)
  23.438 +
  23.439 +lemma is_acc_iface_is_iface: "is_acc_iface G P I \<Longrightarrow> is_iface G I"
  23.440 +by (auto simp add: is_acc_iface_def)
  23.441 +
  23.442 +lemma is_acc_iface_is_accessible: 
  23.443 +  "is_acc_iface G P I \<Longrightarrow> G\<turnstile>(Iface I) accessible_in P"
  23.444 +by (auto simp add: is_acc_iface_def)
  23.445 +
  23.446 +lemma is_acc_type_is_type: "is_acc_type G P T \<Longrightarrow> is_type G T"
  23.447 +by (auto simp add: is_acc_type_def)
  23.448 +
  23.449 +lemma is_acc_iface_is_accessible: "is_acc_type G P T \<Longrightarrow> G\<turnstile>T accessible_in P"
  23.450 +by (auto simp add: is_acc_type_def)
  23.451 +
  23.452 +lemma wt_Methd_is_methd: 
  23.453 +  "E\<turnstile>In1l (Methd C sig)\<Colon>T \<Longrightarrow> is_methd (prg E) C sig"
  23.454 +apply (erule_tac wt_elim_cases)
  23.455 +apply clarsimp
  23.456 +apply (erule is_methdI, assumption)
  23.457 +done
  23.458 +
  23.459 +(* Special versions of some typing rules, better suited to pattern match the
  23.460 + * conclusion (no selectors in the conclusion\<dots>)
  23.461 + *)
  23.462 +
  23.463 +lemma wt_Super:
  23.464 +"\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object;
  23.465 +  class (prg E) C = Some c;D=super c\<rbrakk> 
  23.466 + \<Longrightarrow> E,dt\<Turnstile>Super\<Colon>-Class D"
  23.467 +by (auto elim: wt.Super)
  23.468 + 
  23.469 +lemma wt_Call: 
  23.470 +"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT; E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;  
  23.471 +  max_spec (prg E) (cls E) statT \<lparr>name=mn,parTs=pTs\<rparr> 
  23.472 +    = {((statDeclC,m),pTs')};rT=(resTy m);   
  23.473 + mode = invmode m e\<rbrakk> \<Longrightarrow> E,dt\<Turnstile>{statT,mode}e\<cdot>mn({pTs'}ps)\<Colon>-rT"
  23.474 +by (auto elim: wt.Call)
  23.475 +
  23.476 +
  23.477 +
  23.478 +lemma invocationTypeExpr_noClassD: 
  23.479 +"\<lbrakk> E\<turnstile>e\<Colon>-RefT statT\<rbrakk>
  23.480 + \<Longrightarrow> (\<forall> statC. statT \<noteq> ClassT statC) \<longrightarrow> invmode m e \<noteq> SuperM"
  23.481 +proof -
  23.482 +  assume wt: "E\<turnstile>e\<Colon>-RefT statT"
  23.483 +  show ?thesis
  23.484 +  proof (cases "e=Super")
  23.485 +    case True
  23.486 +    with wt obtain "C" where "statT = ClassT C" by (blast elim: wt_elim_cases)
  23.487 +    then show ?thesis by blast
  23.488 +  next 
  23.489 +    case False then show ?thesis 
  23.490 +      by (auto simp add: invmode_def split: split_if_asm)
  23.491 +  qed
  23.492 +qed
  23.493 +
  23.494 +lemma wt_Super: 
  23.495 +"\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object; class (prg E) C = Some c; D=super c\<rbrakk> 
  23.496 +\<Longrightarrow> E,dt\<Turnstile>Super\<Colon>-Class D"
  23.497 +by (auto elim: wt.Super)
  23.498 +
  23.499 +
  23.500 +lemma wt_FVar:	
  23.501 +"\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; accfield (prg E) (cls E) C fn = Some (fd,f);
  23.502 +  sf=static f; fT=(type f)\<rbrakk> \<Longrightarrow> E,dt\<Turnstile>{fd,sf}e..fn\<Colon>=fT"
  23.503 +by (auto elim: wt.FVar)
  23.504 +
  23.505 +lemma wt_init [iff]: "E,dt\<Turnstile>Init C\<Colon>\<surd> = is_class (prg E) C"
  23.506 +by (auto elim: wt_elim_cases intro: "wt.Init")
  23.507 +
  23.508 +declare wt.Skip [iff]
  23.509 +
  23.510 +lemma wt_StatRef: 
  23.511 +  "is_acc_type (prg E) (pkg E) (RefT rt) \<Longrightarrow> E\<turnstile>StatRef rt\<Colon>-RefT rt"
  23.512 +apply (rule wt.Cast)
  23.513 +apply  (rule wt.Lit)
  23.514 +apply   (simp (no_asm))
  23.515 +apply  (simp (no_asm_simp))
  23.516 +apply (rule cast.widen)
  23.517 +apply (simp (no_asm))
  23.518 +done
  23.519 +
  23.520 +lemma wt_Inj_elim: 
  23.521 +  "\<And>E. E,dt\<Turnstile>t\<Colon>U \<Longrightarrow> case t of 
  23.522 +                       In1 ec \<Rightarrow> (case ec of 
  23.523 +                                    Inl e \<Rightarrow> \<exists>T. U=Inl T  
  23.524 +                                  | Inr s \<Rightarrow> U=Inl (PrimT Void))  
  23.525 +                     | In2 e \<Rightarrow> (\<exists>T. U=Inl T) 
  23.526 +                     | In3 e \<Rightarrow> (\<exists>T. U=Inr T)"
  23.527 +apply (erule wt.induct)
  23.528 +apply auto
  23.529 +done
  23.530 +
  23.531 +ML {*
  23.532 +fun wt_fun name inj rhs =
  23.533 +let
  23.534 +  val lhs = "E,dt\<Turnstile>" ^ inj ^ " t\<Colon>U"
  23.535 +  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
  23.536 +	(K [Auto_tac, ALLGOALS (ftac (thm "wt_Inj_elim")) THEN Auto_tac])
  23.537 +  fun is_Inj (Const (inj,_) $ _) = true
  23.538 +    | is_Inj _                   = false
  23.539 +  fun pred (t as (_ $ (Const ("Pair",_) $
  23.540 +     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
  23.541 +       x))) $ _ )) = is_Inj x
  23.542 +in
  23.543 +  make_simproc name lhs pred (thm name)
  23.544 +end
  23.545 +
  23.546 +val wt_expr_proc  = wt_fun "wt_expr_eq"  "In1l" "\<exists>T.  U=Inl T  \<and> E,dt\<Turnstile>t\<Colon>-T"
  23.547 +val wt_var_proc   = wt_fun "wt_var_eq"   "In2"  "\<exists>T.  U=Inl T  \<and> E,dt\<Turnstile>t\<Colon>=T"
  23.548 +val wt_exprs_proc = wt_fun "wt_exprs_eq" "In3"  "\<exists>Ts. U=Inr Ts \<and> E,dt\<Turnstile>t\<Colon>\<doteq>Ts"
  23.549 +val wt_stmt_proc  = wt_fun "wt_stmt_eq"  "In1r" "U=Inl(PrimT Void)\<and>E,dt\<Turnstile>t\<Colon>\<surd>"
  23.550 +*}
  23.551 +
  23.552 +ML {*
  23.553 +Addsimprocs [wt_expr_proc,wt_var_proc,wt_exprs_proc,wt_stmt_proc]
  23.554 +*}
  23.555 +
  23.556 +lemma Inj_eq_lemma [simp]: 
  23.557 +  "(\<forall>T. (\<exists>T'. T = Inj T' \<and> P T') \<longrightarrow> Q T) = (\<forall>T'. P T' \<longrightarrow> Q (Inj T'))"
  23.558 +by auto
  23.559 +
  23.560 +(* unused *)
  23.561 +lemma single_valued_tys_lemma [rule_format (no_asm)]: 
  23.562 +  "\<forall>S T. G\<turnstile>S\<preceq>T \<longrightarrow> G\<turnstile>T\<preceq>S \<longrightarrow> S = T \<Longrightarrow> E,dt\<Turnstile>t\<Colon>T \<Longrightarrow>  
  23.563 +     G = prg E \<longrightarrow> (\<forall>T'. E,dt\<Turnstile>t\<Colon>T' \<longrightarrow> T  = T')"
  23.564 +apply (cases "E", erule wt.induct)
  23.565 +apply (safe del: disjE)
  23.566 +apply (simp_all (no_asm_use) split del: split_if_asm)
  23.567 +apply (safe del: disjE)
  23.568 +(* 17 subgoals *)
  23.569 +apply (tactic {* ALLGOALS (fn i => if i = 9 then EVERY'[thin_tac "?E,dt\<Turnstile>e0\<Colon>-PrimT Boolean", thin_tac "?E,dt\<Turnstile>e1\<Colon>-?T1", thin_tac "?E,dt\<Turnstile>e2\<Colon>-?T2"] i else thin_tac "All ?P" i) *})
  23.570 +(*apply (safe del: disjE elim!: wt_elim_cases)*)
  23.571 +apply (tactic {*ALLGOALS (eresolve_tac (thms "wt_elim_cases"))*})
  23.572 +apply (simp_all (no_asm_use) split del: split_if_asm)
  23.573 +apply (erule_tac [10] V = "All ?P" in thin_rl) (* Call *)
  23.574 +apply ((blast del: equalityCE dest: sym [THEN trans])+)
  23.575 +done
  23.576 +
  23.577 +(* unused *)
  23.578 +lemma single_valued_tys: 
  23.579 +"ws_prog (prg E) \<Longrightarrow> single_valued {(t,T). E,dt\<Turnstile>t\<Colon>T}"
  23.580 +apply (unfold single_valued_def)
  23.581 +apply clarsimp
  23.582 +apply (rule single_valued_tys_lemma)
  23.583 +apply (auto intro!: widen_antisym)
  23.584 +done
  23.585 +
  23.586 +lemma typeof_empty_is_type [rule_format (no_asm)]: 
  23.587 +  "typeof (\<lambda>a. None) v = Some T \<longrightarrow> is_type G T"
  23.588 +apply (rule val.induct)
  23.589 +apply    	auto
  23.590 +done
  23.591 +
  23.592 +(* unused *)
  23.593 +lemma typeof_is_type [rule_format (no_asm)]: 
  23.594 + "(\<forall>a. v \<noteq> Addr a) \<longrightarrow> (\<exists>T. typeof dt v = Some T \<and> is_type G T)"
  23.595 +apply (rule val.induct)
  23.596 +prefer 5 
  23.597 +apply     fast
  23.598 +apply  (simp_all (no_asm))
  23.599 +done
  23.600 +
  23.601 +end
    24.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    24.2 +++ b/src/HOL/Bali/document/root.tex	Mon Jan 28 17:00:19 2002 +0100
    24.3 @@ -0,0 +1,238 @@
    24.4 +
    24.5 +\documentclass[11pt,a4paper]{book}
    24.6 +\usepackage{isabelle,isabellesym}
    24.7 +
    24.8 +% further packages required for unusual symbols (see also isabellesym.sty)
    24.9 +\usepackage{latexsym}
   24.10 +%\usepackage{amssymb}
   24.11 +%\usepackage[english]{babel}
   24.12 +%\usepackage[latin1]{inputenc}
   24.13 +%\usepackage[only,bigsqcap]{stmaryrd}
   24.14 +%\usepackage{wasysym}
   24.15 +%\usepackage{eufrak}
   24.16 +%\usepackage{textcomp}
   24.17 +%\usepackage{marvosym}
   24.18 +
   24.19 +% this should be the last package used
   24.20 +\usepackage{pdfsetup}
   24.21 +
   24.22 +% proper setup for best-style documents
   24.23 +\urlstyle{rm}
   24.24 +\isabellestyle{it}
   24.25 +
   24.26 +\pagestyle{myheadings}
   24.27 +
   24.28 +\addtolength{\hoffset}{-1,5cm}
   24.29 +\addtolength{\textwidth}{4cm}
   24.30 +\addtolength{\voffset}{-2cm}
   24.31 +\addtolength{\textheight}{4cm}
   24.32 +
   24.33 +%subsection instead of section to make the toc readable
   24.34 +\renewcommand{\thesubsection}{\arabic{subsection}}
   24.35 +\renewcommand{\isamarkupheader}[1]{\newpage\markright{Theory~\isabellecontext}\subsection{#1}}
   24.36 +\renewcommand{\isamarkupsection}[1]{\subsubsection{#1}}
   24.37 +
   24.38 +%remove spaces from the isabelle environment (trivlist makes them too large)
   24.39 +\renewenvironment{isabelle}
   24.40 +{\begin{isabellebody}}
   24.41 +{\end{isabellebody}}
   24.42 +
   24.43 +\begin{document}
   24.44 +
   24.45 +\title{Java Source and Bytecode Formalizations in Isabelle: Bali\\
   24.46 +  {\large -- VerifiCard Project Deliverables -- }}
   24.47 +\author{Gerwin Klein \and Tobias Nipkow \and David von Oheimb \and
   24.48 +  Leonor Prensa Nieto \and Norbert Schirmer \and Martin Strecker}
   24.49 +\maketitle
   24.50 +
   24.51 +\tableofcontents
   24.52 +
   24.53 +\parindent 0pt\parskip 0.5ex
   24.54 +\chapter{Overview}
   24.55 +These theories, called Bali,  model and analyse different aspects of the 
   24.56 +JavaCard \textbf{source language}. 
   24.57 +The basis is an abstract model of the JavaCard source language. 
   24.58 +On it, a type system, an operational semantics and an axiomatic semantics 
   24.59 +(Hoare logic) are built. The execution of a wellformed program (with respect to
   24.60 +the type system) according to the operational semantics is proved to be 
   24.61 +typesafe. The axiomatic semantics is proved to be sound and relative complete 
   24.62 +with respect to the operational semantics.
   24.63 +
   24.64 +We have modelled large parts of the original JavaCard source language. It models
   24.65 +features such as:
   24.66 +\begin{itemize}
   24.67 +\item The basic ``primitive types'' of Java 
   24.68 +\item Classes and related concepts 
   24.69 +\item Class fields and methods
   24.70 +\item Instance fields and methods
   24.71 +\item Interfaces and related concepts 
   24.72 +\item Arrays
   24.73 +\item Static initialisation
   24.74 +\item Static overloading of fields and methods
   24.75 +\item Inheritance, overriding and hiding of methods, dynamic binding
   24.76 +\item All cases of abrupt termination
   24.77 +      \begin{itemize}
   24.78 +        \item Exception throwing and handling
   24.79 +        \item \texttt{break}, \texttt{continue} and \texttt{return} 
   24.80 +      \end{itemize}
   24.81 +\item Packages
   24.82 +\item Access Modifiers (\texttt{private}, \texttt{protected}, \texttt{public})
   24.83 +\end{itemize}
   24.84 +
   24.85 +The following features are missing in Bali wrt.{} JavaCard:
   24.86 +\begin{itemize}
   24.87 +\item Some primitive types (\texttt{byte, short})
   24.88 +\item Most numeric operations, syntactic variants of statements
   24.89 +  (\texttt{do}-loop, \texttt{for}-loop)
   24.90 +\item A ``definite assignment'' check
   24.91 +\end{itemize}
   24.92 +
   24.93 +In addition, features are missing that are not part of the JavaCard
   24.94 +language, such as multithreading and garbage collection. No attempt
   24.95 +has been made to model peculiarities of JavaCard such as the applet
   24.96 +firewall or the transaction mechanism.
   24.97 +
   24.98 +
   24.99 +Overview of the theories:
  24.100 +\begin{description}
  24.101 +\item [Basis.thy] 
  24.102 +Some basic definitions and settings not specific to JavaCard but missing in HOL.
  24.103 +
  24.104 +\item [Table.thy]
  24.105 +Definition and some properties of a lookup table to map various names 
  24.106 +(like class names or method names) to some content (like classes or methods).
  24.107 +
  24.108 +\item[Name.thy]
  24.109 +Definition of various names (class names, variable names, package names,...)
  24.110 +
  24.111 +\item[Value.thy]
  24.112 +JavaCard expression values (Boolean, Integer, Addresses,...)
  24.113 +
  24.114 +\item[Type.thy]
  24.115 +JavaCard types. Primitive types (Boolean, Integer,...) and reference types 
  24.116 +(Classes, Interfaces, Arrays,...)
  24.117 +
  24.118 +\item[Term.thy]
  24.119 +JavaCard terms. Variables, expressions and statements.
  24.120 +
  24.121 +\item[Decl.thy]
  24.122 +Class, interface and program declarations. Recursion operators for the
  24.123 +class and the interface hierarchy. 
  24.124 +
  24.125 +\item[TypeRel.thy]
  24.126 +Various relations on types like the subclass-, subinterface-, widening-, 
  24.127 +narrowing- and casting-relation.
  24.128 +
  24.129 +\item[DeclConcepts.thy]
  24.130 +Advanced concepts on the class and interface hierarchy like inheritance, 
  24.131 +overriding, hiding, accessibility of types and members according to the access 
  24.132 +modifiers, method lookup.
  24.133 +
  24.134 +\item[WellType.thy]
  24.135 +Typesystem on the JavaCard term level.
  24.136 +
  24.137 +\item[WellForm.thy]
  24.138 +Typesystem on the JavaCard class, interface and program level.
  24.139 +
  24.140 +\item[State.thy]
  24.141 +The program state (like object store) for the execution of JavaCard.
  24.142 +Abrupt completion (exceptions, break, continue, return) is modelled as flag
  24.143 +inside the state.
  24.144 +
  24.145 +\item[Eval.thy]
  24.146 +Operational (big step) semantics for JavaCard.
  24.147 +
  24.148 +\item[Example.thy]
  24.149 +An concrete example of a JavaCard program to validate the typesystem and the
  24.150 +operational semantics.
  24.151 +
  24.152 +\item[Conform.thy]
  24.153 +Conformance predicate for states. When does an execution state conform to the
  24.154 +static types of the program given by the typesystem.
  24.155 +
  24.156 +\item[TypeSafe.thy]
  24.157 +Typesafety proof of the execution of JavaCard. ''Welltyped programs don't go
  24.158 +wrong'' or more technical: The execution of a welltyped JavaCard program 
  24.159 +preserves the conformance of execution states.
  24.160 +
  24.161 +\item[Evaln.thy]
  24.162 +Copy of the operational semantics given in Eval.thy expanded with an annotation
  24.163 +for the maximal recursive depth. The semantics is not altered. The annotation
  24.164 +is needed for the soundness proof of the axiomatic semantics.
  24.165 +
  24.166 +\item[AxSem.thy]
  24.167 +An axiomatic semantics (Hoare logic) for JavaCard.
  24.168 +
  24.169 +\item[AxSound.thy]
  24.170 +The soundness proof of the axiomatic semantics with respect to the operational
  24.171 +semantics.
  24.172 +
  24.173 +\item[AxCompl.thy]
  24.174 +The proof of (relative) completeness of the axiomatic semantics with respect
  24.175 +to the operational semantics. 
  24.176 +
  24.177 +\item[AxExample.thy]
  24.178 +An concrete example of the axiomatic semantics at work, applied to prove 
  24.179 +some properties of the JavaCard example given in Example.thy.
  24.180 +\end{description}
  24.181 +
  24.182 +% include generated text of all theories
  24.183 +\chapter{Basis.thy}
  24.184 +\input{../generated/Basis.tex}
  24.185 +\chapter{Table.thy}
  24.186 +\input{../generated/Table.tex}    
  24.187 +
  24.188 +\chapter{Name.thy}
  24.189 +\input{../generated/Name.tex}
  24.190 +\chapter{Value.thy}     
  24.191 +\input{../generated/Value.tex}
  24.192 +\chapter{Type.thy}
  24.193 +\input{../generated/Type.tex}      
  24.194 +\chapter{Term.thy}
  24.195 +\input{../generated/Term.tex}     
  24.196 +\chapter{Decl.thy}
  24.197 +\input{../generated/Decl.tex}          
  24.198 +\chapter{TypeRel.thy}
  24.199 +\input{../generated/TypeRel.tex}   
  24.200 +\chapter{DeclConcepts.thy}
  24.201 +\input{../generated/DeclConcepts.tex}  
  24.202 +
  24.203 +\chapter{WellType.thy}
  24.204 +\input{../generated/WellType.tex}
  24.205 +\chapter{WellForm.thy}
  24.206 +\input{../generated/WellForm.tex}
  24.207 +
  24.208 +\chapter{State.thy}
  24.209 +\input{../generated/State.tex}    
  24.210 +\chapter{Eval.thy}
  24.211 +\input{../generated/Eval.tex}          
  24.212 +
  24.213 +\chapter{Example.thy}
  24.214 +\input{../generated/Example.tex}  
  24.215 +
  24.216 +\chapter{Conform.thy}
  24.217 +\input{../generated/Conform.tex}       
  24.218 +\chapter{TypeSafe.thy}
  24.219 +\input{../generated/TypeSafe.tex}
  24.220 +
  24.221 +\chapter{Evaln.thy}
  24.222 +\input{../generated/Evaln.tex}         
  24.223 +\chapter{AxSem.thy}
  24.224 +\input{../generated/AxSem.tex}      
  24.225 +\chapter{AxSound.thy}
  24.226 +\input{../generated/AxSound.tex}    
  24.227 +\chapter{AxCompl.thy}
  24.228 +\input{../generated/AxCompl.tex}    
  24.229 +\chapter{AxExample.thy}
  24.230 +\input{../generated/AxExample.tex}  
  24.231 +
  24.232 +
  24.233 +
  24.234 +
  24.235 +
  24.236 +
  24.237 +
  24.238 +%\bibliographystyle{abbrv}
  24.239 +%\bibliography{root}
  24.240 +
  24.241 +\end{document}