Isabelle/Bali sources;
authorschirmer
Mon, 28 Jan 2002 17:00:19 +0100
changeset 12854 00d4a435777f
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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/AxCompl.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,645 @@
+(*  Title:      isabelle/Bali/AxCompl.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1999 Technische Universitaet Muenchen
+*)
+
+header {*
+Completeness proof for Axiomatic semantics of Java expressions and statements
+*}
+
+theory AxCompl = AxSem:
+
+text {*
+design issues:
+\begin{itemize}
+\item proof structured by Most General Formulas (-> Thomas Kleymann)
+\end{itemize}
+*}
+section "set of not yet initialzed classes"
+
+constdefs
+
+  nyinitcls :: "prog \<Rightarrow> state \<Rightarrow> qtname set"
+ "nyinitcls G s \<equiv> {C. is_class G C \<and> \<not> initd C s}"
+
+lemma nyinitcls_subset_class: "nyinitcls G s \<subseteq> {C. is_class G C}"
+apply (unfold nyinitcls_def)
+apply fast
+done
+
+lemmas finite_nyinitcls [simp] =
+   finite_is_class [THEN nyinitcls_subset_class [THEN finite_subset], standard]
+
+lemma card_nyinitcls_bound: "card (nyinitcls G s) \<le> card {C. is_class G C}"
+apply (rule nyinitcls_subset_class [THEN finite_is_class [THEN card_mono]])
+done
+
+lemma nyinitcls_set_locals_cong [simp]: 
+  "nyinitcls G (x,set_locals l s) = nyinitcls G (x,s)"
+apply (unfold nyinitcls_def)
+apply (simp (no_asm))
+done
+
+lemma nyinitcls_abrupt_cong [simp]: "nyinitcls G (f x, y) = nyinitcls G (x, y)"
+apply (unfold nyinitcls_def)
+apply (simp (no_asm))
+done
+
+lemma nyinitcls_abupd_cong [simp]:"!!s. nyinitcls G (abupd f s) = nyinitcls G s"
+apply (unfold nyinitcls_def)
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (simp (no_asm))
+done
+
+lemma card_nyinitcls_abrupt_congE [elim!]: 
+        "card (nyinitcls G (x, s)) \<le> n \<Longrightarrow> card (nyinitcls G (y, s)) \<le> n"
+apply (unfold nyinitcls_def)
+apply auto
+done
+
+lemma nyinitcls_new_xcpt_var [simp]: 
+"nyinitcls G (new_xcpt_var vn s) = nyinitcls G s"
+apply (unfold nyinitcls_def)
+apply (induct_tac "s")
+apply (simp (no_asm))
+done
+
+lemma nyinitcls_init_lvars [simp]: 
+  "nyinitcls G ((init_lvars G C sig mode a' pvs) s) = nyinitcls G s"
+apply (induct_tac "s")
+apply (simp (no_asm) add: init_lvars_def2 split add: split_if)
+done
+
+lemma nyinitcls_emptyD: "\<lbrakk>nyinitcls G s = {}; is_class G C\<rbrakk> \<Longrightarrow> initd C s"
+apply (unfold nyinitcls_def)
+apply fast
+done
+
+lemma card_Suc_lemma: "\<lbrakk>card (insert a A) \<le> Suc n; a\<notin>A; finite A\<rbrakk> \<Longrightarrow> card A \<le> n"
+apply (rotate_tac 1)
+apply clarsimp
+done
+
+lemma nyinitcls_le_SucD: 
+"\<lbrakk>card (nyinitcls G (x,s)) \<le> Suc n; \<not>inited C (globs s); class G C=Some y\<rbrakk> \<Longrightarrow> 
+  card (nyinitcls G (x,init_class_obj G C s)) \<le> n"
+apply (subgoal_tac 
+        "nyinitcls G (x,s) = insert C (nyinitcls G (x,init_class_obj G C s))")
+apply  clarsimp
+apply  (erule thin_rl)
+apply  (rule card_Suc_lemma [OF _ _ finite_nyinitcls])
+apply   (auto dest!: not_initedD elim!: 
+              simp add: nyinitcls_def inited_def split add: split_if_asm)
+done
+
+ML {* bind_thm("inited_gext'",permute_prems 0 1 (thm "inited_gext")) *}
+
+lemma nyinitcls_gext: "snd s\<le>|snd s' \<Longrightarrow> nyinitcls G s' \<subseteq> nyinitcls G s"
+apply (unfold nyinitcls_def)
+apply (force dest!: inited_gext')
+done
+
+lemma card_nyinitcls_gext: 
+  "\<lbrakk>snd s\<le>|snd s'; card (nyinitcls G s) \<le> n\<rbrakk>\<Longrightarrow> card (nyinitcls G s') \<le> n"
+apply (rule le_trans)
+apply  (rule card_mono)
+apply   (rule finite_nyinitcls)
+apply  (erule nyinitcls_gext)
+apply assumption
+done
+
+
+section "init-le"
+
+constdefs
+  init_le :: "prog \<Rightarrow> nat \<Rightarrow> state \<Rightarrow> bool"            ("_\<turnstile>init\<le>_"  [51,51] 50)
+ "G\<turnstile>init\<le>n \<equiv> \<lambda>s. card (nyinitcls G s) \<le> n"
+  
+lemma init_le_def2 [simp]: "(G\<turnstile>init\<le>n) s = (card (nyinitcls G s)\<le>n)"
+apply (unfold init_le_def)
+apply auto
+done
+
+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}"
+apply (drule spec)
+apply (erule conseq1)
+apply clarsimp
+apply (rule card_nyinitcls_bound)
+done
+
+section "Most General Triples and Formulas"
+
+constdefs
+
+  remember_init_state :: "state assn"                ("\<doteq>")
+  "\<doteq> \<equiv> \<lambda>Y s Z. s = Z"
+
+lemma remember_init_state_def2 [simp]: "\<doteq> Y = op ="
+apply (unfold remember_init_state_def)
+apply (simp (no_asm))
+done
+
+consts
+  
+  MGF ::"[state assn, term, prog] \<Rightarrow> state triple"   ("{_} _\<succ> {_\<rightarrow>}"[3,65,3]62)
+  MGFn::"[nat       , term, prog] \<Rightarrow> state triple" ("{=:_} _\<succ> {_\<rightarrow>}"[3,65,3]62)
+
+defs
+  
+
+  MGF_def:
+  "{P} t\<succ> {G\<rightarrow>} \<equiv> {P} t\<succ> {\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y,s')}"
+
+  MGFn_def:
+  "{=:n} t\<succ> {G\<rightarrow>} \<equiv> {\<doteq> \<and>. G\<turnstile>init\<le>n} t\<succ> {G\<rightarrow>}"
+
+(* unused *)
+lemma MGF_valid: "G,{}\<Turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
+apply (unfold MGF_def)
+apply (force dest!: evaln_eval simp add: ax_valids_def triple_valid_def2)
+done
+
+lemma MGF_res_eq_lemma [simp]: 
+  "(\<forall>Y' Y s. Y = Y' \<and> P s \<longrightarrow> Q s) = (\<forall>s. P s \<longrightarrow> Q s)"
+apply auto
+done
+
+lemma MGFn_def2: 
+"G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>} = G,A\<turnstile>{\<doteq> \<and>. G\<turnstile>init\<le>n} 
+                    t\<succ> {\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y,s')}"
+apply (unfold MGFn_def MGF_def)
+apply fast
+done
+
+lemma MGF_MGFn_iff: "G,A\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} = (\<forall>n. G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>})"
+apply (simp (no_asm_use) add: MGFn_def2 MGF_def)
+apply safe
+apply  (erule_tac [2] All_init_leD)
+apply (erule conseq1)
+apply clarsimp
+done
+
+lemma MGFnD: 
+"G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>} \<Longrightarrow>  
+ G,A\<turnstile>{(\<lambda>Y' s' s. s' = s           \<and> P s) \<and>. G\<turnstile>init\<le>n}  
+ t\<succ>  {(\<lambda>Y' s' s. G\<turnstile>s\<midarrow>t\<succ>\<rightarrow>(Y',s') \<and> P s) \<and>. G\<turnstile>init\<le>n}"
+apply (unfold init_le_def)
+apply (simp (no_asm_use) add: MGFn_def2)
+apply (erule conseq12)
+apply clarsimp
+apply (erule (1) eval_gext [THEN card_nyinitcls_gext])
+done
+lemmas MGFnD' = MGFnD [of _ _ _ _ "\<lambda>x. True"] 
+
+lemma MGFNormalI: "G,A\<turnstile>{Normal \<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow>  
+  G,(A::state triple set)\<turnstile>{\<doteq>::state assn} t\<succ> {G\<rightarrow>}"
+apply (unfold MGF_def)
+apply (rule ax_Normal_cases)
+apply  (erule conseq1)
+apply  clarsimp
+apply (rule ax_derivs.Abrupt [THEN conseq1])
+apply (clarsimp simp add: Let_def)
+done
+
+lemma MGFNormalD: "G,A\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow> G,A\<turnstile>{Normal \<doteq>} t\<succ> {G\<rightarrow>}"
+apply (unfold MGF_def)
+apply (erule conseq1)
+apply clarsimp
+done
+
+lemma MGFn_NormalI: 
+"G,(A::state triple set)\<turnstile>{Normal((\<lambda>Y' s' s. s'=s \<and> normal s) \<and>. G\<turnstile>init\<le>n)}t\<succ> 
+ {\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y,s')} \<Longrightarrow> G,A\<turnstile>{=:n}t\<succ>{G\<rightarrow>}"
+apply (simp (no_asm_use) add: MGFn_def2)
+apply (rule ax_Normal_cases)
+apply  (erule conseq1)
+apply  clarsimp
+apply (rule ax_derivs.Abrupt [THEN conseq1])
+apply (clarsimp simp add: Let_def)
+done
+
+lemma MGFn_free_wt: 
+  "(\<exists>T L C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) 
+    \<longrightarrow> G,(A::state triple set)\<turnstile>{=:n} t\<succ> {G\<rightarrow>} 
+   \<Longrightarrow> G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>}"
+apply (rule MGFn_NormalI)
+apply (rule ax_free_wt)
+apply (auto elim: conseq12 simp add: MGFn_def MGF_def)
+done
+
+
+section "main lemmas"
+
+declare fun_upd_apply [simp del]
+declare splitI2 [rule del] (*prevents ugly renaming of state variables*)
+
+ML_setup {* 
+Delsimprocs [eval_expr_proc, eval_var_proc, eval_exprs_proc, eval_stmt_proc]
+*} (*prevents modifying rhs of MGF*)
+ML {*
+val eval_css = (claset() delrules [thm "eval.Abrupt"] addSIs (thms "eval.intros") 
+                delrules[thm "eval.Expr", thm "eval.Init", thm "eval.Try"] 
+                addIs   [thm "eval.Expr", thm "eval.Init"]
+                addSEs[thm "eval.Try"] delrules[equalityCE],
+                simpset() addsimps [split_paired_all,Let_def]
+ addsimprocs [eval_expr_proc,eval_var_proc,eval_exprs_proc,eval_stmt_proc]);
+val eval_Force_tac = force_tac eval_css;
+
+val wt_prepare_tac = EVERY'[
+    rtac (thm "MGFn_free_wt"),
+    clarsimp_tac (claset() addSEs (thms "wt_elim_cases"), simpset())]
+val compl_prepare_tac = EVERY'[rtac (thm "MGFn_NormalI"), Simp_tac]
+val forw_hyp_tac = EVERY'[etac (thm "MGFnD'" RS thm "conseq12"), Clarsimp_tac]
+val forw_hyp_eval_Force_tac = 
+         EVERY'[TRY o rtac allI, forw_hyp_tac, eval_Force_tac]
+*}
+
+lemma MGFn_Init: "\<forall>m. Suc m\<le>n \<longrightarrow> (\<forall>t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>}) \<Longrightarrow>  
+  G,(A::state triple set)\<turnstile>{=:n} In1r (Init C)\<succ> {G\<rightarrow>}"
+apply (tactic "wt_prepare_tac 1")
+(* requires is_class G C two times for nyinitcls *)
+apply (tactic "compl_prepare_tac 1")
+apply (rule_tac C = "initd C" in ax_cases)
+apply  (rule ax_derivs.Done [THEN conseq1])
+apply  (clarsimp intro!: init_done)
+apply (rule_tac y = n in nat.exhaust, clarsimp)
+apply  (rule ax_impossible [THEN conseq1])
+apply  (force dest!: nyinitcls_emptyD)
+apply clarsimp
+apply (drule_tac x = "nat" in spec)
+apply clarsimp
+apply (rule_tac Q = " (\<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)
+apply   simp
+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)
+prefer 2
+apply   (force elim!: nyinitcls_le_SucD)
+apply  (simp split add: split_if, rule conjI, clarify)
+apply   (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1")
+apply  clarify
+apply  (drule spec)
+apply  (erule MGFnD' [THEN conseq12])
+apply  (tactic "force_tac (claset(), simpset() addsimprocs[eval_stmt_proc]) 1")
+apply (rule allI)
+apply (drule spec)
+apply (erule MGFnD' [THEN conseq12])
+apply clarsimp
+apply (tactic {* pair_tac "sa" 1 *})
+apply (tactic"clarsimp_tac (claset(), simpset() addsimprocs[eval_stmt_proc]) 1")
+apply (rule eval_Init, force+)
+done
+lemmas MGFn_InitD = MGFn_Init [THEN MGFnD, THEN ax_NormalD]
+
+lemma MGFn_Call: 
+"\<lbrakk>\<forall>C sig. G,(A::state triple set)\<turnstile>{=:n} In1l (Methd C sig)\<succ> {G\<rightarrow>};  
+  G,A\<turnstile>{=:n} In1l e\<succ> {G\<rightarrow>}; G,A\<turnstile>{=:n} In3 ps\<succ> {G\<rightarrow>}\<rbrakk> \<Longrightarrow>  
+  G,A\<turnstile>{=:n} In1l ({statT,mode}e\<cdot>mn({pTs'}ps))\<succ> {G\<rightarrow>}"
+apply (tactic "wt_prepare_tac 1") (* required for equating mode = invmode m e *)
+apply (tactic "compl_prepare_tac 1")
+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)
+apply  (erule MGFnD [THEN ax_NormalD])
+apply safe
+apply  (erule_tac V = "All ?P" in thin_rl, tactic "forw_hyp_eval_Force_tac 1")
+apply (drule spec, drule spec)
+apply (erule MGFnD' [THEN conseq12])
+apply (tactic "clarsimp_tac eval_css 1")
+apply (erule (1) eval_Call)
+apply   (rule HOL.refl)
+apply  (simp (no_asm_simp))+
+done
+
+lemma MGF_altern: "G,A\<turnstile>{Normal (\<doteq> \<and>. p)} t\<succ> {G\<rightarrow>} =  
+ 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)} 
+  t\<succ> {\<lambda>Y s Z. (Y,s) = Z}"
+apply (unfold MGF_def)
+apply (auto del: conjI elim!: conseq12)
+apply (case_tac "\<exists>w s. G\<turnstile>Norm sa \<midarrow>t\<succ>\<rightarrow> (w,s) ")
+apply  (fast dest: unique_eval)
+apply clarsimp
+apply (erule thin_rl)
+apply (erule thin_rl)
+apply (drule split_paired_All [THEN subst])
+apply (clarsimp elim!: state_not_single)
+done
+
+
+lemma MGFn_Loop: 
+"\<lbrakk>G,(A::state triple set)\<turnstile>{=:n} In1l expr\<succ> {G\<rightarrow>};G,A\<turnstile>{=:n} In1r stmnt\<succ> {G\<rightarrow>} \<rbrakk> 
+\<Longrightarrow> 
+  G,A\<turnstile>{=:n} In1r (l\<bullet> While(expr) stmnt)\<succ> {G\<rightarrow>}"
+apply (rule MGFn_NormalI, simp)
+apply (rule_tac p2 = "\<lambda>s. card (nyinitcls G s) \<le> n" in 
+          MGF_altern [unfolded MGF_def, THEN iffD2, THEN conseq1])
+prefer 2
+apply  clarsimp
+apply (rule_tac P' = 
+"((\<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) 
+  \<and>. (\<lambda>s. card (nyinitcls G s) \<le> n))" in conseq12)
+prefer 2
+apply  clarsimp
+apply  (tactic "smp_tac 1 1", erule_tac V = "All ?P" in thin_rl)
+apply  (rule_tac [2] P' = " (\<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)
+apply   (force dest!: eval.Loop split add: split_if_asm)
+prefer 2
+apply  (erule MGFnD' [THEN conseq12])
+apply  clarsimp
+apply  (erule_tac V = "card (nyinitcls G s') \<le> n" in thin_rl)
+apply  (tactic "eval_Force_tac 1")
+apply (erule MGFnD' [THEN conseq12] , clarsimp)
+apply (rule conjI, erule exI)
+apply (tactic "clarsimp_tac eval_css 1")
+apply (case_tac "a")
+prefer 2
+apply  (clarsimp)
+apply (clarsimp split add: split_if)
+apply (rule conjI, (tactic {* force_tac (claset() addSDs [thm "eval.Loop"],
+  simpset() addsimps [split_paired_all] addsimprocs [eval_stmt_proc]) 1*})+)
+done
+
+lemma MGFn_lemma [rule_format (no_asm)]: 
+ "\<forall>n C sig. G,(A::state triple set)\<turnstile>{=:n} In1l (Methd C sig)\<succ> {G\<rightarrow>} \<Longrightarrow>  
+  \<forall>t. G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>}"
+apply (rule full_nat_induct)
+apply (rule allI)
+apply (drule_tac x = n in spec)
+apply (drule_tac psi = "All ?P" in asm_rl)
+apply (subgoal_tac "\<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>}")
+apply  (tactic "Clarify_tac 2")
+apply  (induct_tac "t")
+apply    (induct_tac "a")
+apply     fast+
+apply (rule var_expr_stmt.induct)
+(* 28 subgoals *)
+prefer 14 apply fast (* Methd *)
+prefer 13 apply (erule (2) MGFn_Call)
+apply (erule_tac [!] V = "All ?P" in thin_rl) (* assumptions on Methd *)
+apply (erule_tac [24] MGFn_Init)
+prefer 19 apply (erule (1) MGFn_Loop)
+apply (tactic "ALLGOALS compl_prepare_tac")
+
+apply (rule ax_derivs.LVar [THEN conseq1], tactic "eval_Force_tac 1")
+
+apply (rule ax_derivs.FVar)
+apply  (erule MGFn_InitD)
+apply (tactic "forw_hyp_eval_Force_tac 1")
+
+apply (rule ax_derivs.AVar)
+apply  (erule MGFnD [THEN ax_NormalD])
+apply (tactic "forw_hyp_eval_Force_tac 1")
+
+apply (rule ax_derivs.NewC)
+apply (erule MGFn_InitD [THEN conseq2])
+apply (tactic "eval_Force_tac 1")
+
+apply (rule_tac Q = "(\<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)
+apply  (simp add: init_comp_ty_def split add: split_if)
+apply   (rule conjI, clarsimp)
+apply   (erule MGFn_InitD [THEN conseq2])
+apply   (tactic "clarsimp_tac eval_css 1")
+apply  clarsimp
+apply  (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1")
+apply (tactic "forw_hyp_eval_Force_tac 1")
+
+apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast],tactic"eval_Force_tac 1")
+
+apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst],tactic"eval_Force_tac 1")
+apply (rule ax_derivs.Lit [THEN conseq1], tactic "eval_Force_tac 1")
+apply (rule ax_derivs.Super [THEN conseq1], tactic "eval_Force_tac 1")
+apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc],tactic"eval_Force_tac 1")
+
+apply (rule ax_derivs.Ass)
+apply  (erule MGFnD [THEN ax_NormalD])
+apply (tactic "forw_hyp_eval_Force_tac 1")
+
+apply (rule ax_derivs.Cond)
+apply  (erule MGFnD [THEN ax_NormalD])
+apply (rule allI)
+apply (rule ax_Normal_cases)
+prefer 2
+apply  (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
+apply  (tactic "eval_Force_tac 1")
+apply (case_tac "b")
+apply  (simp, tactic "forw_hyp_eval_Force_tac 1")
+apply (simp, tactic "forw_hyp_eval_Force_tac 1")
+
+apply (rule_tac Q = " (\<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)
+ apply (erule MGFn_InitD [THEN conseq2])
+ apply (tactic "eval_Force_tac 1")
+apply (tactic "forw_hyp_tac 1")
+apply (tactic {* clarsimp_tac (eval_css delsimps2 [split_paired_all]) 1 *})
+apply (erule (1) eval.Body)
+
+apply (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1")
+
+apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr],tactic"eval_Force_tac 1")
+
+apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
+apply (tactic "clarsimp_tac eval_css 1")
+
+apply (rule ax_derivs.Comp)
+apply  (erule MGFnD [THEN ax_NormalD])
+apply (tactic "forw_hyp_eval_Force_tac 1")
+
+apply (rule ax_derivs.If)
+apply  (erule MGFnD [THEN ax_NormalD])
+apply (rule allI)
+apply (rule ax_Normal_cases)
+prefer 2
+apply  (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
+apply  (tactic "eval_Force_tac 1")
+apply (case_tac "b")
+apply  (simp, tactic "forw_hyp_eval_Force_tac 1")
+apply (simp, tactic "forw_hyp_eval_Force_tac 1")
+
+apply (rule ax_derivs.Do [THEN conseq1])
+apply (tactic {* force_tac (eval_css addsimps2 [thm "abupd_def2"]) 1 *})
+
+apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
+apply (tactic "clarsimp_tac eval_css 1")
+
+apply (rule_tac Q = " (\<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)
+apply   (tactic "eval_Force_tac 3")
+apply  (tactic "forw_hyp_eval_Force_tac 2")
+apply (erule MGFnD [THEN ax_NormalD, THEN conseq2])
+apply (tactic "clarsimp_tac eval_css 1")
+apply (force elim: sxalloc_gext [THEN card_nyinitcls_gext])
+
+apply (rule_tac Q = " (\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>stmt1\<rightarrow> s') \<and>. G\<turnstile>init\<le>n" in ax_derivs.Fin)
+apply  (tactic "forw_hyp_eval_Force_tac 1")
+apply (rule allI)
+apply (tactic "forw_hyp_tac 1")
+apply (tactic {* pair_tac "sb" 1 *})
+apply (tactic"clarsimp_tac (claset(),simpset() addsimprocs [eval_stmt_proc]) 1")
+apply (drule (1) eval.Fin)
+apply clarsimp
+
+apply (rule ax_derivs.Nil [THEN conseq1], tactic "eval_Force_tac 1")
+
+apply (rule ax_derivs.Cons)
+apply  (erule MGFnD [THEN ax_NormalD])
+apply (tactic "forw_hyp_eval_Force_tac 1")
+done
+
+lemma MGF_asm: "\<forall>C sig. is_methd G C sig \<longrightarrow> G,A\<turnstile>{\<doteq>} In1l (Methd C sig)\<succ> {G\<rightarrow>} \<Longrightarrow>
+  G,(A::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
+apply (simp (no_asm_use) add: MGF_MGFn_iff)
+apply (rule allI)
+apply (rule MGFn_lemma)
+apply (intro strip)
+apply (rule MGFn_free_wt)
+apply (force dest: wt_Methd_is_methd)
+done
+
+declare splitI2 [intro!]
+ML_setup {*
+Addsimprocs [ eval_expr_proc, eval_var_proc, eval_exprs_proc, eval_stmt_proc]
+*}
+
+
+section "nested version"
+
+lemma nesting_lemma' [rule_format (no_asm)]: "[| !!A ts. ts <= A ==> P A ts; 
+  !!A pn. !b:bdy pn. P (insert (mgf_call pn) A) {mgf b} ==> P A {mgf_call pn}; 
+  !!A t. !pn:U. P A {mgf_call pn} ==> P A {mgf t};  
+          finite U; uA = mgf_call`U |] ==>  
+  !A. A <= uA --> n <= card uA --> card A = card uA - n --> (!t. P A {mgf t})"
+proof -
+  assume ax_derivs_asm:    "!!A ts. ts <= A ==> P A ts"
+  assume MGF_nested_Methd: "!!A pn. !b:bdy pn. P (insert (mgf_call pn) A) 
+                                                  {mgf b} ==> P A {mgf_call pn}"
+  assume MGF_asm:          "!!A t. !pn:U. P A {mgf_call pn} ==> P A {mgf t}"
+  assume "finite U" "uA = mgf_call`U"
+  then show ?thesis
+    apply -
+    apply (induct_tac "n")
+    apply  (tactic "ALLGOALS Clarsimp_tac")
+    apply  (tactic "dtac (permute_prems 0 1 card_seteq) 1")
+    apply    simp
+    apply   (erule finite_imageI)
+    apply  (simp add: MGF_asm ax_derivs_asm)
+    apply (rule MGF_asm)
+    apply (rule ballI)
+    apply (case_tac "mgf_call pn : A")
+    apply  (fast intro: ax_derivs_asm)
+    apply (rule MGF_nested_Methd)
+    apply (rule ballI)
+    apply (drule spec, erule impE, erule_tac [2] impE, erule_tac [3] impE, 
+           erule_tac [4] spec)
+    apply   fast
+    apply  (erule Suc_leD)
+    apply (drule finite_subset)
+    apply (erule finite_imageI)
+    apply auto
+    apply arith
+  done
+qed
+
+lemma nesting_lemma [rule_format (no_asm)]: "[| !!A ts. ts <= A ==> P A ts; 
+  !!A pn. !b:bdy pn. P (insert (mgf (f pn)) A) {mgf b} ==> P A {mgf (f pn)}; 
+          !!A t. !pn:U. P A {mgf (f pn)} ==> P A {mgf t}; 
+          finite U |] ==> P {} {mgf t}"
+proof -
+  assume 2: "!!A pn. !b:bdy pn. P (insert (mgf (f pn)) A) {mgf b} ==> P A {mgf (f pn)}"
+  assume 3: "!!A t. !pn:U. P A {mgf (f pn)} ==> P A {mgf t}"
+  assume "!!A ts. ts <= A ==> P A ts" "finite U"
+  then show ?thesis
+    apply -
+    apply (rule_tac mgf = "mgf" in nesting_lemma')
+    apply (erule_tac [2] 2)
+    apply (rule_tac [2] 3)
+    apply (rule_tac [6] le_refl)
+    apply auto
+  done
+qed
+
+lemma MGF_nested_Methd: "\<lbrakk>  
+  G,insert ({Normal \<doteq>} In1l (Methd  C sig) \<succ>{G\<rightarrow>}) A\<turnstile>  
+            {Normal \<doteq>} In1l (body G C sig) \<succ>{G\<rightarrow>}  
+ \<rbrakk> \<Longrightarrow>  G,A\<turnstile>{Normal \<doteq>} In1l (Methd  C sig) \<succ>{G\<rightarrow>}"
+apply (unfold MGF_def)
+apply (rule ax_MethdN)
+apply (erule conseq2)
+apply clarsimp
+apply (erule MethdI)
+done
+
+lemma MGF_deriv: "ws_prog G \<Longrightarrow> G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
+apply (rule MGFNormalI)
+apply (rule_tac mgf = "\<lambda>t. {Normal \<doteq>} t\<succ> {G\<rightarrow>}" and 
+                bdy = "\<lambda> (C,sig) .{In1l (body G C sig) }" and 
+                f = "\<lambda> (C,sig) . In1l (Methd C sig) " in nesting_lemma)
+apply    (erule ax_derivs.asm)
+apply   (clarsimp simp add: split_tupled_all)
+apply   (erule MGF_nested_Methd)
+apply  (erule_tac [2] finite_is_methd)
+apply (rule MGF_asm [THEN MGFNormalD])
+apply clarify
+apply (rule MGFNormalI)
+apply force
+done
+
+
+section "simultaneous version"
+
+lemma MGF_simult_Methd_lemma: "finite ms \<Longrightarrow>  
+  G,A\<union> (\<lambda>(C,sig). {Normal \<doteq>} In1l (Methd  C sig)\<succ> {G\<rightarrow>}) ` ms  
+     |\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} In1l (body G C sig)\<succ> {G\<rightarrow>}) ` ms \<Longrightarrow>  
+  G,A|\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} In1l (Methd  C sig)\<succ> {G\<rightarrow>}) ` ms"
+apply (unfold MGF_def)
+apply (rule ax_derivs.Methd [unfolded mtriples_def])
+apply (erule ax_finite_pointwise)
+prefer 2
+apply  (rule ax_derivs.asm)
+apply  fast
+apply clarsimp
+apply (rule conseq2)
+apply  (erule (1) ax_methods_spec)
+apply clarsimp
+apply (erule eval_Methd)
+done
+
+lemma MGF_simult_Methd: "ws_prog G \<Longrightarrow> 
+   G,({}::state triple set)|\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} In1l (Methd C sig)\<succ> {G\<rightarrow>}) 
+   ` Collect (split (is_methd G)) "
+apply (frule finite_is_methd)
+apply (rule MGF_simult_Methd_lemma)
+apply  assumption
+apply (erule ax_finite_pointwise)
+prefer 2
+apply  (rule ax_derivs.asm)
+apply  blast
+apply clarsimp
+apply (rule MGF_asm [THEN MGFNormalD])
+apply clarify
+apply (rule MGFNormalI)
+apply force
+done
+
+lemma MGF_deriv: "ws_prog G \<Longrightarrow> G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
+apply (rule MGF_asm)
+apply (intro strip)
+apply (rule MGFNormalI)
+apply (rule ax_derivs.weaken)
+apply  (erule MGF_simult_Methd)
+apply force
+done
+
+
+section "corollaries"
+
+lemma MGF_complete: "G,{}\<Turnstile>{P} t\<succ> {Q} \<Longrightarrow> G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow>
+  G,({}::state triple set)\<turnstile>{P::state assn} t\<succ> {Q}"
+apply (rule ax_no_hazard)
+apply (unfold MGF_def)
+apply (erule conseq12)
+apply (simp (no_asm_use) add: ax_valids_def triple_valid_def)
+apply (fast dest!: eval_evaln)
+done
+
+theorem ax_complete: "ws_prog G \<Longrightarrow>  
+  G,{}\<Turnstile>{P::state assn} t\<succ> {Q} \<Longrightarrow> G,({}::state triple set)\<turnstile>{P} t\<succ> {Q}"
+apply (erule MGF_complete)
+apply (erule MGF_deriv)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/AxExample.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,270 @@
+(*  Title:      isabelle/Bali/AxExample.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   2000 Technische Universitaet Muenchen
+*)
+header {* Example of a proof based on the Bali axiomatic semantics *}
+
+theory AxExample = AxSem + Example:
+
+constdefs
+  arr_inv :: "st \<Rightarrow> bool"
+ "arr_inv \<equiv> \<lambda>s. \<exists>obj a T el. globs s (Stat Base) = Some obj \<and>
+                              values obj (Inl (arr, Base)) = Some (Addr a) \<and>
+                              heap s a = Some \<lparr>tag=Arr T 2,values=el\<rparr>"
+
+lemma arr_inv_new_obj: 
+"\<And>a. \<lbrakk>arr_inv s; new_Addr (heap s)=Some a\<rbrakk> \<Longrightarrow> arr_inv (gupd(Inl a\<mapsto>x) s)"
+apply (unfold arr_inv_def)
+apply (force dest!: new_AddrD2)
+done
+
+lemma arr_inv_set_locals [simp]: "arr_inv (set_locals l s) = arr_inv s"
+apply (unfold arr_inv_def)
+apply (simp (no_asm))
+done
+
+lemma arr_inv_gupd_Stat [simp]: 
+  "Base \<noteq> C \<Longrightarrow> arr_inv (gupd(Stat C\<mapsto>obj) s) = arr_inv s"
+apply (unfold arr_inv_def)
+apply (simp (no_asm_simp))
+done
+
+lemma ax_inv_lupd [simp]: "arr_inv (lupd(x\<mapsto>y) s) = arr_inv s"
+apply (unfold arr_inv_def)
+apply (simp (no_asm))
+done
+
+
+declare split_if_asm [split del]
+declare lvar_def [simp]
+
+ML {*
+fun inst1_tac s t = instantiate_tac [(s,t)];
+val ax_tac = REPEAT o rtac allI THEN'
+             resolve_tac(thm "ax_Skip"::thm "ax_StatRef"::thm "ax_MethdN"::
+                         thm "ax_Alloc"::thm "ax_Alloc_Arr"::
+                         thm "ax_SXAlloc_Normal"::
+                         funpow 7 tl (thms "ax_derivs.intros"))
+*}
+
+
+theorem ax_test: "tprg,({}::'a triple set)\<turnstile> 
+  {Normal (\<lambda>Y s Z::'a. heap_free four s \<and> \<not>initd Base s \<and> \<not> initd Ext s)} 
+  .test [Class Base]. {\<lambda>Y s Z. abrupt s = Some (Xcpt (Std IndOutBound))}"
+apply (unfold test_def arr_viewed_from_def)
+apply (tactic "ax_tac 1" (*;;*))
+defer
+apply  (tactic "ax_tac 1" (* Try *))
+defer
+apply    (tactic {* inst1_tac "Q1" 
+                 "\<lambda>Y s Z. arr_inv (snd s) \<and> tprg,s\<turnstile>catch SXcpt NullPointer" *})
+prefer 2
+apply    simp
+apply   (rule_tac P' = "Normal (\<lambda>Y s Z. arr_inv (snd s))" in conseq1)
+prefer 2
+apply    clarsimp
+apply   (rule_tac Q' = "(\<lambda>Y s Z. ?Q Y s Z)\<leftarrow>=False\<down>=\<diamondsuit>" in conseq2)
+prefer 2
+apply    simp
+apply   (tactic "ax_tac 1" (* While *))
+prefer 2
+apply    (rule ax_impossible [THEN conseq1], clarsimp)
+apply   (rule_tac P' = "Normal ?P" in conseq1)
+prefer 2
+apply    clarsimp
+apply   (tactic "ax_tac 1")
+apply   (tactic "ax_tac 1" (* AVar *))
+prefer 2
+apply    (rule ax_subst_Val_allI)
+apply    (tactic {* inst1_tac "P'21" "\<lambda>u a. Normal (?PP a\<leftarrow>?x) u" *})
+apply    (simp del: avar_def2 peek_and_def2)
+apply    (tactic "ax_tac 1")
+apply   (tactic "ax_tac 1")
+      (* just for clarification: *)
+apply   (rule_tac Q' = "Normal (\<lambda>Var:(v, f) u ua. fst (snd (avar tprg (Intg 2) v u)) = Some (Xcpt (Std IndOutBound)))" in conseq2)
+prefer 2
+apply    (clarsimp simp add: split_beta)
+apply   (tactic "ax_tac 1" (* FVar *))
+apply    (tactic "ax_tac 2" (* StatRef *))
+apply   (rule ax_derivs.Done [THEN conseq1])
+apply   (clarsimp simp add: arr_inv_def inited_def in_bounds_def)
+defer
+apply  (rule ax_SXAlloc_catch_SXcpt)
+apply  (rule_tac Q' = "(\<lambda>Y (x, s) Z. x = Some (Xcpt (Std NullPointer)) \<and> arr_inv s) \<and>. heap_free two" in conseq2)
+prefer 2
+apply   (simp add: arr_inv_new_obj)
+apply  (tactic "ax_tac 1") 
+apply  (rule_tac C = "Ext" in ax_Call_known_DynT)
+apply     (unfold DynT_prop_def)
+apply     (simp (no_asm))
+apply    (intro strip)
+apply    (rule_tac P' = "Normal ?P" in conseq1)
+apply     (tactic "ax_tac 1" (* Methd *))
+apply     (rule ax_thin [OF _ empty_subsetI])
+apply     (simp (no_asm) add: body_def2)
+apply     (tactic "ax_tac 1" (* Body *))
+(* apply       (rule_tac [2] ax_derivs.Abrupt) *)
+defer
+apply      (simp (no_asm))
+apply      (tactic "ax_tac 1")
+apply      (tactic "ax_tac 1") (* Ass *)
+prefer 2
+apply       (rule ax_subst_Var_allI)
+apply       (tactic {* inst1_tac "P'27" "\<lambda>a vs l vf. ?PP a vs l vf\<leftarrow>?x \<and>. ?p" *})
+apply       (rule allI)
+apply       (tactic {* simp_tac (simpset() delloop "split_all_tac" delsimps [thm "peek_and_def2"]) 1 *})
+apply       (rule ax_derivs.Abrupt)
+apply      (simp (no_asm))
+apply      (tactic "ax_tac 1" (* FVar *))
+apply       (tactic "ax_tac 2", tactic "ax_tac 2", tactic "ax_tac 2")
+apply      (tactic "ax_tac 1")
+apply     clarsimp
+apply     (tactic {* inst1_tac "R14" "\<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)" *})
+prefer 5
+apply     (rule ax_derivs.Done [THEN conseq1], force)
+apply    force
+apply   (rule ax_subst_Val_allI)
+apply   (tactic {* inst1_tac "P'33" "\<lambda>u a. Normal (?PP a\<leftarrow>?x) u" *})
+apply   (simp (no_asm) del: peek_and_def2)
+apply   (tactic "ax_tac 1")
+prefer 2
+apply   (rule ax_subst_Val_allI)
+apply    (tactic {* inst1_tac "P'4" "\<lambda>aa v. Normal (?QQ aa v\<leftarrow>?y)" *})
+apply    (simp del: peek_and_def2)
+apply    (tactic "ax_tac 1")
+apply   (tactic "ax_tac 1")
+apply  (tactic "ax_tac 1")
+apply  (tactic "ax_tac 1")
+(* end method call *)
+apply (simp (no_asm))
+    (* just for clarification: *)
+apply (rule_tac Q' = "Normal ((\<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> 
+ invocation_declclass tprg IntVir s (the (locals s (VName e))) (ClassT Base)  
+     \<lparr>name = foo, parTs = [Class Base]\<rparr> = Ext)) \<and>. initd Ext \<and>. heap_free two)"
+  in conseq2)
+prefer 2
+apply  clarsimp
+apply (tactic "ax_tac 1")
+apply (tactic "ax_tac 1")
+defer
+apply  (rule ax_subst_Var_allI)
+apply  (tactic {* inst1_tac "P'14" "\<lambda>u vf. Normal (?PP vf \<and>. ?p) u" *})
+apply  (simp (no_asm) del: split_paired_All peek_and_def2)
+apply  (tactic "ax_tac 1" (* NewC *))
+apply  (tactic "ax_tac 1" (* ax_Alloc *))
+     (* just for clarification: *)
+apply  (rule_tac Q' = "Normal ((\<lambda>Y s Z. arr_inv (store s) \<and> vf=lvar (VName e) (store s)) \<and>. heap_free tree \<and>. initd Ext)" in conseq2)
+prefer 2
+apply   (simp add: invocation_declclass_def dynmethd_def)
+apply   (unfold dynlookup_def)
+apply   (simp add: dynmethd_Ext_foo)
+apply   (force elim!: arr_inv_new_obj atleast_free_SucD atleast_free_weaken)
+     (* begin init *)
+apply  (rule ax_InitS)
+apply     force
+apply    (simp (no_asm))
+apply   (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *})
+apply   (rule ax_Init_Skip_lemma)
+apply  (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *})
+apply  (rule ax_InitS [THEN conseq1] (* init Base *))
+apply      force
+apply     (simp (no_asm))
+apply    (unfold arr_viewed_from_def)
+apply    (rule allI)
+apply    (rule_tac P' = "Normal ?P" in conseq1)
+apply     (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *})
+apply     (tactic "ax_tac 1")
+apply     (tactic "ax_tac 1")
+apply     (rule_tac [2] ax_subst_Var_allI)
+apply      (tactic {* inst1_tac "P'29" "\<lambda>vf l vfa. Normal (?P vf l vfa)" *})
+apply     (tactic {* simp_tac (simpset() delloop "split_all_tac" delsimps [split_paired_All, thm "peek_and_def2"]) 2 *})
+apply      (tactic "ax_tac 2" (* NewA *))
+apply       (tactic "ax_tac 3" (* ax_Alloc_Arr *))
+apply       (tactic "ax_tac 3")
+apply      (tactic {* inst1_tac "P" "\<lambda>vf l vfa. Normal (?P vf l vfa\<leftarrow>\<diamondsuit>)" *})
+apply      (tactic {* simp_tac (simpset() delloop "split_all_tac") 2 *})
+apply      (tactic "ax_tac 2")
+apply     (tactic "ax_tac 1" (* FVar *))
+apply      (tactic "ax_tac 2" (* StatRef *))
+apply     (rule ax_derivs.Done [THEN conseq1])
+apply     (tactic {* inst1_tac "Q22" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf=lvar (VName e) (snd s)) \<and>. heap_free four \<and>. initd Base \<and>. initd Ext)" *})
+apply     (clarsimp split del: split_if)
+apply     (frule atleast_free_weaken [THEN atleast_free_weaken])
+apply     (drule initedD)
+apply     (clarsimp elim!: atleast_free_SucD simp add: arr_inv_def)
+apply    force
+apply   (tactic {* simp_tac (simpset() delloop "split_all_tac") 1 *})
+apply   (rule ax_triv_Init_Object [THEN peek_and_forget2, THEN conseq1])
+apply     (rule wf_tprg)
+apply    clarsimp
+apply   (tactic {* inst1_tac "P22" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf = lvar (VName e) (snd s)) \<and>. heap_free four \<and>. initd Ext)" *})
+apply   clarsimp
+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)" *})
+apply  clarsimp
+     (* end init *)
+apply (rule conseq1)
+apply (tactic "ax_tac 1")
+apply clarsimp
+done
+
+(*
+while (true) {
+  if (i) {throw xcpt;}
+  else i=j
+}
+*)
+lemma Loop_Xcpt_benchmark: 
+ "Q = (\<lambda>Y (x,s) Z. x \<noteq> None \<longrightarrow> the_Bool (the (locals s i))) \<Longrightarrow>  
+  G,({}::'a triple set)\<turnstile>{Normal (\<lambda>Y s Z::'a. True)}  
+  .lab1\<bullet> While(Lit (Bool True)) (If(Acc (LVar i)) (Throw (Acc (LVar xcpt))) Else
+        (Expr (Ass (LVar i) (Acc (LVar j))))). {Q}"
+apply (rule_tac P' = "Q" and Q' = "Q\<leftarrow>=False\<down>=\<diamondsuit>" in conseq12)
+apply  safe
+apply  (tactic "ax_tac 1" (* Loop *))
+apply   (rule ax_Normal_cases)
+prefer 2
+apply    (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
+apply   (rule conseq1)
+apply    (tactic "ax_tac 1")
+apply   clarsimp
+prefer 2
+apply  clarsimp
+apply (tactic "ax_tac 1" (* If *))
+apply  (tactic 
+  {* inst1_tac "P'21" "Normal (\<lambda>s.. (\<lambda>Y s Z. True)\<down>=Val (the (locals s i)))" *})
+apply  (tactic "ax_tac 1")
+apply  (rule conseq1)
+apply   (tactic "ax_tac 1")
+apply  clarsimp
+apply (rule allI)
+apply (rule ax_escape)
+apply auto
+apply  (rule conseq1)
+apply   (tactic "ax_tac 1" (* Throw *))
+apply   (tactic "ax_tac 1")
+apply   (tactic "ax_tac 1")
+apply  clarsimp
+apply (rule_tac Q' = "Normal (\<lambda>Y s Z. True)" in conseq2)
+prefer 2
+apply  clarsimp
+apply (rule conseq1)
+apply  (tactic "ax_tac 1")
+apply  (tactic "ax_tac 1")
+prefer 2
+apply   (rule ax_subst_Var_allI)
+apply   (tactic {* inst1_tac "P'29" "\<lambda>b Y ba Z vf. \<lambda>Y (x,s) Z. x=None \<and> snd vf = snd (lvar i s)" *})
+apply   (rule allI)
+apply   (rule_tac P' = "Normal ?P" in conseq1)
+prefer 2
+apply    clarsimp
+apply   (tactic "ax_tac 1")
+apply   (rule conseq1)
+apply    (tactic "ax_tac 1")
+apply   clarsimp
+apply  (tactic "ax_tac 1")
+apply clarsimp
+done
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/AxSem.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,1191 @@
+(*  Title:      isabelle/Bali/AxSem.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1998 Technische Universitaet Muenchen
+*)
+
+header {* Axiomatic semantics of Java expressions and statements 
+          (see also Eval.thy)
+        *}
+
+theory AxSem = Evaln + TypeSafe:
+
+text {*
+design issues:
+\begin{itemize}
+\item a strong version of validity for triples with premises, namely one that 
+      takes the recursive depth needed to complete execution, enables 
+      correctness proof
+\item auxiliary variables are handled first-class (-> Thomas Kleymann)
+\item expressions not flattened to elementary assignments (as usual for 
+      axiomatic semantics) but treated first-class => explicit result value 
+      handling
+\item intermediate values not on triple, but on assertion level 
+      (with result entry)
+\item multiple results with semantical substitution mechnism not requiring a 
+      stack 
+\item because of dynamic method binding, terms need to be dependent on state.
+  this is also useful for conditional expressions and statements
+\item result values in triples exactly as in eval relation (also for xcpt 
+      states)
+\item validity: additional assumption of state conformance and well-typedness,
+  which is required for soundness and thus rule hazard required of completeness
+\end{itemize}
+
+restrictions:
+\begin{itemize}
+\item all triples in a derivation are of the same type (due to weak 
+      polymorphism)
+\end{itemize}
+*}
+
+
+
+types  res = vals (* result entry *)
+syntax
+  Val  :: "val      \<Rightarrow> res"
+  Var  :: "var      \<Rightarrow> res"
+  Vals :: "val list \<Rightarrow> res"
+translations
+  "Val  x"     => "(In1 x)"
+  "Var  x"     => "(In2 x)"
+  "Vals x"     => "(In3 x)"
+
+syntax
+  "Val_"    :: "[pttrn] => pttrn"     ("Val:_"  [951] 950)
+  "Var_"    :: "[pttrn] => pttrn"     ("Var:_"  [951] 950)
+  "Vals_"   :: "[pttrn] => pttrn"     ("Vals:_" [951] 950)
+
+translations
+  "\<lambda>Val:v . b"  == "(\<lambda>v. b) \<circ> the_In1"
+  "\<lambda>Var:v . b"  == "(\<lambda>v. b) \<circ> the_In2"
+  "\<lambda>Vals:v. b"  == "(\<lambda>v. b) \<circ> the_In3"
+
+  (* relation on result values, state and auxiliary variables *)
+types 'a assn   =        "res \<Rightarrow> state \<Rightarrow> 'a \<Rightarrow> bool"
+translations
+      "res"    <= (type) "AxSem.res"
+      "a assn" <= (type) "vals \<Rightarrow> state \<Rightarrow> a \<Rightarrow> bool"
+
+constdefs
+  assn_imp   :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> bool"             (infixr "\<Rightarrow>" 25)
+ "P \<Rightarrow> Q \<equiv> \<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z"
+  
+lemma assn_imp_def2 [iff]: "(P \<Rightarrow> Q) = (\<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z)"
+apply (unfold assn_imp_def)
+apply (rule HOL.refl)
+done
+
+
+section "assertion transformers"
+
+subsection "peek-and"
+
+constdefs
+  peek_and   :: "'a assn \<Rightarrow> (state \<Rightarrow>  bool) \<Rightarrow> 'a assn" (infixl "\<and>." 13)
+ "P \<and>. p \<equiv> \<lambda>Y s Z. P Y s Z \<and> p s"
+
+lemma peek_and_def2 [simp]: "peek_and P p Y s = (\<lambda>Z. (P Y s Z \<and> p s))"
+apply (unfold peek_and_def)
+apply (simp (no_asm))
+done
+
+lemma peek_and_Not [simp]: "(P \<and>. (\<lambda>s. \<not> f s)) = (P \<and>. Not \<circ> f)"
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+lemma peek_and_and [simp]: "peek_and (peek_and P p) p = peek_and P p"
+apply (unfold peek_and_def)
+apply (simp (no_asm))
+done
+
+lemma peek_and_commut: "(P \<and>. p \<and>. q) = (P \<and>. q \<and>. p)"
+apply (rule ext)
+apply (rule ext)
+apply (rule ext)
+apply auto
+done
+
+syntax
+  Normal     :: "'a assn \<Rightarrow> 'a assn"
+translations
+  "Normal P" == "P \<and>. normal"
+
+lemma peek_and_Normal [simp]: "peek_and (Normal P) p = Normal (peek_and P p)"
+apply (rule ext)
+apply (rule ext)
+apply (rule ext)
+apply auto
+done
+
+subsection "assn-supd"
+
+constdefs
+  assn_supd  :: "'a assn \<Rightarrow> (state \<Rightarrow> state) \<Rightarrow> 'a assn" (infixl ";." 13)
+ "P ;. f \<equiv> \<lambda>Y s' Z. \<exists>s. P Y s Z \<and> s' = f s"
+
+lemma assn_supd_def2 [simp]: "assn_supd P f Y s' Z = (\<exists>s. P Y s Z \<and> s' = f s)"
+apply (unfold assn_supd_def)
+apply (simp (no_asm))
+done
+
+subsection "supd-assn"
+
+constdefs
+  supd_assn  :: "(state \<Rightarrow> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr ".;" 13)
+ "f .; P \<equiv> \<lambda>Y s. P Y (f s)"
+
+
+lemma supd_assn_def2 [simp]: "(f .; P) Y s = P Y (f s)"
+apply (unfold supd_assn_def)
+apply (simp (no_asm))
+done
+
+lemma supd_assn_supdD [elim]: "((f .; Q) ;. f) Y s Z \<Longrightarrow> Q Y s Z"
+apply auto
+done
+
+lemma supd_assn_supdI [elim]: "Q Y s Z \<Longrightarrow> (f .; (Q ;. f)) Y s Z"
+apply (auto simp del: split_paired_Ex)
+done
+
+subsection "subst-res"
+
+constdefs
+  subst_res   :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn"              ("_\<leftarrow>_"  [60,61] 60)
+ "P\<leftarrow>w \<equiv> \<lambda>Y. P w"
+
+lemma subst_res_def2 [simp]: "(P\<leftarrow>w) Y = P w"
+apply (unfold subst_res_def)
+apply (simp (no_asm))
+done
+
+lemma subst_subst_res [simp]: "P\<leftarrow>w\<leftarrow>v = P\<leftarrow>w"
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+lemma peek_and_subst_res [simp]: "(P \<and>. p)\<leftarrow>w = (P\<leftarrow>w \<and>. p)"
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+(*###Do not work for some strange (unification?) reason
+lemma subst_res_Val_beta [simp]: "(\<lambda>Y. P (the_In1 Y))\<leftarrow>Val v = (\<lambda>Y. P v)"
+apply (rule ext)
+by simp
+
+lemma subst_res_Var_beta [simp]: "(\<lambda>Y. P (the_In2 Y))\<leftarrow>Var vf = (\<lambda>Y. P vf)";
+apply (rule ext)
+by simp
+
+lemma subst_res_Vals_beta [simp]: "(\<lambda>Y. P (the_In3 Y))\<leftarrow>Vals vs = (\<lambda>Y. P vs)";
+apply (rule ext)
+by simp
+*)
+
+subsection "subst-Bool"
+
+constdefs
+  subst_Bool  :: "'a assn \<Rightarrow> bool \<Rightarrow> 'a assn"             ("_\<leftarrow>=_" [60,61] 60)
+ "P\<leftarrow>=b \<equiv> \<lambda>Y s Z. \<exists>v. P (Val v) s Z \<and> (normal s \<longrightarrow> the_Bool v=b)"
+
+lemma subst_Bool_def2 [simp]: 
+"(P\<leftarrow>=b) Y s Z = (\<exists>v. P (Val v) s Z \<and> (normal s \<longrightarrow> the_Bool v=b))"
+apply (unfold subst_Bool_def)
+apply (simp (no_asm))
+done
+
+lemma subst_Bool_the_BoolI: "P (Val b) s Z \<Longrightarrow> (P\<leftarrow>=the_Bool b) Y s Z"
+apply auto
+done
+
+subsection "peek-res"
+
+constdefs
+  peek_res    :: "(res \<Rightarrow> 'a assn) \<Rightarrow> 'a assn"
+ "peek_res Pf \<equiv> \<lambda>Y. Pf Y Y"
+
+syntax
+"@peek_res"  :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"            ("\<lambda>_:. _" [0,3] 3)
+translations
+  "\<lambda>w:. P"   == "peek_res (\<lambda>w. P)"
+
+lemma peek_res_def2 [simp]: "peek_res P Y = P Y Y"
+apply (unfold peek_res_def)
+apply (simp (no_asm))
+done
+
+lemma peek_res_subst_res [simp]: "peek_res P\<leftarrow>w = P w\<leftarrow>w"
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+(* unused *)
+lemma peek_subst_res_allI: 
+ "(\<And>a. T a (P (f a)\<leftarrow>f a)) \<Longrightarrow> \<forall>a. T a (peek_res P\<leftarrow>f a)"
+apply (rule allI)
+apply (simp (no_asm))
+apply fast
+done
+
+subsection "ign-res"
+
+constdefs
+  ign_res    ::  "        'a assn \<Rightarrow> 'a assn"            ("_\<down>" [1000] 1000)
+  "P\<down>        \<equiv> \<lambda>Y s Z. \<exists>Y. P Y s Z"
+
+lemma ign_res_def2 [simp]: "P\<down> Y s Z = (\<exists>Y. P Y s Z)"
+apply (unfold ign_res_def)
+apply (simp (no_asm))
+done
+
+lemma ign_ign_res [simp]: "P\<down>\<down> = P\<down>"
+apply (rule ext)
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+lemma ign_subst_res [simp]: "P\<down>\<leftarrow>w = P\<down>"
+apply (rule ext)
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+lemma peek_and_ign_res [simp]: "(P \<and>. p)\<down> = (P\<down> \<and>. p)"
+apply (rule ext)
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+subsection "peek-st"
+
+constdefs
+  peek_st    :: "(st \<Rightarrow> 'a assn) \<Rightarrow> 'a assn"
+ "peek_st P \<equiv> \<lambda>Y s. P (store s) Y s"
+
+syntax
+"@peek_st"   :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"            ("\<lambda>_.. _" [0,3] 3)
+translations
+  "\<lambda>s.. P"   == "peek_st (\<lambda>s. P)"
+
+lemma peek_st_def2 [simp]: "(\<lambda>s.. Pf s) Y s = Pf (store s) Y s"
+apply (unfold peek_st_def)
+apply (simp (no_asm))
+done
+
+lemma peek_st_triv [simp]: "(\<lambda>s.. P) = P"
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+lemma peek_st_st [simp]: "(\<lambda>s.. \<lambda>s'.. P s s') = (\<lambda>s.. P s s)"
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+lemma peek_st_split [simp]: "(\<lambda>s.. \<lambda>Y s'. P s Y s') = (\<lambda>Y s. P (store s) Y s)"
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+lemma peek_st_subst_res [simp]: "(\<lambda>s.. P s)\<leftarrow>w = (\<lambda>s.. P s\<leftarrow>w)"
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+lemma peek_st_Normal [simp]: "(\<lambda>s..(Normal (P s))) = Normal (\<lambda>s.. P s)"
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+subsection "ign-res-eq"
+
+constdefs
+  ign_res_eq :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn"               ("_\<down>=_"  [60,61] 60)
+ "P\<down>=w       \<equiv> \<lambda>Y:. P\<down> \<and>. (\<lambda>s. Y=w)"
+
+lemma ign_res_eq_def2 [simp]: "(P\<down>=w) Y s Z = ((\<exists>Y. P Y s Z) \<and> Y=w)"
+apply (unfold ign_res_eq_def)
+apply auto
+done
+
+lemma ign_ign_res_eq [simp]: "(P\<down>=w)\<down> = P\<down>"
+apply (rule ext)
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+(* unused *)
+lemma ign_res_eq_subst_res: "P\<down>=w\<leftarrow>w = P\<down>"
+apply (rule ext)
+apply (rule ext)
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+(* unused *)
+lemma subst_Bool_ign_res_eq: "((P\<leftarrow>=b)\<down>=x) Y s Z = ((P\<leftarrow>=b) Y s Z  \<and> Y=x)"
+apply (simp (no_asm))
+done
+
+subsection "RefVar"
+
+constdefs
+  RefVar    :: "(state \<Rightarrow> vvar \<times> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn"(infixr "..;" 13)
+ "vf ..; P \<equiv> \<lambda>Y s. let (v,s') = vf s in P (Var v) s'"
+ 
+lemma RefVar_def2 [simp]: "(vf ..; P) Y s =  
+  P (Var (fst (vf s))) (snd (vf s))"
+apply (unfold RefVar_def Let_def)
+apply (simp (no_asm) add: split_beta)
+done
+
+subsection "allocation"
+
+constdefs
+  Alloc      :: "prog \<Rightarrow> obj_tag \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
+ "Alloc G otag P \<equiv> \<lambda>Y s Z.
+                   \<forall>s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a\<rightarrow> s'\<longrightarrow> P (Val (Addr a)) s' Z"
+
+  SXAlloc     :: "prog \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
+ "SXAlloc G P \<equiv> \<lambda>Y s Z. \<forall>s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<longrightarrow> P Y s' Z"
+
+
+lemma Alloc_def2 [simp]: "Alloc G otag P Y s Z =  
+       (\<forall>s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a\<rightarrow> s'\<longrightarrow> P (Val (Addr a)) s' Z)"
+apply (unfold Alloc_def)
+apply (simp (no_asm))
+done
+
+lemma SXAlloc_def2 [simp]: 
+  "SXAlloc G P Y s Z = (\<forall>s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<longrightarrow> P Y s' Z)"
+apply (unfold SXAlloc_def)
+apply (simp (no_asm))
+done
+
+section "validity"
+
+constdefs
+  type_ok  :: "prog \<Rightarrow> term \<Rightarrow> state \<Rightarrow> bool"
+ "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)"
+
+datatype    'a triple = triple "('a assn)" "term" "('a assn)" (** should be
+something like triple = \<forall>'a. triple ('a assn) term ('a assn)   **)
+                                        ("{(1_)}/ _>/ {(1_)}"      [3,65,3]75)
+types    'a triples = "'a triple set"
+
+syntax
+
+  var_triple   :: "['a assn, var         ,'a assn] \<Rightarrow> 'a triple"
+                                         ("{(1_)}/ _=>/ {(1_)}"    [3,80,3] 75)
+  expr_triple  :: "['a assn, expr        ,'a assn] \<Rightarrow> 'a triple"
+                                         ("{(1_)}/ _->/ {(1_)}"    [3,80,3] 75)
+  exprs_triple :: "['a assn, expr list   ,'a assn] \<Rightarrow> 'a triple"
+                                         ("{(1_)}/ _#>/ {(1_)}"    [3,65,3] 75)
+  stmt_triple  :: "['a assn, stmt,        'a assn] \<Rightarrow> 'a triple"
+                                         ("{(1_)}/ ._./ {(1_)}"     [3,65,3] 75)
+
+syntax (xsymbols)
+
+  triple       :: "['a assn, term        ,'a assn] \<Rightarrow> 'a triple"
+                                         ("{(1_)}/ _\<succ>/ {(1_)}"     [3,65,3] 75)
+  var_triple   :: "['a assn, var         ,'a assn] \<Rightarrow> 'a triple"
+                                         ("{(1_)}/ _=\<succ>/ {(1_)}"    [3,80,3] 75)
+  expr_triple  :: "['a assn, expr        ,'a assn] \<Rightarrow> 'a triple"
+                                         ("{(1_)}/ _-\<succ>/ {(1_)}"    [3,80,3] 75)
+  exprs_triple :: "['a assn, expr list   ,'a assn] \<Rightarrow> 'a triple"
+                                         ("{(1_)}/ _\<doteq>\<succ>/ {(1_)}"    [3,65,3] 75)
+
+translations
+  "{P} e-\<succ> {Q}" == "{P} In1l e\<succ> {Q}"
+  "{P} e=\<succ> {Q}" == "{P} In2  e\<succ> {Q}"
+  "{P} e\<doteq>\<succ> {Q}" == "{P} In3  e\<succ> {Q}"
+  "{P} .c. {Q}" == "{P} In1r c\<succ> {Q}"
+
+lemma inj_triple: "inj (\<lambda>(P,t,Q). {P} t\<succ> {Q})"
+apply (rule injI)
+apply auto
+done
+
+lemma triple_inj_eq: "({P} t\<succ> {Q} = {P'} t'\<succ> {Q'} ) = (P=P' \<and> t=t' \<and> Q=Q')"
+apply auto
+done
+
+constdefs
+  mtriples  :: "('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<Rightarrow> 'sig \<Rightarrow> expr) \<Rightarrow> 
+                ('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<times>  'sig) set \<Rightarrow> 'a triples"
+                                     ("{{(1_)}/ _-\<succ>/ {(1_)} | _}"[3,65,3,65]75)
+ "{{P} tf-\<succ> {Q} | ms} \<equiv> (\<lambda>(C,sig). {Normal(P C sig)} tf C sig-\<succ> {Q C sig})`ms"
+  
+consts
+
+ triple_valid :: "prog \<Rightarrow> nat \<Rightarrow>        'a triple  \<Rightarrow> bool"
+                                                (   "_\<Turnstile>_:_" [61,0, 58] 57)
+    ax_valids :: "prog \<Rightarrow> 'b triples \<Rightarrow> 'a triples \<Rightarrow> bool"
+                                                ("_,_|\<Turnstile>_"   [61,58,58] 57)
+    ax_derivs :: "prog \<Rightarrow> ('b triples \<times> 'a triples) set"
+
+syntax
+
+ triples_valid:: "prog \<Rightarrow> nat \<Rightarrow>         'a triples \<Rightarrow> bool"
+                                                (  "_||=_:_" [61,0, 58] 57)
+     ax_valid :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
+                                                ( "_,_|=_"   [61,58,58] 57)
+     ax_Derivs:: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triples \<Rightarrow> bool"
+                                                ("_,_||-_"   [61,58,58] 57)
+     ax_Deriv :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
+                                                ( "_,_|-_"   [61,58,58] 57)
+
+syntax (xsymbols)
+
+ triples_valid:: "prog \<Rightarrow> nat \<Rightarrow>         'a triples \<Rightarrow> bool"
+                                                (  "_|\<Turnstile>_:_" [61,0, 58] 57)
+     ax_valid :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
+                                                ( "_,_\<Turnstile>_"   [61,58,58] 57)
+     ax_Derivs:: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triples \<Rightarrow> bool"
+                                                ("_,_|\<turnstile>_"   [61,58,58] 57)
+     ax_Deriv :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
+                                                ( "_,_\<turnstile>_"   [61,58,58] 57)
+
+defs  triple_valid_def:  "G\<Turnstile>n:t  \<equiv> case t of {P} t\<succ> {Q} \<Rightarrow>
+                          \<forall>Y s Z. P Y s Z \<longrightarrow> type_ok G t s \<longrightarrow>
+                          (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> Q Y' s' Z)"
+translations         "G|\<Turnstile>n:ts" == "Ball ts (triple_valid G n)"
+defs   ax_valids_def:"G,A|\<Turnstile>ts  \<equiv>  \<forall>n. G|\<Turnstile>n:A \<longrightarrow> G|\<Turnstile>n:ts"
+translations         "G,A \<Turnstile>t"  == "G,A|\<Turnstile>{t}"
+                     "G,A|\<turnstile>ts" == "(A,ts) \<in> ax_derivs G"
+                     "G,A \<turnstile>t"  == "G,A|\<turnstile>{t}"
+
+lemma triple_valid_def2: "G\<Turnstile>n:{P} t\<succ> {Q} =  
+ (\<forall>Y s Z. P Y s Z 
+  \<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> 
+  (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s')\<longrightarrow> Q Y' s' Z))"
+apply (unfold triple_valid_def type_ok_def)
+apply (simp (no_asm))
+done
+
+
+declare split_paired_All [simp del] split_paired_Ex [simp del] 
+declare split_if     [split del] split_if_asm     [split del] 
+        option.split [split del] option.split_asm [split del]
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac";
+claset_ref () := claset () delSWrapper "split_all_tac"
+*}
+
+
+inductive "ax_derivs G" intros
+
+  empty: " G,A|\<turnstile>{}"
+  insert:"\<lbrakk>G,A\<turnstile>t; G,A|\<turnstile>ts\<rbrakk> \<Longrightarrow>
+          G,A|\<turnstile>insert t ts"
+
+  asm:   "ts\<subseteq>A \<Longrightarrow> G,A|\<turnstile>ts"
+
+(* could be added for convenience and efficiency, but is not necessary
+  cut:   "\<lbrakk>G,A'|\<turnstile>ts; G,A|\<turnstile>A'\<rbrakk> \<Longrightarrow>
+           G,A |\<turnstile>ts"
+*)
+  weaken:"\<lbrakk>G,A|\<turnstile>ts'; ts \<subseteq> ts'\<rbrakk> \<Longrightarrow> G,A|\<turnstile>ts"
+
+  conseq:"\<forall>Y s Z . P  Y s Z  \<longrightarrow> (\<exists>P' Q'. G,A\<turnstile>{P'} t\<succ> {Q'} \<and> (\<forall>Y' s'. 
+         (\<forall>Y   Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>
+                                 Q  Y' s' Z ))
+                                         \<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q }"
+
+  hazard:"G,A\<turnstile>{P \<and>. Not \<circ> type_ok G t} t\<succ> {Q}"
+
+  Abrupt:  "G,A\<turnstile>{P\<leftarrow>(arbitrary3 t) \<and>. Not \<circ> normal} t\<succ> {P}"
+
+  (* variables *)
+  LVar:  " G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Var (lvar vn s))} LVar vn=\<succ> {P}"
+
+  FVar: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Q};
+          G,A\<turnstile>{Q} e-\<succ> {\<lambda>Val:a:. fvar C stat fn a ..; R}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} {C,stat}e..fn=\<succ> {R}"
+
+  AVar:  "\<lbrakk>G,A\<turnstile>{Normal P} e1-\<succ> {Q};
+          \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} e2-\<succ> {\<lambda>Val:i:. avar G i a ..; R}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} e1.[e2]=\<succ> {R}"
+  (* expressions *)
+
+  NewC: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Alloc G (CInst C) Q}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}"
+
+  NewA: "\<lbrakk>G,A\<turnstile>{Normal P} .init_comp_ty T. {Q};  G,A\<turnstile>{Q} e-\<succ>
+	  {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} New T[e]-\<succ> {R}"
+
+  Cast: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
+          abupd (raise_if (\<not>G,s\<turnstile>v fits T) ClassCast) .; Q\<leftarrow>Val v}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} Cast T e-\<succ> {Q}"
+
+  Inst: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
+                  Q\<leftarrow>Val (Bool (v\<noteq>Null \<and> G,s\<turnstile>v fits RefT T))}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} e InstOf T-\<succ> {Q}"
+
+  Lit:                          "G,A\<turnstile>{Normal (P\<leftarrow>Val v)} Lit v-\<succ> {P}"
+
+  Super:" G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Val (val_this s))} Super-\<succ> {P}"
+
+  Acc:  "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {\<lambda>Var:(v,f):. Q\<leftarrow>Val v}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} Acc va-\<succ> {Q}"
+
+  Ass:  "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {Q};
+     \<forall>vf. G,A\<turnstile>{Q\<leftarrow>Var vf} e-\<succ> {\<lambda>Val:v:. assign (snd vf) v .; R}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} va:=e-\<succ> {R}"
+
+  Cond: "\<lbrakk>G,A \<turnstile>{Normal P} e0-\<succ> {P'};
+          \<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} (if b then e1 else e2)-\<succ> {Q}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} e0 ? e1 : e2-\<succ> {Q}"
+
+  Call: 
+"\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q}; \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {R a};
+  \<forall>a vs invC declC l. G,A\<turnstile>{(R a\<leftarrow>Vals vs \<and>.
+ (\<lambda>s. declC=invocation_declclass G mode (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> \<and>
+      invC = invocation_class mode (store s) a statT \<and>
+         l = locals (store s)) ;.
+      init_lvars G declC \<lparr>name=mn,parTs=pTs\<rparr> mode a vs) \<and>.
+      (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)}
+ Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}\<rbrakk> \<Longrightarrow>
+         G,A\<turnstile>{Normal P} {statT,mode}e\<cdot>mn({pTs}args)-\<succ> {S}"
+
+  Methd:"\<lbrakk>G,A\<union> {{P} Methd-\<succ> {Q} | ms} |\<turnstile> {{P} body G-\<succ> {Q} | ms}\<rbrakk> \<Longrightarrow>
+                                 G,A|\<turnstile>{{P} Methd-\<succ>  {Q} | ms}"
+
+  Body: "\<lbrakk>G,A\<turnstile>{Normal P} .Init D. {Q}; 
+  G,A\<turnstile>{Q} .c. {\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>(In1 (the (locals s Result)))}\<rbrakk> 
+    \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} Body D c-\<succ> {R}"
+  
+  (* expression lists *)
+
+  Nil:                          "G,A\<turnstile>{Normal (P\<leftarrow>Vals [])} []\<doteq>\<succ> {P}"
+
+  Cons: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q};
+          \<forall>v. G,A\<turnstile>{Q\<leftarrow>Val v} es\<doteq>\<succ> {\<lambda>Vals:vs:. R\<leftarrow>Vals (v#vs)}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} e#es\<doteq>\<succ> {R}"
+
+  (* statements *)
+
+  Skip:                         "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit>)} .Skip. {P}"
+
+  Expr: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} .Expr e. {Q}"
+
+  Lab: "\<lbrakk>G,A\<turnstile>{Normal P} .c. {abupd (absorb (Break l)) .; Q}\<rbrakk> \<Longrightarrow>
+                           G,A\<turnstile>{Normal P} .l\<bullet> c. {Q}"
+
+  Comp: "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q};
+          G,A\<turnstile>{Q} .c2. {R}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} .c1;;c2. {R}"
+
+  If:   "\<lbrakk>G,A \<turnstile>{Normal P} e-\<succ> {P'};
+          \<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c1 else c2). {Q}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} .If(e) c1 Else c2. {Q}"
+(* unfolding variant of Loop, not needed here
+  LoopU:"\<lbrakk>G,A \<turnstile>{Normal P} e-\<succ> {P'};
+          \<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c;;While(e) c else Skip).{Q}\<rbrakk>
+         \<Longrightarrow>              G,A\<turnstile>{Normal P} .While(e) c. {Q}"
+*)
+  Loop: "\<lbrakk>G,A\<turnstile>{P} e-\<succ> {P'}; 
+          G,A\<turnstile>{Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}\<rbrakk> \<Longrightarrow>
+                            G,A\<turnstile>{P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}"
+(** Beware of polymorphic_Loop below: should be identical terms **)
+  
+  Do: "G,A\<turnstile>{Normal (abupd (\<lambda>a. (Some (Jump j))) .; P\<leftarrow>\<diamondsuit>)} .Do j. {P}"
+
+  Throw:"\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} .Throw e. {Q}"
+
+  Try:  "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {SXAlloc G Q};
+          G,A\<turnstile>{Q \<and>. (\<lambda>s.  G,s\<turnstile>catch C) ;. new_xcpt_var vn} .c2. {R};
+              (Q \<and>. (\<lambda>s. \<not>G,s\<turnstile>catch C)) \<Rightarrow> R\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} .Try c1 Catch(C vn) c2. {R}"
+
+  Fin:  "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q};
+      \<forall>x. G,A\<turnstile>{Q \<and>. (\<lambda>s. x = fst s) ;. abupd (\<lambda>x. None)}
+              .c2. {abupd (abrupt_if (x\<noteq>None) x) .; R}\<rbrakk> \<Longrightarrow>
+                                 G,A\<turnstile>{Normal P} .c1 Finally c2. {R}"
+
+  Done:                       "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit> \<and>. initd C)} .Init C. {P}"
+
+  Init: "\<lbrakk>the (class G C) = c;
+          G,A\<turnstile>{Normal ((P \<and>. Not \<circ> initd C) ;. supd (init_class_obj G C))}
+              .(if C = Object then Skip else Init (super c)). {Q};
+      \<forall>l. G,A\<turnstile>{Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars empty}
+              .init c. {set_lvars l .; R}\<rbrakk> \<Longrightarrow>
+                               G,A\<turnstile>{Normal (P \<and>. Not \<circ> initd C)} .Init C. {R}"
+
+axioms (** these terms are the same as above, but with generalized typing **)
+  polymorphic_conseq:
+        "\<forall>Y s Z . P  Y s Z  \<longrightarrow> (\<exists>P' Q'. G,A\<turnstile>{P'} t\<succ> {Q'} \<and> (\<forall>Y' s'. 
+        (\<forall>Y   Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>
+                                Q  Y' s' Z ))
+                                         \<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q }"
+
+  polymorphic_Loop:
+        "\<lbrakk>G,A\<turnstile>{P} e-\<succ> {P'}; 
+          G,A\<turnstile>{Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}\<rbrakk> \<Longrightarrow>
+                            G,A\<turnstile>{P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}"
+
+constdefs
+ adapt_pre :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
+"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)"
+
+
+section "rules derived by induction"
+
+lemma cut_valid: "\<lbrakk>G,A'|\<Turnstile>ts; G,A|\<Turnstile>A'\<rbrakk> \<Longrightarrow> G,A|\<Turnstile>ts"
+apply (unfold ax_valids_def)
+apply fast
+done
+
+(*if cut is available
+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>  
+       G,A|\<turnstile>ts"
+b y etac ax_derivs.cut 1;
+b y eatac ax_derivs.asm 1 1;
+qed "ax_thin";
+*)
+lemma ax_thin [rule_format (no_asm)]: 
+  "G,(A'::'a triple set)|\<turnstile>(ts::'a triple set) \<Longrightarrow> \<forall>A. A' \<subseteq> A \<longrightarrow> G,A|\<turnstile>ts"
+apply (erule ax_derivs.induct)
+apply                (tactic "ALLGOALS(EVERY'[Clarify_tac,REPEAT o smp_tac 1])")
+apply                (rule ax_derivs.empty)
+apply               (erule (1) ax_derivs.insert)
+apply              (fast intro: ax_derivs.asm)
+(*apply           (fast intro: ax_derivs.cut) *)
+apply            (fast intro: ax_derivs.weaken)
+apply           (rule ax_derivs.conseq, intro strip, tactic "smp_tac 3 1",clarify, tactic "smp_tac 1 1",rule exI, rule exI, erule (1) conjI)
+(* 31 subgoals *)
+prefer 16 (* Methd *)
+apply (rule ax_derivs.Methd, drule spec, erule mp, fast)
+apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) (thms "ax_derivs.intros")) 
+                     THEN_ALL_NEW Blast_tac) *})
+apply (erule ax_derivs.Call)
+apply   clarify 
+apply   blast
+
+apply   (rule allI)+ 
+apply   (drule spec)+
+apply   blast
+done
+
+lemma ax_thin_insert: "G,(A::'a triple set)\<turnstile>(t::'a triple) \<Longrightarrow> G,insert x A\<turnstile>t"
+apply (erule ax_thin)
+apply fast
+done
+
+lemma subset_mtriples_iff: 
+  "ts \<subseteq> {{P} mb-\<succ> {Q} | ms} = (\<exists>ms'. ms'\<subseteq>ms \<and>  ts = {{P} mb-\<succ> {Q} | ms'})"
+apply (unfold mtriples_def)
+apply (rule subset_image_iff)
+done
+
+lemma weaken: 
+ "G,(A::'a triple set)|\<turnstile>(ts'::'a triple set) \<Longrightarrow> !ts. ts \<subseteq> ts' \<longrightarrow> G,A|\<turnstile>ts"
+apply (erule ax_derivs.induct)
+(*36 subgoals*)
+apply       (tactic "ALLGOALS strip_tac")
+apply       (tactic {* ALLGOALS(REPEAT o (EVERY'[dtac (thm "subset_singletonD"),
+         etac disjE, fast_tac (claset() addSIs [thm "ax_derivs.empty"])]))*})
+apply       (tactic "TRYALL hyp_subst_tac")
+apply       (simp, rule ax_derivs.empty)
+apply      (drule subset_insertD)
+apply      (blast intro: ax_derivs.insert)
+apply     (fast intro: ax_derivs.asm)
+(*apply  (blast intro: ax_derivs.cut) *)
+apply   (fast intro: ax_derivs.weaken)
+apply  (rule ax_derivs.conseq, clarify, tactic "smp_tac 3 1", blast(* unused *))
+(*31 subgoals*)
+apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) (thms "ax_derivs.intros")) 
+                   THEN_ALL_NEW Fast_tac) *})
+(*1 subgoal*)
+apply (clarsimp simp add: subset_mtriples_iff)
+apply (rule ax_derivs.Methd)
+apply (drule spec)
+apply (erule impE)
+apply  (rule exI)
+apply  (erule conjI)
+apply  (rule HOL.refl)
+oops (* dead end, Methd is to blame *)
+
+
+section "rules derived from conseq"
+
+lemma conseq12: "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q'};  
+ \<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>  
+  Q Y' s' Z)\<rbrakk>  
+  \<Longrightarrow>  G,A\<turnstile>{P ::'a assn} t\<succ> {Q }"
+apply (rule polymorphic_conseq)
+apply clarsimp
+apply blast
+done
+
+(*unused, but nice variant*)
+lemma conseq12': "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q'}; \<forall>s Y' s'.  
+       (\<forall>Y Z. P' Y s Z \<longrightarrow> Q' Y' s' Z) \<longrightarrow>  
+       (\<forall>Y Z. P  Y s Z \<longrightarrow> Q  Y' s' Z)\<rbrakk>  
+  \<Longrightarrow>  G,A\<turnstile>{P } t\<succ> {Q }"
+apply (erule conseq12)
+apply fast
+done
+
+lemma conseq12_from_conseq12': "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q'};  
+ \<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>  
+  Q Y' s' Z)\<rbrakk>  
+  \<Longrightarrow>  G,A\<turnstile>{P } t\<succ> {Q }"
+apply (erule conseq12')
+apply blast
+done
+
+lemma conseq1: "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q}; P \<Rightarrow> P'\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q}"
+apply (erule conseq12)
+apply blast
+done
+
+lemma conseq2: "\<lbrakk>G,A\<turnstile>{P} t\<succ> {Q'}; Q' \<Rightarrow> Q\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
+apply (erule conseq12)
+apply blast
+done
+
+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>  
+  G,A\<turnstile>{P} t\<succ> {Q}"
+apply (rule polymorphic_conseq)
+apply force
+done
+
+(* unused *)
+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}"
+apply (rule ax_escape (* unused *))
+apply clarify
+apply (rule conseq12)
+apply  fast
+apply auto
+done
+(*alternative (more direct) proof:
+apply (rule ax_derivs.conseq) *)(* unused *)(*
+apply (fast)
+*)
+
+
+lemma ax_impossible [intro]: "G,A\<turnstile>{\<lambda>Y s Z. False} t\<succ> {Q}"
+apply (rule ax_escape)
+apply clarify
+done
+
+(* unused *)
+lemma ax_nochange_lemma: "\<lbrakk>P Y s; All (op = w)\<rbrakk> \<Longrightarrow> P w s"
+apply auto
+done
+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}"
+apply (erule conseq12)
+apply auto
+apply (erule (1) ax_nochange_lemma)
+done
+
+(* unused *)
+lemma ax_trivial: "G,A\<turnstile>{P}  t\<succ> {\<lambda>Y s Z. True}"
+apply (rule polymorphic_conseq(* unused *))
+apply auto
+done
+
+(* unused *)
+lemma ax_disj: "\<lbrakk>G,A\<turnstile>{P1} t\<succ> {Q1}; G,A\<turnstile>{P2} t\<succ> {Q2}\<rbrakk> \<Longrightarrow>  
+  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}"
+apply (rule ax_escape (* unused *))
+apply safe
+apply  (erule conseq12, fast)+
+done
+
+(* unused *)
+lemma ax_supd_shuffle: "(\<exists>Q. G,A\<turnstile>{P} .c1. {Q} \<and> G,A\<turnstile>{Q ;. f} .c2. {R}) =  
+       (\<exists>Q'. G,A\<turnstile>{P} .c1. {f .; Q'} \<and> G,A\<turnstile>{Q'} .c2. {R})"
+apply (best elim!: conseq1 conseq2)
+done
+
+lemma ax_cases: "\<lbrakk>G,A\<turnstile>{P \<and>.       C} t\<succ> {Q};  
+                       G,A\<turnstile>{P \<and>. Not \<circ> C} t\<succ> {Q}\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
+apply (unfold peek_and_def)
+apply (rule ax_escape)
+apply clarify
+apply (case_tac "C s")
+apply  (erule conseq12, force)+
+done
+(*alternative (more direct) proof:
+apply (rule rtac ax_derivs.conseq) *)(* unused *)(*
+apply clarify
+apply (case_tac "C s")
+apply  force+
+*)
+
+lemma ax_adapt: "G,A\<turnstile>{P} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}"
+apply (unfold adapt_pre_def)
+apply (erule conseq12)
+apply fast
+done
+
+lemma adapt_pre_adapts: "G,A\<Turnstile>{P} t\<succ> {Q} \<longrightarrow> G,A\<Turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}"
+apply (unfold adapt_pre_def)
+apply (simp add: ax_valids_def triple_valid_def2)
+apply fast
+done
+
+
+lemma adapt_pre_weakest: 
+"\<forall>G (A::'a triple set) t. G,A\<Turnstile>{P} t\<succ> {Q} \<longrightarrow> G,A\<Turnstile>{P'} t\<succ> {Q'} \<Longrightarrow>  
+  P' \<Rightarrow> adapt_pre P Q (Q'::'a assn)"
+apply (unfold adapt_pre_def)
+apply (drule spec)
+apply (drule_tac x = "{}" in spec)
+apply (drule_tac x = "In1r Skip" in spec)
+apply (simp add: ax_valids_def triple_valid_def2)
+oops
+
+(*
+Goal "\<forall>(A::'a triple set) t. G,A\<Turnstile>{P} t\<succ> {Q} \<longrightarrow> G,A\<Turnstile>{P'} t\<succ> {Q'} \<Longrightarrow>  
+  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}"
+b y fatac ax_sound 1 1;
+b y asm_full_simp_tac (simpset() addsimps [ax_valids_def,triple_valid_def2]) 1;
+b y rtac ax_no_hazard 1; 
+b y etac conseq12 1;
+b y Clarify_tac 1;
+b y case_tac "\<forall>Z. \<not>P Y s Z" 1;
+b y smp_tac 2 1;
+b y etac thin_rl 1;
+b y etac thin_rl 1;
+b y clarsimp_tac (claset(), simpset() addsimps [type_ok_def]) 1;
+b y subgoal_tac "G|\<Turnstile>n:A" 1;
+b y smp_tac 1 1;
+b y smp_tac 3 1;
+b y etac impE 1;
+ back();
+ b y Fast_tac 1;
+b y 
+b y rotate_tac 2 1;
+b y etac thin_rl 1;
+b y  etac thin_rl 2;
+b y  etac thin_rl 2;
+b y  Clarify_tac 2;
+b y  dtac spec 2;
+b y  EVERY'[dtac spec, mp_tac] 2;
+b y  thin_tac "\<forall>n Y s Z. ?PP n Y s Z" 2;
+b y  thin_tac "P' Y s Z" 2;
+b y  Blast_tac 2;
+b y smp_tac 3 1;
+b y case_tac "\<forall>Z. \<not>P Y s Z" 1;
+b y dres_inst_tac [("x","In1r Skip")] spec 1;
+b y Full_simp_tac 1;
+*)
+
+lemma peek_and_forget1_Normal: 
+ "G,A\<turnstile>{Normal P} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{Normal (P \<and>. p)} t\<succ> {Q}"
+apply (erule conseq1)
+apply (simp (no_asm))
+done
+
+lemma peek_and_forget1: "G,A\<turnstile>{P} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{P \<and>. p} t\<succ> {Q}"
+apply (erule conseq1)
+apply (simp (no_asm))
+done
+
+lemmas ax_NormalD = peek_and_forget1 [of _ _ _ _ _ normal] 
+
+lemma peek_and_forget2: "G,A\<turnstile>{P} t\<succ> {Q \<and>. p} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
+apply (erule conseq2)
+apply (simp (no_asm))
+done
+
+lemma ax_subst_Val_allI: "\<forall>v. G,A\<turnstile>{(P'               v )\<leftarrow>Val v} t\<succ> {Q v} \<Longrightarrow>  
+      \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In1 w))\<leftarrow>Val v} t\<succ> {Q v}"
+apply (force elim!: conseq1)
+done
+
+lemma ax_subst_Var_allI: "\<forall>v. G,A\<turnstile>{(P'               v )\<leftarrow>Var v} t\<succ> {Q v} \<Longrightarrow>  
+      \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In2 w))\<leftarrow>Var v} t\<succ> {Q v}"
+apply (force elim!: conseq1)
+done
+
+lemma ax_subst_Vals_allI: "(\<forall>v. G,A\<turnstile>{(     P'          v )\<leftarrow>Vals v} t\<succ> {Q v}) \<Longrightarrow>  
+       \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In3 w))\<leftarrow>Vals v} t\<succ> {Q v}"
+apply (force elim!: conseq1)
+done
+
+
+section "alternative axioms"
+
+lemma ax_Lit2: 
+  "G,(A::'a triple set)\<turnstile>{Normal P::'a assn} Lit v-\<succ> {Normal (P\<down>=Val v)}"
+apply (rule ax_derivs.Lit [THEN conseq1])
+apply force
+done
+lemma ax_Lit2_test_complete: 
+  "G,(A::'a triple set)\<turnstile>{Normal (P\<leftarrow>Val v)::'a assn} Lit v-\<succ> {P}"
+apply (rule ax_Lit2 [THEN conseq2])
+apply force
+done
+
+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))}"
+apply (rule ax_derivs.LVar [THEN conseq1])
+apply force
+done
+
+lemma ax_Super2: "G,(A::'a triple set)\<turnstile>
+  {Normal P::'a assn} Super-\<succ> {Normal (\<lambda>s.. P\<down>=Val (val_this s))}"
+apply (rule ax_derivs.Super [THEN conseq1])
+apply force
+done
+
+lemma ax_Nil2: 
+  "G,(A::'a triple set)\<turnstile>{Normal P::'a assn} []\<doteq>\<succ> {Normal (P\<down>=Vals [])}"
+apply (rule ax_derivs.Nil [THEN conseq1])
+apply force
+done
+
+
+section "misc derived structural rules"
+
+(* unused *)
+lemma ax_finite_mtriples_lemma: "\<lbrakk>F \<subseteq> ms; finite ms; \<forall>(C,sig)\<in>ms. 
+    G,(A::'a triple set)\<turnstile>{Normal (P C sig)::'a assn} mb C sig-\<succ> {Q C sig}\<rbrakk> \<Longrightarrow> 
+       G,A|\<turnstile>{{P} mb-\<succ> {Q} | F}"
+apply (frule (1) finite_subset)
+apply (erule make_imp)
+apply (erule thin_rl)
+apply (erule finite_induct)
+apply  (unfold mtriples_def)
+apply  (clarsimp intro!: ax_derivs.empty ax_derivs.insert)+
+apply force
+done
+lemmas ax_finite_mtriples = ax_finite_mtriples_lemma [OF subset_refl]
+
+lemma ax_derivs_insertD: 
+ "G,(A::'a triple set)|\<turnstile>insert (t::'a triple) ts \<Longrightarrow> G,A\<turnstile>t \<and> G,A|\<turnstile>ts"
+apply (fast intro: ax_derivs.weaken)
+done
+
+lemma ax_methods_spec: 
+"\<lbrakk>G,(A::'a triple set)|\<turnstile>split f ` ms; (C,sig) \<in> ms\<rbrakk>\<Longrightarrow> G,A\<turnstile>((f C sig)::'a triple)"
+apply (erule ax_derivs.weaken)
+apply (force del: image_eqI intro: rev_image_eqI)
+done
+
+(* this version is used to avoid using the cut rule *)
+lemma ax_finite_pointwise_lemma [rule_format]: "\<lbrakk>F \<subseteq> ms; finite ms\<rbrakk> \<Longrightarrow>  
+  ((\<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>  
+      G,A|\<turnstile>split f ` F \<longrightarrow> G,A|\<turnstile>split g ` F"
+apply (frule (1) finite_subset)
+apply (erule make_imp)
+apply (erule thin_rl)
+apply (erule finite_induct)
+apply  clarsimp+
+apply (drule ax_derivs_insertD)
+apply (rule ax_derivs.insert)
+apply  (simp (no_asm_simp) only: split_tupled_all)
+apply  (auto elim: ax_methods_spec)
+done
+lemmas ax_finite_pointwise = ax_finite_pointwise_lemma [OF subset_refl]
+ 
+lemma ax_no_hazard: 
+  "G,(A::'a triple set)\<turnstile>{P \<and>. type_ok G t} t\<succ> {Q::'a assn} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
+apply (erule ax_cases)
+apply (rule ax_derivs.hazard [THEN conseq1])
+apply force
+done
+
+lemma ax_free_wt: 
+ "(\<exists>T L C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) 
+  \<longrightarrow> G,(A::'a triple set)\<turnstile>{Normal P} t\<succ> {Q::'a assn} \<Longrightarrow> 
+  G,A\<turnstile>{Normal P} t\<succ> {Q}"
+apply (rule ax_no_hazard)
+apply (rule ax_escape)
+apply clarify
+apply (erule mp [THEN conseq12])
+apply  (auto simp add: type_ok_def)
+done
+
+ML {*
+bind_thms ("ax_Abrupts", sum3_instantiate (thm "ax_derivs.Abrupt"))
+*}
+declare ax_Abrupts [intro!]
+
+lemmas ax_Normal_cases = ax_cases [of _ _ normal]
+
+lemma ax_Skip [intro!]: "G,(A::'a triple set)\<turnstile>{P\<leftarrow>\<diamondsuit>} .Skip. {P::'a assn}"
+apply (rule ax_Normal_cases)
+apply  (rule ax_derivs.Skip)
+apply fast
+done
+lemmas ax_SkipI = ax_Skip [THEN conseq1, standard]
+
+
+section "derived rules for methd call"
+
+lemma ax_Call_known_DynT: 
+"\<lbrakk>G\<turnstile>IntVir\<rightarrow>C\<preceq>statT; 
+  \<forall>a vs l. G,A\<turnstile>{(R a\<leftarrow>Vals vs \<and>. (\<lambda>s. l = locals (store s)) ;.
+  init_lvars G C \<lparr>name=mn,parTs=pTs\<rparr> IntVir a vs)} 
+    Methd C \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}; 
+  \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ>  
+       {R a \<and>. (\<lambda>s. C = obj_class (the (heap (store s) (the_Addr a))) \<and>
+                     C = invocation_declclass 
+                            G IntVir (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> )};  
+       G,(A::'a triple set)\<turnstile>{Normal P} e-\<succ> {Q::'a assn}\<rbrakk>  
+   \<Longrightarrow> G,A\<turnstile>{Normal P} {statT,IntVir}e\<cdot>mn({pTs}args)-\<succ> {S}"
+apply (erule ax_derivs.Call)
+apply  safe
+apply  (erule spec)
+apply (rule ax_escape, clarsimp)
+apply (drule spec, drule spec, drule spec,erule conseq12)
+apply force
+done
+
+
+lemma ax_Call_Static: 
+ "\<lbrakk>\<forall>a vs l. G,A\<turnstile>{R a\<leftarrow>Vals vs \<and>. (\<lambda>s. l = locals (store s)) ;.  
+               init_lvars G C \<lparr>name=mn,parTs=pTs\<rparr> Static any_Addr vs}  
+              Methd C \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}; 
+  G,A\<turnstile>{Normal P} e-\<succ> {Q};
+  \<forall> a. G,(A::'a triple set)\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {(R::val \<Rightarrow> 'a assn)  a 
+  \<and>. (\<lambda> s. C=invocation_declclass 
+                G Static (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr>)}
+\<rbrakk>  \<Longrightarrow>  G,A\<turnstile>{Normal P} {statT,Static}e\<cdot>mn({pTs}args)-\<succ> {S}"
+apply (erule ax_derivs.Call)
+apply  safe
+apply  (erule spec)
+apply (rule ax_escape, clarsimp)
+apply (erule_tac V = "?P \<longrightarrow> ?Q" in thin_rl)
+apply (drule spec,drule spec,drule spec, erule conseq12)
+apply (force simp add: init_lvars_def)
+done
+
+lemma ax_Methd1: 
+ "\<lbrakk>G,A\<union>{{P} Methd-\<succ> {Q} | ms}|\<turnstile> {{P} body G-\<succ> {Q} | ms}; (C,sig)\<in> ms\<rbrakk> \<Longrightarrow> 
+       G,A\<turnstile>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
+apply (drule ax_derivs.Methd)
+apply (unfold mtriples_def)
+apply (erule (1) ax_methods_spec)
+done
+
+lemma ax_MethdN: 
+"G,insert({Normal P} Methd  C sig-\<succ> {Q}) A\<turnstile> 
+          {Normal P} body G C sig-\<succ> {Q} \<Longrightarrow>  
+      G,A\<turnstile>{Normal P} Methd   C sig-\<succ> {Q}"
+apply (rule ax_Methd1)
+apply  (rule_tac [2] singletonI)
+apply (unfold mtriples_def)
+apply clarsimp
+done
+
+lemma ax_StatRef: 
+  "G,(A::'a triple set)\<turnstile>{Normal (P\<leftarrow>Val Null)} StatRef rt-\<succ> {P::'a assn}"
+apply (rule ax_derivs.Cast)
+apply (rule ax_Lit2 [THEN conseq2])
+apply clarsimp
+done
+
+section "rules derived from Init and Done"
+
+  lemma ax_InitS: "\<lbrakk>the (class G C) = c; C \<noteq> Object;  
+     \<forall>l. G,A\<turnstile>{Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars empty}  
+            .init c. {set_lvars l .; R};   
+         G,A\<turnstile>{Normal ((P \<and>. Not \<circ> initd C) ;. supd (init_class_obj G C))}  
+  .Init (super c). {Q}\<rbrakk> \<Longrightarrow>  
+  G,(A::'a triple set)\<turnstile>{Normal (P \<and>. Not \<circ> initd C)} .Init C. {R::'a assn}"
+apply (erule ax_derivs.Init)
+apply  (simp (no_asm_simp))
+apply assumption
+done
+
+lemma ax_Init_Skip_lemma: 
+"\<forall>l. G,(A::'a triple set)\<turnstile>{P\<leftarrow>\<diamondsuit> \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars l'}
+  .Skip. {(set_lvars l .; P)::'a assn}"
+apply (rule allI)
+apply (rule ax_SkipI)
+apply clarsimp
+done
+
+lemma ax_triv_InitS: "\<lbrakk>the (class G C) = c;init c = Skip; C \<noteq> Object; 
+       P\<leftarrow>\<diamondsuit> \<Rightarrow> (supd (init_class_obj G C) .; P);  
+       G,A\<turnstile>{Normal (P \<and>. initd C)} .Init (super c). {(P \<and>. initd C)\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>  
+       G,(A::'a triple set)\<turnstile>{Normal P\<leftarrow>\<diamondsuit>} .Init C. {(P \<and>. initd C)::'a assn}"
+apply (rule_tac C = "initd C" in ax_cases)
+apply  (rule conseq1, rule ax_derivs.Done, clarsimp)
+apply (simp (no_asm))
+apply (erule (1) ax_InitS)
+apply  simp
+apply  (rule ax_Init_Skip_lemma)
+apply (erule conseq1)
+apply force
+done
+
+lemma ax_Init_Object: "wf_prog G \<Longrightarrow> G,(A::'a triple set)\<turnstile>
+  {Normal ((supd (init_class_obj G Object) .; P\<leftarrow>\<diamondsuit>) \<and>. Not \<circ> initd Object)} 
+       .Init Object. {(P \<and>. initd Object)::'a assn}"
+apply (rule ax_derivs.Init)
+apply   (drule class_Object, force)
+apply (simp_all (no_asm))
+apply (rule_tac [2] ax_Init_Skip_lemma)
+apply (rule ax_SkipI, force)
+done
+
+lemma ax_triv_Init_Object: "\<lbrakk>wf_prog G;  
+       (P::'a assn) \<Rightarrow> (supd (init_class_obj G Object) .; P)\<rbrakk> \<Longrightarrow>  
+  G,(A::'a triple set)\<turnstile>{Normal P\<leftarrow>\<diamondsuit>} .Init Object. {P \<and>. initd Object}"
+apply (rule_tac C = "initd Object" in ax_cases)
+apply  (rule conseq1, rule ax_derivs.Done, clarsimp)
+apply (erule ax_Init_Object [THEN conseq1])
+apply force
+done
+
+
+section "introduction rules for Alloc and SXAlloc"
+
+lemma ax_SXAlloc_Normal: "G,A\<turnstile>{P} .c. {Normal Q} \<Longrightarrow> G,A\<turnstile>{P} .c. {SXAlloc G Q}"
+apply (erule conseq2)
+apply (clarsimp elim!: sxalloc_elim_cases simp add: split_tupled_all)
+done
+
+lemma ax_Alloc: 
+  "G,A\<turnstile>{P} t\<succ> {Normal (\<lambda>Y (x,s) Z. (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
+ Q (Val (Addr a)) (Norm(init_obj G (CInst C) (Heap a) s)) Z)) \<and>. 
+    heap_free (Suc (Suc 0))}
+   \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Alloc G (CInst C) Q}"
+apply (erule conseq2)
+apply (auto elim!: halloc_elim_cases)
+done
+
+lemma ax_Alloc_Arr: 
+ "G,A\<turnstile>{P} t\<succ> {\<lambda>Val:i:. Normal (\<lambda>Y (x,s) Z. \<not>the_Intg i<0 \<and>  
+  (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
+  Q (Val (Addr a)) (Norm (init_obj G (Arr T (the_Intg i)) (Heap a) s)) Z)) \<and>. 
+   heap_free (Suc (Suc 0))} \<Longrightarrow>  
+ G,A\<turnstile>{P} t\<succ> {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T(the_Intg i)) Q}"
+apply (erule conseq2)
+apply (auto elim!: halloc_elim_cases)
+done
+
+lemma ax_SXAlloc_catch_SXcpt: 
+ "\<lbrakk>G,A\<turnstile>{P} t\<succ> {(\<lambda>Y (x,s) Z. x=Some (Xcpt (Std xn)) \<and>  
+  (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
+  Q Y (Some (Xcpt (Loc a)),init_obj G (CInst (SXcpt xn)) (Heap a) s) Z))  
+  \<and>. heap_free (Suc (Suc 0))}\<rbrakk> \<Longrightarrow>  
+  G,A\<turnstile>{P} t\<succ> {SXAlloc G (\<lambda>Y s Z. Q Y s Z \<and> G,s\<turnstile>catch SXcpt xn)}"
+apply (erule conseq2)
+apply (auto elim!: sxalloc_elim_cases halloc_elim_cases)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/AxSound.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,434 @@
+(*  Title:      isabelle/Bali/AxSound.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1999 Technische Universitaet Muenchen
+*)
+header {* Soundness proof for Axiomatic semantics of Java expressions and 
+          statements
+       *}
+
+
+
+theory AxSound = AxSem:
+
+section "validity"
+
+consts
+
+ triple_valid2:: "prog \<Rightarrow> nat \<Rightarrow>        'a triple  \<Rightarrow> bool"
+                                                (   "_\<Turnstile>_\<Colon>_"[61,0, 58] 57)
+    ax_valids2:: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool"
+                                                ("_,_|\<Turnstile>\<Colon>_" [61,58,58] 57)
+
+defs  triple_valid2_def: "G\<Turnstile>n\<Colon>t \<equiv> case t of {P} t\<succ> {Q} \<Rightarrow>
+ \<forall>Y s Z. P Y s Z \<longrightarrow> (\<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>
+ (\<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))))"
+
+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)"
+
+lemma triple_valid2_def2: "G\<Turnstile>n\<Colon>{P} t\<succ> {Q} =  
+ (\<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s')\<longrightarrow>  
+  (\<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>  
+  Q Y' s' Z \<and> s'\<Colon>\<preceq>(G,L)))))"
+apply (unfold triple_valid2_def)
+apply (simp (no_asm) add: split_paired_All)
+apply blast
+done
+
+lemma triple_valid2_eq [rule_format (no_asm)]: 
+  "wf_prog G ==> triple_valid2 G = triple_valid G"
+apply (rule ext)
+apply (rule ext)
+apply (rule triple.induct)
+apply (simp (no_asm) add: triple_valid_def2 triple_valid2_def2)
+apply (rule iffI)
+apply  fast
+apply clarify
+apply (tactic "smp_tac 3 1")
+apply (case_tac "normal s")
+apply  clarsimp
+apply  (blast dest: evaln_eval eval_type_sound [THEN conjunct1])
+apply clarsimp
+done
+
+lemma ax_valids2_eq: "wf_prog G \<Longrightarrow> G,A|\<Turnstile>\<Colon>ts = G,A|\<Turnstile>ts"
+apply (unfold ax_valids_def ax_valids2_def)
+apply (force simp add: triple_valid2_eq)
+done
+
+lemma triple_valid2_Suc [rule_format (no_asm)]: "G\<Turnstile>Suc n\<Colon>t \<longrightarrow> G\<Turnstile>n\<Colon>t"
+apply (induct_tac "t")
+apply (subst triple_valid2_def2)
+apply (subst triple_valid2_def2)
+apply (fast intro: evaln_nonstrict_Suc)
+done
+
+lemma Methd_triple_valid2_0: "G\<Turnstile>0\<Colon>{Normal P} Methd C sig-\<succ> {Q}"
+apply (clarsimp elim!: evaln_elim_cases simp add: triple_valid2_def2)
+done
+
+lemma Methd_triple_valid2_SucI: 
+"\<lbrakk>G\<Turnstile>n\<Colon>{Normal P} body G C sig-\<succ>{Q}\<rbrakk> 
+  \<Longrightarrow> G\<Turnstile>Suc n\<Colon>{Normal P} Methd C sig-\<succ> {Q}"
+apply (simp (no_asm_use) add: triple_valid2_def2)
+apply (intro strip, tactic "smp_tac 3 1", clarify)
+apply (erule wt_elim_cases, erule evaln_elim_cases)
+apply (unfold body_def Let_def)
+apply clarsimp
+apply blast
+done
+
+lemma triples_valid2_Suc: 
+ "Ball ts (triple_valid2 G (Suc n)) \<Longrightarrow> Ball ts (triple_valid2 G n)"
+apply (fast intro: triple_valid2_Suc)
+done
+
+lemma "G|\<Turnstile>n:insert t A = (G\<Turnstile>n:t \<and> G|\<Turnstile>n:A)";
+oops
+
+
+section "soundness"
+
+lemma Methd_sound: 
+"\<lbrakk>G,A\<union>  {{P} Methd-\<succ> {Q} | ms}|\<Turnstile>\<Colon>{{P} body G-\<succ> {Q} | ms}\<rbrakk> \<Longrightarrow> 
+  G,A|\<Turnstile>\<Colon>{{P} Methd-\<succ> {Q} | ms}"
+apply (unfold ax_valids2_def mtriples_def)
+apply (rule allI)
+apply (induct_tac "n")
+apply  (clarify, tactic {* pair_tac "x" 1 *}, simp (no_asm))
+apply  (fast intro: Methd_triple_valid2_0)
+apply (clarify, tactic {* pair_tac "xa" 1 *}, simp (no_asm))
+apply (drule triples_valid2_Suc)
+apply (erule (1) notE impE)
+apply (drule_tac x = na in spec)
+apply (tactic {* auto_tac (claset() addSIs [thm "Methd_triple_valid2_SucI"],
+   simpset() addsimps [ball_Un] addloop ("split_all_tac", split_all_tac)) *})
+done
+
+
+lemma valids2_inductI: "\<forall>s t n Y' s'. G\<turnstile>s\<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> t = c \<longrightarrow>    
+  Ball A (triple_valid2 G n) \<longrightarrow> (\<forall>Y Z. P Y s Z \<longrightarrow>  
+  (\<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>  
+  Q Y' s' Z \<and> s'\<Colon>\<preceq>(G, L)))) \<Longrightarrow>  
+  G,A|\<Turnstile>\<Colon>{ {P} c\<succ> {Q}}"
+apply (simp (no_asm) add: ax_valids2_def triple_valid2_def2)
+apply clarsimp
+done
+
+ML_setup {*
+Delsimprocs [evaln_expr_proc,evaln_var_proc,evaln_exprs_proc,evaln_stmt_proc]
+*}
+
+lemma Loop_sound: "\<lbrakk>G,A|\<Turnstile>\<Colon>{ {P} e-\<succ> {P'}};  
+       G,A|\<Turnstile>\<Colon>{ {Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}}\<rbrakk> \<Longrightarrow>  
+       G,A|\<Turnstile>\<Colon>{ {P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}}"
+apply (rule valids2_inductI)
+apply ((rule allI)+, rule impI, tactic {* pair_tac "s" 1*}, tactic {* pair_tac "s'" 1*})
+apply (erule evaln.induct)
+apply  simp_all (* takes half a minute *)
+apply  clarify
+apply  (erule_tac V = "G,A|\<Turnstile>\<Colon>{ {?P'} .c. {?P}}" in thin_rl)
+apply  (simp_all (no_asm_use) add: ax_valids2_def triple_valid2_def2)
+apply  (tactic "smp_tac 1 1", tactic "smp_tac 3 1", force)
+apply clarify
+apply (rule wt_elim_cases, assumption)
+apply (tactic "smp_tac 1 1", tactic "smp_tac 1 1", tactic "smp_tac 3 1", 
+       tactic "smp_tac 2 1", tactic "smp_tac 1 1")
+apply (erule impE,simp (no_asm),blast)
+apply (simp add: imp_conjL split_tupled_all split_paired_All)
+apply (case_tac "the_Bool b")
+apply  clarsimp
+apply  (case_tac "a")
+apply (simp_all)
+apply clarsimp
+apply  (erule_tac V = "c = l\<bullet> While(e) c \<longrightarrow> ?P" in thin_rl)
+apply (blast intro: conforms_absorb)
+apply blast+
+done
+
+declare subst_Bool_def2 [simp del]
+lemma all_empty: "(!x. P) = P"
+by simp
+lemma sound_valid2_lemma: 
+"\<lbrakk>\<forall>v n. Ball A (triple_valid2 G n) \<longrightarrow> P v n; Ball A (triple_valid2 G n)\<rbrakk>
+ \<Longrightarrow>P v n"
+by blast
+ML {*
+val fullsimptac = full_simp_tac(simpset() delsimps [thm "all_empty"]);
+val sound_prepare_tac = EVERY'[REPEAT o thin_tac "?x \<in> ax_derivs G",
+ full_simp_tac (simpset()addsimps[thm "ax_valids2_def",thm "triple_valid2_def2",
+                                  thm "imp_conjL"] delsimps[thm "all_empty"]),
+ Clarify_tac];
+val sound_elim_tac = EVERY'[eresolve_tac (thms "evaln_elim_cases"), 
+        TRY o eresolve_tac (thms "wt_elim_cases"), fullsimptac, Clarify_tac];
+val sound_valid2_tac = REPEAT o FIRST'[smp_tac 1, 
+                  datac (thm "sound_valid2_lemma") 1];
+val sound_forw_hyp_tac = 
+ EVERY'[smp_tac 3 
+          ORELSE' EVERY'[dtac spec,dtac spec, dtac spec,etac impE, Fast_tac] 
+          ORELSE' EVERY'[dtac spec,dtac spec,etac impE, Fast_tac],
+        fullsimptac, 
+        smp_tac 2,TRY o smp_tac 1,
+        TRY o EVERY'[etac impE, TRY o rtac impI, 
+        atac ORELSE' (EVERY' [REPEAT o rtac exI,Blast_tac]),
+        fullsimptac, Clarify_tac, TRY o smp_tac 1]]
+*}
+(* ### rtac conjI,rtac HOL.refl *)
+lemma Call_sound: 
+"\<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}};
+  \<forall>a vs invC declC l. G,A|\<Turnstile>\<Colon>{ {(R a\<leftarrow>Vals vs \<and>.  
+   (\<lambda>s. declC = invocation_declclass 
+                    G mode (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> \<and>
+         invC = invocation_class mode (store s) a statT \<and>
+            l = locals (store s)) ;.  
+   init_lvars G declC \<lparr>name=mn,parTs=pTs\<rparr> mode a vs) \<and>.  
+   (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)}  
+   Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}}\<rbrakk> \<Longrightarrow>  
+  G,A|\<Turnstile>\<Colon>{ {Normal P} {statT,mode}e\<cdot>mn({pTs}ps)-\<succ> {S}}"
+apply (tactic "EVERY'[sound_prepare_tac, sound_elim_tac, sound_valid2_tac] 1")
+apply (rename_tac x1 s1 x2 s2 ab bb v vs m pTsa statDeclC)
+apply (tactic "smp_tac 6 1")
+apply (tactic "sound_forw_hyp_tac 1")
+apply (tactic "sound_forw_hyp_tac 1")
+apply (drule max_spec2mheads)
+apply (drule evaln_eval, drule (3) eval_ts)
+apply (drule evaln_eval, frule (3) evals_ts)
+apply (drule spec,erule impE,rule exI, blast)
+(* apply (drule spec,drule spec,drule spec,erule impE,rule exI,blast) *)
+apply (case_tac "if static m then x2 else (np a') x2")
+defer 1
+apply  (rename_tac x, subgoal_tac "(Some x, s2)\<Colon>\<preceq>(G, L)" (* used two times *))
+prefer 2 
+apply   (force split add: split_if_asm)
+apply  (simp del: if_raise_if_None)
+apply  (tactic "smp_tac 2 1")
+apply (simp only: init_lvars_def2 invmode_Static_eq)
+apply (clarsimp simp del: resTy_mthd)
+apply  (drule spec,erule swap,erule conforms_set_locals [OF _ lconf_empty])
+apply clarsimp
+apply (drule Null_staticD)
+apply (drule eval_gext', drule (1) conf_gext, frule (3) DynT_propI)
+apply (erule impE) apply blast
+apply (subgoal_tac 
+ "G\<turnstile>invmode (mhd (statDeclC,m)) e
+     \<rightarrow>invocation_class (invmode m e) s2 a' statT\<preceq>statT")
+defer   apply simp
+apply (drule (3) DynT_mheadsD,simp,simp)
+apply (clarify, drule wf_mdeclD1, clarify)
+apply (frule ty_expr_is_type) apply simp
+apply (subgoal_tac "invmode (mhd (statDeclC,m)) e = IntVir \<longrightarrow> a' \<noteq> Null")
+defer   apply simp
+apply (frule (2) wt_MethdI)
+apply clarify
+apply (drule (2) conforms_init_lvars)
+apply   (simp) 
+apply   (assumption)+
+apply   simp
+apply   (assumption)+
+apply   (rule impI) apply simp
+apply   simp
+apply   simp
+apply   (rule Ball_weaken)
+apply     assumption
+apply     (force simp add: is_acc_type_def)
+apply (tactic "smp_tac 2 1")
+apply simp
+apply (tactic "smp_tac 1 1")
+apply (erule_tac V = "?P \<longrightarrow> ?Q" in thin_rl) 
+apply (erule impE)
+apply   (rule exI)+ 
+apply   (subgoal_tac "is_static dm = (static m)") 
+prefer 2  apply (simp add: member_is_static_simp)
+apply   (simp only: )
+apply   (simp only: sig.simps)
+apply (force dest!: evaln_eval eval_gext' elim: conforms_return 
+             del: impCE simp add: init_lvars_def2)
+done
+
+lemma Init_sound: "\<lbrakk>wf_prog G; the (class G C) = c;  
+      G,A|\<Turnstile>\<Colon>{ {Normal ((P \<and>. Not \<circ> initd C) ;. supd (init_class_obj G C))}  
+             .(if C = Object then Skip else Init (super c)). {Q}};  
+  \<forall>l. G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars empty}  
+            .init c. {set_lvars l .; R}}\<rbrakk> \<Longrightarrow>  
+      G,A|\<Turnstile>\<Colon>{ {Normal (P \<and>. Not \<circ> initd C)} .Init C. {R}}"
+apply (tactic "EVERY'[sound_prepare_tac, sound_elim_tac,sound_valid2_tac] 1")
+apply (tactic {* instantiate_tac [("l24","\<lambda> n Y Z sa Y' s' L y a b aa ba ab bb. locals b")]*})
+apply (clarsimp simp add: split_paired_Ex)
+apply (drule spec, drule spec, drule spec, erule impE)
+apply  (erule_tac V = "All ?P" in thin_rl, fast)
+apply clarsimp
+apply (tactic "smp_tac 2 1", drule spec, erule impE, 
+       erule (3) conforms_init_class_obj)
+apply (drule (1) wf_prog_cdecl)
+apply (erule impE, rule exI,erule_tac V = "All ?P" in thin_rl,
+       force dest: wf_cdecl_supD split add: split_if simp add: is_acc_class_def)
+apply clarify
+apply (drule spec)
+apply (drule spec)
+apply (drule spec)
+apply  (erule impE)
+apply ( fast)
+apply (simp (no_asm_use) del: empty_def2)
+apply (tactic "smp_tac 2 1")
+apply (drule spec, erule impE, erule conforms_set_locals, rule lconf_empty)
+apply (erule impE,rule impI,rule exI, erule wf_cdecl_wt_init)
+apply clarsimp
+apply (erule (1) conforms_return, force dest: evaln_eval eval_gext')
+done
+
+lemma all_conjunct2: "\<forall>l. P' l \<and> P l \<Longrightarrow> \<forall>l. P l"
+by fast
+
+lemma all4_conjunct2: 
+  "\<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"
+by fast
+
+lemmas sound_lemmas = Init_sound Loop_sound Methd_sound
+
+lemma ax_sound2: "wf_prog G \<Longrightarrow> G,A|\<turnstile>ts \<Longrightarrow> G,A|\<Turnstile>\<Colon>ts"
+apply (erule ax_derivs.induct)
+prefer 20 (* Call *)
+apply (erule (1) Call_sound) apply simp apply force apply force 
+
+apply (tactic {* TRYALL (eresolve_tac (thms "sound_lemmas") THEN_ALL_NEW 
+    eresolve_tac [asm_rl, thm "all_conjunct2", thm "all4_conjunct2"]) *})
+
+apply(tactic "COND (has_fewer_prems(30+3)) (ALLGOALS sound_prepare_tac) no_tac")
+
+               (*empty*)
+apply        fast (* insert *)
+apply       fast (* asm *)
+(*apply    fast *) (* cut *)
+apply     fast (* weaken *)
+apply    (tactic "smp_tac 3 1", clarify, tactic "smp_tac 1 1", frule evaln_eval,
+(* conseq *)case_tac"fst s",clarsimp simp add: eval_type_sound [THEN conjunct1],
+clarsimp)
+apply   (simp (no_asm_use) add: type_ok_def, drule evaln_eval,fast) (* hazard *)
+apply  force (* Abrupt *)
+
+(* 25 subgoals *)
+apply (tactic {* ALLGOALS sound_elim_tac *})(* LVar, Lit, Super, Nil, Skip,Do *)
+apply (tactic {* ALLGOALS (asm_simp_tac (noAll_simpset() 
+                          delsimps [thm "all_empty"])) *})    (* Done *)
+(* for FVar *)
+apply (frule wf_ws_prog) 
+apply (frule ty_expr_is_type [THEN type_is_class, 
+                              THEN accfield_declC_is_class])
+apply (simp,simp,simp) 
+apply (frule_tac [4] wt_init_comp_ty) (* for NewA*)
+apply (tactic "ALLGOALS sound_valid2_tac")
+apply (tactic "TRYALL sound_forw_hyp_tac") (* Cast, Inst, Acc, Expr *)
+apply (tactic {* TRYALL (EVERY'[dtac spec, TRY o EVERY'[rotate_tac ~1, 
+  dtac spec], dtac conjunct2, smp_tac 1, 
+  TRY o dres_inst_tac [("P","P'")] (thm "subst_Bool_the_BoolI")]) *})
+apply (frule_tac [14] x = x1 in conforms_NormI)  (* for Fin *)
+
+
+(* 15 subgoals *)
+(* FVar *)
+apply (tactic "sound_forw_hyp_tac 1")
+apply (clarsimp simp add: fvar_def2 Let_def split add: split_if_asm)
+
+(* AVar *)
+(*
+apply (drule spec, drule spec, erule impE, fast)
+apply (simp)
+apply (tactic "smp_tac 2 1")
+apply (tactic "smp_tac 1 1")
+apply (erule impE)
+apply (rule impI)
+apply (rule exI)+
+apply blast
+apply (clarsimp simp add: avar_def2)
+*)
+apply (tactic "sound_forw_hyp_tac 1")
+apply (clarsimp simp add: avar_def2)
+
+(* NewC *)
+apply (clarsimp simp add: is_acc_class_def)
+apply (erule (1) halloc_conforms, simp, simp)
+
+(* NewA *)
+apply (tactic "sound_forw_hyp_tac 1")
+apply (rule conjI,blast)
+apply (erule (1) halloc_conforms, simp, simp, simp add: is_acc_type_def)
+
+(* Ass *)
+apply (tactic "sound_forw_hyp_tac 1")
+apply (case_tac "aa")
+prefer 2
+apply  clarsimp
+apply (drule evaln_eval)+
+apply (frule (3) eval_ts)
+apply clarsimp
+apply (frule (3) evar_ts [THEN conjunct2])
+apply (frule wf_ws_prog)
+apply (drule (2) conf_widen)
+apply (drule_tac "s1.0" = b in eval_gext')
+apply (clarsimp simp add: assign_conforms_def)
+
+(* Cond *)
+
+apply (tactic "smp_tac 3 1") apply (tactic "smp_tac 2 1") 
+apply (tactic "smp_tac 1 1") apply (erule impE) 
+apply (rule impI,rule exI) 
+apply (rule_tac x = "if the_Bool b then T1 else T2" in exI)
+apply (force split add: split_if)
+apply assumption
+
+(* Body *)
+apply (tactic "sound_forw_hyp_tac 1")
+apply (rule conforms_absorb,assumption)
+
+(* Lab *)
+apply (tactic "sound_forw_hyp_tac 1")
+apply (rule conforms_absorb,assumption)
+
+(* If *)
+apply (tactic "sound_forw_hyp_tac 1")
+apply (tactic "sound_forw_hyp_tac 1")
+apply (force split add: split_if)
+
+(* Throw *)
+apply (drule evaln_eval, drule (3) eval_ts)
+apply clarsimp
+apply (drule (3) Throw_lemma)
+apply clarsimp
+
+(* Try *)
+apply (frule (1) sxalloc_type_sound)
+apply (erule sxalloc_elim_cases2)
+apply  (tactic "smp_tac 3 1")
+apply  (clarsimp split add: option.split_asm)
+apply (clarsimp split add: option.split_asm)
+apply (tactic "smp_tac 1 1")
+apply (simp only: split add: split_if_asm)
+prefer 2
+apply  (tactic "smp_tac 3 1", erule_tac V = "All ?P" in thin_rl, clarsimp)
+apply (drule spec, erule_tac V = "All ?P" in thin_rl, drule spec, drule spec, 
+       erule impE, force)
+apply (frule (2) Try_lemma)
+apply clarsimp
+apply (fast elim!: conforms_deallocL)
+
+(* Fin *)
+apply (tactic "sound_forw_hyp_tac 1")
+apply (case_tac "x1", force)
+apply clarsimp
+apply (drule evaln_eval, drule (4) Fin_lemma)
+done
+
+
+
+declare subst_Bool_def2 [simp]
+
+theorem ax_sound: 
+ "wf_prog G \<Longrightarrow> G,(A::'a triple set)|\<turnstile>(ts::'a triple set) \<Longrightarrow> G,A|\<Turnstile>ts"
+apply (subst ax_valids2_eq [symmetric])
+apply  assumption
+apply (erule (1) ax_sound2)
+done
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Basis.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,370 @@
+(*  Title:      isabelle/Bali/Basis.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+
+*)
+header {* Definitions extending HOL as logical basis of Bali *}
+
+theory Basis = Main:
+
+ML_setup {*
+Unify.search_bound := 40;
+Unify.trace_bound  := 40;
+
+quick_and_dirty:=true;
+
+Pretty.setmargin 77;
+goals_limit:=2;
+*}
+(*print_depth 100;*)
+(*Syntax.ambiguity_level := 1;*)
+
+section "misc"
+
+declare same_fstI [intro!] (*### TO HOL/Wellfounded_Relations *)
+
+(* ###TO HOL/???.ML?? *)
+ML {*
+fun make_simproc name pat pred thm = Simplifier.mk_simproc name
+   [Thm.read_cterm (Thm.sign_of_thm thm) (pat, HOLogic.typeT)] 
+   (K (K (fn s => if pred s then None else Some (standard (mk_meta_eq thm)))))
+*}
+
+declare split_if_asm  [split] option.split [split] option.split_asm [split]
+ML {*
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
+*}
+declare if_weak_cong [cong del] option.weak_case_cong [cong del]
+declare length_Suc_conv [iff];
+
+(*###to be phased out *)
+ML {*
+bind_thm ("make_imp", rearrange_prems [1,0] mp)
+*}
+
+lemma Collect_split_eq: "{p. P (split f p)} = {(a,b). P (f a b)}"
+apply auto
+done
+
+lemma subset_insertD: 
+  "A <= insert x B ==> A <= B & x ~: A | (EX B'. A = insert x B' & B' <= B)"
+apply (case_tac "x:A")
+apply (rule disjI2)
+apply (rule_tac x = "A-{x}" in exI)
+apply fast+
+done
+
+syntax
+  "3" :: nat   ("3")
+  "4" :: nat   ("4")
+translations
+ "3" == "Suc 2"
+ "4" == "Suc 3"
+
+(*unused*)
+lemma range_bool_domain: "range f = {f True, f False}"
+apply auto
+apply (case_tac "xa")
+apply auto
+done
+
+(* context (theory "Transitive_Closure") *)
+lemma irrefl_tranclI': "r^-1 Int r^+ = {} ==> !x. (x, x) ~: r^+"
+apply (rule allI)
+apply (erule irrefl_tranclI)
+done
+
+lemma trancl_rtrancl_trancl:
+"\<lbrakk>(x,y)\<in>r^+; (y,z)\<in>r^*\<rbrakk> \<Longrightarrow> (x,z)\<in>r^+"
+by (auto dest: tranclD rtrancl_trans rtrancl_into_trancl2)
+
+lemma rtrancl_into_trancl3:
+"\<lbrakk>(a,b)\<in>r^*; a\<noteq>b\<rbrakk> \<Longrightarrow> (a,b)\<in>r^+"
+apply (drule rtranclD)
+apply auto
+done
+
+lemma rtrancl_into_rtrancl2: 
+  "\<lbrakk> (a, b) \<in>  r; (b, c) \<in> r^* \<rbrakk> \<Longrightarrow> (a, c) \<in>  r^*"
+by (auto intro: r_into_rtrancl rtrancl_trans)
+
+lemma triangle_lemma:
+ "\<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> 
+ \<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*"
+proof -
+  note converse_rtrancl_induct = converse_rtrancl_induct [consumes 1]
+  note converse_rtranclE = converse_rtranclE [consumes 1] 
+  assume unique: "\<And> a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b=c"
+  assume "(a,x)\<in>r\<^sup>*" 
+  then show "(a,y)\<in>r\<^sup>* \<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*"
+  proof (induct rule: converse_rtrancl_induct)
+    assume "(x,y)\<in>r\<^sup>*"
+    then show ?thesis 
+      by blast
+  next
+    fix a v
+    assume a_v_r: "(a, v) \<in> r" and
+          v_x_rt: "(v, x) \<in> r\<^sup>*" and
+          a_y_rt: "(a, y) \<in> r\<^sup>*"  and
+             hyp: "(v, y) \<in> r\<^sup>* \<Longrightarrow> (x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"
+    from a_y_rt 
+    show "(x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"
+    proof (cases rule: converse_rtranclE)
+      assume "a=y"
+      with a_v_r v_x_rt have "(y,x) \<in> r\<^sup>*"
+	by (auto intro: r_into_rtrancl rtrancl_trans)
+      then show ?thesis 
+	by blast
+    next
+      fix w 
+      assume a_w_r: "(a, w) \<in> r" and
+            w_y_rt: "(w, y) \<in> r\<^sup>*"
+      from a_v_r a_w_r unique 
+      have "v=w" 
+	by auto
+      with w_y_rt hyp 
+      show ?thesis
+	by blast
+    qed
+  qed
+qed
+
+
+lemma rtrancl_cases [consumes 1, case_names Refl Trancl]:
+ "\<lbrakk>(a,b)\<in>r\<^sup>*;  a = b \<Longrightarrow> P; (a,b)\<in>r\<^sup>+ \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+apply (erule rtranclE)
+apply (auto dest: rtrancl_into_trancl1)
+done
+
+(* ### To Transitive_Closure *)
+theorems converse_rtrancl_induct 
+ = converse_rtrancl_induct [consumes 1,case_names Id Step]
+
+theorems converse_trancl_induct 
+         = converse_trancl_induct [consumes 1,case_names Single Step]
+
+(* context (theory "Set") *)
+lemma Ball_weaken:"\<lbrakk>Ball s P;\<And> x. P x\<longrightarrow>Q x\<rbrakk>\<Longrightarrow>Ball s Q"
+by auto
+
+(* context (theory "Finite") *)
+lemma finite_SetCompr2: "[| finite (Collect P); !y. P y --> finite (range (f y)) |] ==>  
+  finite {f y x |x y. P y}"
+apply (subgoal_tac "{f y x |x y. P y} = UNION (Collect P) (%y. range (f y))")
+prefer 2 apply  fast
+apply (erule ssubst)
+apply (erule finite_UN_I)
+apply fast
+done
+
+
+(* ### TO theory "List" *)
+lemma list_all2_trans: "\<forall> a b c. P1 a b \<longrightarrow> P2 b c \<longrightarrow> P3 a c \<Longrightarrow>
+ \<forall>xs2 xs3. list_all2 P1 xs1 xs2 \<longrightarrow> list_all2 P2 xs2 xs3 \<longrightarrow> list_all2 P3 xs1 xs3"
+apply (induct_tac "xs1")
+apply simp
+apply (rule allI)
+apply (induct_tac "xs2")
+apply simp
+apply (rule allI)
+apply (induct_tac "xs3")
+apply auto
+done
+
+
+section "pairs"
+
+lemma surjective_pairing5: "p = (fst p, fst (snd p), fst (snd (snd p)), fst (snd (snd (snd p))), 
+  snd (snd (snd (snd p))))"
+apply auto
+done
+
+lemma fst_splitE [elim!]: 
+"[| fst s' = x';  !!x s. [| s' = (x,s);  x = x' |] ==> Q |] ==> Q"
+apply (cut_tac p = "s'" in surjective_pairing)
+apply auto
+done
+
+lemma fst_in_set_lemma [rule_format (no_asm)]: "(x, y) : set l --> x : fst ` set l"
+apply (induct_tac "l")
+apply  auto
+done
+
+
+section "quantifiers"
+
+(*###to be phased out *)
+ML {* 
+fun noAll_simpset () = simpset() setmksimps 
+	mksimps (filter (fn (x,_) => x<>"All") mksimps_pairs)
+*}
+
+lemma All_Ex_refl_eq2 [simp]: 
+ "(!x. (? b. x = f b & Q b) \<longrightarrow> P x) = (!b. Q b --> P (f b))"
+apply auto
+done
+
+lemma ex_ex_miniscope1 [simp]:
+  "(EX w v. P w v & Q v) = (EX v. (EX w. P w v) & Q v)"
+apply auto
+done
+
+lemma ex_miniscope2 [simp]:
+  "(EX v. P v & Q & R v) = (Q & (EX v. P v & R v))" 
+apply auto
+done
+
+lemma ex_reorder31: "(\<exists>z x y. P x y z) = (\<exists>x y z. P x y z)"
+apply auto
+done
+
+lemma All_Ex_refl_eq1 [simp]: "(!x. (? b. x = f b) --> P x) = (!b. P (f b))"
+apply auto
+done
+
+
+section "sums"
+
+hide const In0 In1
+
+syntax
+  fun_sum :: "('a => 'c) => ('b => 'c) => (('a+'b) => 'c)" (infixr "'(+')"80)
+translations
+ "fun_sum" == "sum_case"
+
+consts    the_Inl  :: "'a + 'b \<Rightarrow> 'a"
+          the_Inr  :: "'a + 'b \<Rightarrow> 'b"
+primrec  "the_Inl (Inl a) = a"
+primrec  "the_Inr (Inr b) = b"
+
+datatype ('a, 'b, 'c) sum3 = In1 'a | In2 'b | In3 'c
+
+consts    the_In1  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'a"
+          the_In2  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'b"
+          the_In3  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'c"
+primrec  "the_In1 (In1 a) = a"
+primrec  "the_In2 (In2 b) = b"
+primrec  "the_In3 (In3 c) = c"
+
+syntax
+	 In1l	:: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
+	 In1r	:: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
+translations
+	"In1l e" == "In1 (Inl e)"
+	"In1r c" == "In1 (Inr c)"
+
+ML {*
+fun sum3_instantiate thm = map (fn s => simplify(simpset()delsimps[not_None_eq])
+ (read_instantiate [("t","In"^s^" ?x")] thm)) ["1l","2","3","1r"]
+*}
+(* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *)
+
+translations
+  "option"<= (type) "Option.option"
+  "list"  <= (type) "List.list"
+  "sum3"  <= (type) "Basis.sum3"
+
+
+section "quantifiers for option type"
+
+syntax
+  Oall :: "[pttrn, 'a option, bool] => bool"   ("(3! _:_:/ _)" [0,0,10] 10)
+  Oex  :: "[pttrn, 'a option, bool] => bool"   ("(3? _:_:/ _)" [0,0,10] 10)
+
+syntax (symbols)
+  Oall :: "[pttrn, 'a option, bool] => bool"   ("(3\<forall>_\<in>_:/ _)"  [0,0,10] 10)
+  Oex  :: "[pttrn, 'a option, bool] => bool"   ("(3\<exists>_\<in>_:/ _)"  [0,0,10] 10)
+
+translations
+  "! x:A: P"    == "! x:o2s A. P"
+  "? x:A: P"    == "? x:o2s A. P"
+
+
+section "unique association lists"
+
+constdefs
+  unique   :: "('a \<times> 'b) list \<Rightarrow> bool"
+ "unique \<equiv> nodups \<circ> map fst"
+
+lemma uniqueD [rule_format (no_asm)]: 
+"unique l--> (!x y. (x,y):set l --> (!x' y'. (x',y'):set l --> x=x'-->  y=y'))"
+apply (unfold unique_def o_def)
+apply (induct_tac "l")
+apply  (auto dest: fst_in_set_lemma)
+done
+
+lemma unique_Nil [simp]: "unique []"
+apply (unfold unique_def)
+apply (simp (no_asm))
+done
+
+lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))"
+apply (unfold unique_def)
+apply  (auto dest: fst_in_set_lemma)
+done
+
+lemmas unique_ConsI = conjI [THEN unique_Cons [THEN iffD2], standard]
+
+lemma unique_single [simp]: "!!p. unique [p]"
+apply auto
+done
+
+lemma unique_ConsD: "unique (x#xs) ==> unique xs"
+apply (simp add: unique_def)
+done
+
+lemma unique_append [rule_format (no_asm)]: "unique l' ==> unique l -->  
+  (!(x,y):set l. !(x',y'):set l'. x' ~= x) --> unique (l @ l')"
+apply (induct_tac "l")
+apply  (auto dest: fst_in_set_lemma)
+done
+
+lemma unique_map_inj [rule_format (no_asm)]: "unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)"
+apply (induct_tac "l")
+apply  (auto dest: fst_in_set_lemma simp add: inj_eq)
+done
+
+lemma map_of_SomeI [rule_format (no_asm)]: "unique l --> (k, x) : set l --> map_of l k = Some x"
+apply (induct_tac "l")
+apply auto
+done
+
+
+section "list patterns"
+
+consts
+  lsplit         :: "[['a, 'a list] => 'b, 'a list] => 'b"
+defs
+  lsplit_def:    "lsplit == %f l. f (hd l) (tl l)"
+(*  list patterns -- extends pre-defined type "pttrn" used in abstractions *)
+syntax
+  "_lpttrn"    :: "[pttrn,pttrn] => pttrn"     ("_#/_" [901,900] 900)
+translations
+  "%y#x#xs. b"  == "lsplit (%y x#xs. b)"
+  "%x#xs  . b"  == "lsplit (%x xs  . b)"
+
+lemma lsplit [simp]: "lsplit c (x#xs) = c x xs"
+apply (unfold lsplit_def)
+apply (simp (no_asm))
+done
+
+lemma lsplit2 [simp]: "lsplit P (x#xs) y z = P x xs y z"
+apply (unfold lsplit_def)
+apply simp
+done 
+
+
+section "dummy pattern for quantifiers, let, etc."
+
+syntax
+  "@dummy_pat"   :: pttrn    ("'_")
+
+parse_translation {*
+let fun dummy_pat_tr [] = Free ("_",dummyT)
+  | dummy_pat_tr ts = raise TERM ("dummy_pat_tr", ts);
+in [("@dummy_pat", dummy_pat_tr)] 
+end
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Conform.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,447 @@
+(*  Title:      isabelle/Bali/Conform.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+
+header {* Conformance notions for the type soundness proof for Java *}
+
+theory Conform = State:
+
+text {*
+design issues:
+\begin{itemize}
+\item lconf allows for (arbitrary) inaccessible values
+\item ''conforms'' does not directly imply that the dynamic types of all 
+      objects on the heap are indeed existing classes. Yet this can be 
+      inferred for all referenced objs.
+\end{itemize}
+*}
+
+types	env_ = "prog \<times> (lname, ty) table" (* same as env of WellType.thy *)
+
+
+section "extension of global store"
+
+constdefs
+
+  gext    :: "st \<Rightarrow> st \<Rightarrow> bool"                ("_\<le>|_"       [71,71]   70)
+   "s\<le>|s' \<equiv> \<forall>r. \<forall>obj\<in>globs s r: \<exists>obj'\<in>globs s' r: tag obj'= tag obj"
+
+lemma gext_objD: 
+"\<lbrakk>s\<le>|s'; globs s r = Some obj\<rbrakk> 
+\<Longrightarrow> \<exists>obj'. globs s' r = Some obj' \<and> tag obj' = tag obj"
+apply (simp only: gext_def)
+by force
+
+lemma rev_gext_objD: 
+"\<lbrakk>globs s r = Some obj; s\<le>|s'\<rbrakk> 
+ \<Longrightarrow> \<exists>obj'. globs s' r = Some obj' \<and> tag obj' = tag obj"
+by (auto elim: gext_objD)
+
+lemma init_class_obj_inited: 
+   "init_class_obj G C s1\<le>|s2 \<Longrightarrow> inited C (globs s2)"
+apply (unfold inited_def init_obj_def)
+apply (auto dest!: gext_objD)
+done
+
+lemma gext_refl [intro!, simp]: "s\<le>|s"
+apply (unfold gext_def)
+apply (fast del: fst_splitE)
+done
+
+lemma gext_gupd [simp, elim!]: "\<And>s. globs s r = None \<Longrightarrow> s\<le>|gupd(r\<mapsto>x)s"
+by (auto simp: gext_def)
+
+lemma gext_new [simp, elim!]: "\<And>s. globs s r = None \<Longrightarrow> s\<le>|init_obj G oi r s"
+apply (simp only: init_obj_def)
+apply (erule_tac gext_gupd)
+done
+
+lemma gext_trans [elim]: "\<And>X. \<lbrakk>s\<le>|s'; s'\<le>|s''\<rbrakk> \<Longrightarrow> s\<le>|s''" 
+by (force simp: gext_def)
+
+lemma gext_upd_gobj [intro!]: "s\<le>|upd_gobj r n v s"
+apply (simp only: gext_def)
+apply auto
+apply (case_tac "ra = r")
+apply auto
+apply (case_tac "globs s r = None")
+apply auto
+done
+
+lemma gext_cong1 [simp]: "set_locals l s1\<le>|s2 = s1\<le>|s2"
+by (auto simp: gext_def)
+
+lemma gext_cong2 [simp]: "s1\<le>|set_locals l s2 = s1\<le>|s2"
+by (auto simp: gext_def)
+
+
+lemma gext_lupd1 [simp]: "lupd(vn\<mapsto>v)s1\<le>|s2 = s1\<le>|s2"
+by (auto simp: gext_def)
+
+lemma gext_lupd2 [simp]: "s1\<le>|lupd(vn\<mapsto>v)s2 = s1\<le>|s2"
+by (auto simp: gext_def)
+
+
+lemma inited_gext: "\<lbrakk>inited C (globs s); s\<le>|s'\<rbrakk> \<Longrightarrow> inited C (globs s')"
+apply (unfold inited_def)
+apply (auto dest: gext_objD)
+done
+
+
+section "value conformance"
+
+constdefs
+
+  conf  :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool"    ("_,_\<turnstile>_\<Colon>\<preceq>_"   [71,71,71,71] 70)
+	   "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"
+
+lemma conf_cong [simp]: "G,set_locals l s\<turnstile>v\<Colon>\<preceq>T = G,s\<turnstile>v\<Colon>\<preceq>T"
+by (auto simp: conf_def)
+
+lemma conf_lupd [simp]: "G,lupd(vn\<mapsto>va)s\<turnstile>v\<Colon>\<preceq>T = G,s\<turnstile>v\<Colon>\<preceq>T"
+by (auto simp: conf_def)
+
+lemma conf_PrimT [simp]: "\<forall>dt. typeof dt v = Some (PrimT t) \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>PrimT t"
+apply (simp add: conf_def)
+done
+
+lemma conf_litval [rule_format (no_asm)]: 
+  "typeof (\<lambda>a. None) v = Some T \<longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T"
+apply (unfold conf_def)
+apply (rule val.induct)
+apply auto
+done
+
+lemma conf_Null [simp]: "G,s\<turnstile>Null\<Colon>\<preceq>T = G\<turnstile>NT\<preceq>T"
+by (simp add: conf_def)
+
+lemma conf_Addr: 
+  "G,s\<turnstile>Addr a\<Colon>\<preceq>T = (\<exists>obj. heap s a = Some obj \<and> G\<turnstile>obj_ty obj\<preceq>T)"
+by (auto simp: conf_def)
+
+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"
+apply (rule conf_Addr [THEN iffD2])
+by fast
+
+lemma defval_conf [rule_format (no_asm), elim]: 
+  "is_type G T \<longrightarrow> G,s\<turnstile>default_val T\<Colon>\<preceq>T"
+apply (unfold conf_def)
+apply (induct "T")
+apply (auto intro: prim_ty.induct)
+done
+
+lemma conf_widen [rule_format (no_asm), elim]: 
+  "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'"
+apply (unfold conf_def)
+apply (rule val.induct)
+apply (auto elim: ws_widen_trans)
+done
+
+lemma conf_gext [rule_format (no_asm), elim]: 
+  "G,s\<turnstile>v\<Colon>\<preceq>T \<longrightarrow> s\<le>|s' \<longrightarrow> G,s'\<turnstile>v\<Colon>\<preceq>T"
+apply (unfold gext_def conf_def)
+apply (rule val.induct)
+apply force+
+done
+
+
+lemma conf_list_widen [rule_format (no_asm)]: 
+"ws_prog G \<Longrightarrow>  
+  \<forall>Ts Ts'. list_all2 (conf G s) vs Ts 
+           \<longrightarrow>   G\<turnstile>Ts[\<preceq>] Ts' \<longrightarrow> list_all2 (conf G s) vs Ts'"
+apply (unfold widens_def)
+apply (rule list_all2_trans)
+apply auto
+done
+
+lemma conf_RefTD [rule_format (no_asm)]: 
+ "G,s\<turnstile>a'\<Colon>\<preceq>RefT T 
+  \<longrightarrow> a' = Null \<or> (\<exists>a obj T'. a' = Addr a \<and> heap s a = Some obj \<and>  
+                    obj_ty obj = T' \<and> G\<turnstile>T'\<preceq>RefT T)"
+apply (unfold conf_def)
+apply (induct_tac "a'")
+apply (auto dest: widen_PrimT)
+done
+
+
+section "value list conformance"
+
+constdefs
+
+  lconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool"
+                                                ("_,_\<turnstile>_[\<Colon>\<preceq>]_" [71,71,71,71] 70)
+           "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"
+
+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"
+by (force simp: lconf_def)
+
+
+lemma lconf_cong [simp]: "\<And>s. G,set_locals x s\<turnstile>l[\<Colon>\<preceq>]L = G,s\<turnstile>l[\<Colon>\<preceq>]L"
+by (auto simp: lconf_def)
+
+lemma lconf_lupd [simp]: "G,lupd(vn\<mapsto>v)s\<turnstile>l[\<Colon>\<preceq>]L = G,s\<turnstile>l[\<Colon>\<preceq>]L"
+by (auto simp: lconf_def)
+
+(* unused *)
+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"
+by (auto simp: lconf_def)
+
+lemma lconf_upd: "\<lbrakk>G,s\<turnstile>l[\<Colon>\<preceq>]L; G,s\<turnstile>v\<Colon>\<preceq>T; L vn = Some T\<rbrakk> \<Longrightarrow>  
+  G,s\<turnstile>l(vn\<mapsto>v)[\<Colon>\<preceq>]L"
+by (auto simp: lconf_def)
+
+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)"
+by (auto simp: lconf_def)
+
+lemma lconf_map_sum [simp]: 
+ "G,s\<turnstile>l1 (+) l2[\<Colon>\<preceq>]L1 (+) L2 = (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>l2[\<Colon>\<preceq>]L2)"
+apply (unfold lconf_def)
+apply safe
+apply (case_tac [3] "n")
+apply (force split add: sum.split)+
+done
+
+lemma lconf_ext_list [rule_format (no_asm)]: "
+ \<And>X. \<lbrakk>G,s\<turnstile>l[\<Colon>\<preceq>]L\<rbrakk> \<Longrightarrow> 
+      \<forall>vs Ts. nodups vns \<longrightarrow> length Ts = length vns 
+      \<longrightarrow> list_all2 (conf G s) vs Ts \<longrightarrow> G,s\<turnstile>l(vns[\<mapsto>]vs)[\<Colon>\<preceq>]L(vns[\<mapsto>]Ts)"
+apply (unfold lconf_def)
+apply (induct_tac "vns")
+apply  clarsimp
+apply clarsimp
+apply (frule list_all2_lengthD)
+apply clarsimp
+done
+
+
+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"
+apply (simp only: lconf_def)
+apply safe
+apply (drule spec)
+apply (drule ospec)
+apply auto
+done 
+
+
+lemma lconf_gext [elim]: "\<lbrakk>G,s\<turnstile>l[\<Colon>\<preceq>]L; s\<le>|s'\<rbrakk> \<Longrightarrow> G,s'\<turnstile>l[\<Colon>\<preceq>]L"
+apply (simp only: lconf_def)
+apply fast
+done
+
+lemma lconf_empty [simp, intro!]: "G,s\<turnstile>vs[\<Colon>\<preceq>]empty"
+apply (unfold lconf_def)
+apply force
+done
+
+lemma lconf_init_vals [intro!]: 
+	" \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<Colon>\<preceq>]fs"
+apply (unfold lconf_def)
+apply force
+done
+
+
+section "object conformance"
+
+constdefs
+
+  oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool"  ("_,_\<turnstile>_\<Colon>\<preceq>\<surd>_"  [71,71,71,71] 70)
+	   "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r \<equiv> G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and> 
+                           (case r of 
+		              Heap a \<Rightarrow> is_type G (obj_ty obj) 
+                            | Stat C \<Rightarrow> True)"
+(*
+lemma oconf_def2:  "G,s\<turnstile>\<lparr>tag=oi,values=fs\<rparr>\<Colon>\<preceq>\<surd>r =  
+  (G,s\<turnstile>fs[\<Colon>\<preceq>]var_tys G oi r \<and> 
+  (case r of Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=fs\<rparr>) | Stat C \<Rightarrow> True))"
+by (simp add: oconf_def Let_def)
+*)
+(*
+lemma oconf_def2:  "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r =  
+  (G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and> 
+  (case r of Heap a \<Rightarrow> is_type G (obj_ty obj) | Stat C \<Rightarrow> True))"
+by (simp add: oconf_def Let_def)
+*)
+lemma oconf_is_type: "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>Heap a \<Longrightarrow> is_type G (obj_ty obj)"
+by (auto simp: oconf_def Let_def)
+
+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"
+by (simp add: oconf_def) 
+
+lemma oconf_cong [simp]: "G,set_locals l s\<turnstile>obj\<Colon>\<preceq>\<surd>r = G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r"
+by (auto simp: oconf_def Let_def)
+
+lemma oconf_init_obj_lemma: 
+"\<lbrakk>\<And>C c. class G C = Some c \<Longrightarrow> unique (DeclConcepts.fields G C);  
+  \<And>C c f fld. \<lbrakk>class G C = Some c; 
+                table_of (DeclConcepts.fields G C) f = Some fld \<rbrakk> 
+            \<Longrightarrow> is_type G (type fld);  
+  (case r of 
+     Heap a \<Rightarrow> is_type G (obj_ty obj) 
+  | Stat C \<Rightarrow> is_class G C)
+\<rbrakk> \<Longrightarrow>  G,s\<turnstile>obj \<lparr>values:=init_vals (var_tys G (tag obj) r)\<rparr>\<Colon>\<preceq>\<surd>r"
+apply (auto simp add: oconf_def)
+apply (drule_tac var_tys_Some_eq [THEN iffD1]) 
+defer
+apply (subst obj_ty_cong)
+apply(auto dest!: fields_table_SomeD obj_ty_CInst1 obj_ty_Arr1
+           split add: sum.split_asm obj_tag.split_asm)
+done
+
+(*
+lemma oconf_init_obj_lemma: 
+"\<lbrakk>\<And>C c. class G C = Some c \<Longrightarrow> unique (fields G C);  
+  \<And>C c f fld. \<lbrakk>class G C = Some c; table_of (fields G C) f = Some fld \<rbrakk> 
+            \<Longrightarrow> is_type G (type fld);  
+  (case r of 
+     Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=fs\<rparr>) 
+  | Stat C \<Rightarrow> is_class G C)
+\<rbrakk> \<Longrightarrow>  G,s\<turnstile>\<lparr>tag=oi, values=init_vals (var_tys G oi r)\<rparr>\<Colon>\<preceq>\<surd>r"
+apply (auto simp add: oconf_def)
+apply (drule_tac var_tys_Some_eq [THEN iffD1]) 
+defer
+apply (subst obj_ty_eq)
+apply(auto dest!: fields_table_SomeD split add: sum.split_asm obj_tag.split_asm)
+done
+*)
+
+
+section "state conformance"
+
+constdefs
+
+  conforms :: "state \<Rightarrow> env_ \<Rightarrow> bool"          (     "_\<Colon>\<preceq>_"   [71,71]      70)
+   "xs\<Colon>\<preceq>E \<equiv> let (G, L) = E; s = snd xs; l = locals s in
+      (\<forall>r. \<forall>obj\<in>globs s r:           G,s\<turnstile>obj   \<Colon>\<preceq>\<surd>r) \<and>
+                  \<spacespace>                   G,s\<turnstile>l    [\<Colon>\<preceq>]L\<spacespace> \<and>
+      (\<forall>a. fst xs=Some(Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable))"
+
+section "conforms"
+
+lemma conforms_globsD: 
+"\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); globs s r = Some obj\<rbrakk> \<Longrightarrow> G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r"
+by (auto simp: conforms_def Let_def)
+
+lemma conforms_localD: "(x, s)\<Colon>\<preceq>(G, L) \<Longrightarrow> G,s\<turnstile>locals s[\<Colon>\<preceq>]L"
+by (auto simp: conforms_def Let_def)
+
+lemma conforms_XcptLocD: "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); x = Some (Xcpt (Loc a))\<rbrakk> \<Longrightarrow>  
+	  G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)"
+by (auto simp: conforms_def Let_def)
+
+lemma conforms_RefTD: 
+ "\<lbrakk>G,s\<turnstile>a'\<Colon>\<preceq>RefT t; a' \<noteq> Null; (x,s) \<Colon>\<preceq>(G, L)\<rbrakk> \<Longrightarrow>  
+   \<exists>a obj. a' = Addr a \<and> globs s (Inl a) = Some obj \<and>  
+   G\<turnstile>obj_ty obj\<preceq>RefT t \<and> is_type G (obj_ty obj)"
+apply (drule_tac conf_RefTD)
+apply clarsimp
+apply (rule conforms_globsD [THEN oconf_is_type])
+apply auto
+done
+
+lemma conforms_Jump [iff]:
+  "((Some (Jump j), s)\<Colon>\<preceq>(G, L)) = (Norm s\<Colon>\<preceq>(G, L))"
+by (auto simp: conforms_def)
+
+lemma conforms_StdXcpt [iff]: 
+  "((Some (Xcpt (Std xn)), s)\<Colon>\<preceq>(G, L)) = (Norm s\<Colon>\<preceq>(G, L))"
+by (auto simp: conforms_def)
+
+lemma conforms_raise_if [iff]: 
+  "((raise_if c xn x, s)\<Colon>\<preceq>(G, L)) = ((x, s)\<Colon>\<preceq>(G, L))"
+by (auto simp: abrupt_if_def)
+
+
+lemma conforms_NormI: "(x, s)\<Colon>\<preceq>(G, L) \<Longrightarrow> Norm s\<Colon>\<preceq>(G, L)"
+by (auto simp: conforms_def Let_def)
+
+
+lemma conforms_absorb [rule_format]:
+  "(a, b)\<Colon>\<preceq>(G, L) \<longrightarrow> (absorb j a, b)\<Colon>\<preceq>(G, L)"
+apply (rule impI)
+apply ( case_tac a)
+apply (case_tac "absorb j a")
+apply auto
+apply (case_tac "absorb j (Some a)",auto)
+apply (erule conforms_NormI)
+done
+
+lemma conformsI: "\<lbrakk>\<forall>r. \<forall>obj\<in>globs s r: G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r;  
+     G,s\<turnstile>locals s[\<Colon>\<preceq>]L;  
+     \<forall>a. x = Some (Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)\<rbrakk> \<Longrightarrow> 
+  (x, s)\<Colon>\<preceq>(G, L)"
+by (auto simp: conforms_def Let_def)
+
+lemma conforms_xconf: "\<lbrakk>(x, s)\<Colon>\<preceq>(G,L);   
+ \<forall>a. x' = Some (Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)\<rbrakk> \<Longrightarrow> 
+ (x',s)\<Colon>\<preceq>(G,L)"
+by (fast intro: conformsI elim: conforms_globsD conforms_localD)
+
+lemma conforms_lupd: 
+ "\<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)"
+by (force intro: conformsI lconf_upd dest: conforms_globsD conforms_localD 
+                                           conforms_XcptLocD simp: oconf_def)
+
+
+lemmas conforms_allocL_aux = conforms_localD [THEN lconf_ext]
+
+lemma conforms_allocL: 
+  "\<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))"
+by (force intro: conformsI dest: conforms_globsD 
+          elim: conforms_XcptLocD conforms_allocL_aux simp: oconf_def)
+
+lemmas conforms_deallocL_aux = conforms_localD [THEN lconf_deallocL]
+
+lemma conforms_deallocL: "\<And>s.\<lbrakk>s\<Colon>\<preceq>(G, L(vn\<mapsto>T)); L vn = None\<rbrakk> \<Longrightarrow> s\<Colon>\<preceq>(G,L)"
+by (fast intro: conformsI dest: conforms_globsD 
+         elim: conforms_XcptLocD conforms_deallocL_aux)
+
+lemma conforms_gext: "\<lbrakk>(x, s)\<Colon>\<preceq>(G,L); s\<le>|s';  
+  \<forall>r. \<forall>obj\<in>globs s' r: G,s'\<turnstile>obj\<Colon>\<preceq>\<surd>r;  
+   locals s'=locals s\<rbrakk> \<Longrightarrow> (x,s')\<Colon>\<preceq>(G,L)"
+by (force intro!: conformsI dest: conforms_localD conforms_XcptLocD)
+
+
+lemma conforms_xgext: 
+  "\<lbrakk>(x ,s)\<Colon>\<preceq>(G,L); (x', s')\<Colon>\<preceq>(G, L); s'\<le>|s\<rbrakk> \<Longrightarrow> (x',s)\<Colon>\<preceq>(G,L)"
+apply (erule_tac conforms_xconf)
+apply (fast dest: conforms_XcptLocD)
+done
+
+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> 
+\<Longrightarrow>  (x, gupd(r\<mapsto>obj)s)\<Colon>\<preceq>(G, L)"
+apply (rule conforms_gext)
+apply    auto
+apply (force dest: conforms_globsD simp add: oconf_def)+
+done
+
+lemma conforms_upd_gobj: "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L); globs s r = Some obj; 
+  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)"
+apply (rule conforms_gext)
+apply auto
+apply (drule (1) conforms_globsD)
+apply (simp add: oconf_def)
+apply safe
+apply (rule lconf_upd)
+apply auto
+apply (simp only: obj_ty_cong) 
+apply (force dest: conforms_globsD intro!: lconf_upd 
+       simp add: oconf_def cong del: sum.weak_case_cong)
+done
+
+lemma conforms_set_locals: 
+  "\<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)"
+apply (auto intro!: conformsI dest: conforms_globsD 
+            elim!: conforms_XcptLocD simp add: oconf_def)
+done
+
+lemma conforms_return: "\<And>s'. \<lbrakk>(x,s)\<Colon>\<preceq>(G, L); (x',s')\<Colon>\<preceq>(G, L'); s\<le>|s'\<rbrakk> \<Longrightarrow>  
+  (x',set_locals (locals s) s')\<Colon>\<preceq>(G, L)"
+apply (rule conforms_xconf)
+prefer 2 apply (force dest: conforms_XcptLocD)
+apply (erule conforms_gext)
+apply (force dest: conforms_globsD)+
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Decl.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,867 @@
+(*  Title:      isabelle/Bali/Decl.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* Field, method, interface, and class declarations, whole Java programs
+*}
+
+(** order is significant, because of clash for "var" **)
+theory Decl = Term + Table:
+
+text {*
+improvements:
+\begin{itemize}
+\item clarification and correction of some aspects of the package/access concept
+  (Also submitted as bug report to the Java Bug Database:
+   Bug Id: 4485402 and Bug Id: 4493343 
+   http://developer.java.sun.com/developer/bugParade/index.jshtml
+  )
+\end{itemize}
+simplifications:
+\begin{itemize}
+\item the only field and method modifiers are static and the access modifiers
+\item no constructors, which may be simulated by new + suitable methods
+\item there is just one global initializer per class, which can simulate all 
+      others
+
+\item no throws clause
+\item a void method is replaced by one that returns Unit (of dummy type Void)
+
+\item no interface fields
+
+\item every class has an explicit superclass (unused for Object)
+\item the (standard) methods of Object and of standard exceptions are not 
+      specified
+
+\item no main method
+\end{itemize}
+*}
+
+subsection {* Modifier*}
+
+subsubsection {* Access modifier *}
+
+datatype acc_modi (* access modifier *)
+         = Private | Package | Protected | Public 
+
+text {* 
+We can define a linear order for the access modifiers. With Private yielding the
+most restrictive access and public the most liberal access policy:
+  Private < Package < Protected < Public
+*}
+ 
+instance acc_modi:: ord
+by intro_classes
+
+defs (overloaded)
+less_acc_def: 
+ "a < (b::acc_modi) 
+      \<equiv> (case a of
+             Private    \<Rightarrow> (b=Package \<or> b=Protected \<or> b=Public)
+           | Package    \<Rightarrow> (b=Protected \<or> b=Public)
+           | Protected  \<Rightarrow> (b=Public)
+           | Public     \<Rightarrow> False)"
+le_acc_def:
+ "a \<le> (b::acc_modi) \<equiv> (a = b) \<or> (a < b)"
+
+instance acc_modi:: order
+proof (intro_classes)
+  fix x y z::acc_modi
+  {
+  show "x \<le> x"               \<spacespace>\<spacespace>    -- reflexivity
+    by (auto simp add: le_acc_def)
+  next
+  assume "x \<le> y" "y \<le> z"           -- transitivity 
+  thus "x \<le> z"
+    by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split)
+  next
+  assume "x \<le> y" "y \<le> x"           -- antisymmetry
+  thus "x = y"
+  proof -
+    have "\<forall> x y. x < (y::acc_modi) \<and> y < x \<longrightarrow> False"
+      by (auto simp add: less_acc_def split add: acc_modi.split)
+    with prems show ?thesis
+      by  (auto simp add: le_acc_def)
+  qed
+  next
+  show "(x < y) = (x \<le> y \<and> x \<noteq> y)"
+    by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split) 
+  }
+qed
+  
+instance acc_modi:: linorder
+proof intro_classes
+  fix x y:: acc_modi
+  show  "x \<le> y \<or> y \<le> x"   
+  by (auto simp add: less_acc_def le_acc_def split add: acc_modi.split)
+qed
+
+lemma acc_modi_top [simp]: "Public \<le> a \<Longrightarrow> a = Public"
+by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
+
+lemma acc_modi_top1 [simp, intro!]: "a \<le> Public"
+by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
+
+lemma acc_modi_le_Public: 
+"a \<le> Public \<Longrightarrow> a=Private \<or> a = Package \<or> a=Protected \<or> a=Public"
+by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
+
+lemma acc_modi_bottom: "a \<le> Private \<Longrightarrow> a = Private"
+by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
+
+lemma acc_modi_Private_le: 
+"Private \<le> a \<Longrightarrow> a=Private \<or> a = Package \<or> a=Protected \<or> a=Public"
+by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
+
+lemma acc_modi_Package_le: 
+  "Package \<le> a \<Longrightarrow> a = Package \<or> a=Protected \<or> a=Public"
+by (auto simp add: le_acc_def less_acc_def split: acc_modi.split)
+
+lemma acc_modi_le_Package: 
+  "a \<le> Package \<Longrightarrow> a=Private \<or> a = Package"
+by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
+
+lemma acc_modi_Protected_le: 
+  "Protected \<le> a \<Longrightarrow> a=Protected \<or> a=Public"
+by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
+
+lemma acc_modi_le_Protected: 
+  "a \<le> Protected  \<Longrightarrow> a=Private \<or> a = Package \<or> a = Protected"
+by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits)
+
+
+lemmas acc_modi_le_Dests = acc_modi_top           acc_modi_le_Public
+                           acc_modi_Private_le    acc_modi_bottom
+                           acc_modi_Package_le    acc_modi_le_Package
+                           acc_modi_Protected_le  acc_modi_le_Protected
+
+lemma acc_modi_Package_le_cases 
+ [consumes 1,case_names Package Protected Public]:
+ "Package \<le> m \<Longrightarrow> ( m = Package \<Longrightarrow> P m) \<Longrightarrow> (m=Protected \<Longrightarrow> P m) \<Longrightarrow> 
+   (m=Public \<Longrightarrow> P m) \<Longrightarrow> P m"
+by (auto dest: acc_modi_Package_le)
+
+
+subsubsection {* Static Modifier *}
+types stat_modi = bool (* modifier: static *)
+
+subsection {* Declaration (base "class" for member,interface and class
+ declarations *}
+
+record decl =
+        access :: acc_modi
+
+translations
+  "decl" <= (type) "\<lparr>access::acc_modi\<rparr>"
+  "decl" <= (type) "\<lparr>access::acc_modi,\<dots>::'a\<rparr>"
+
+subsection {* Member (field or method)*}
+record  member = decl +
+         static :: stat_modi
+
+translations
+  "member" <= (type) "\<lparr>access::acc_modi,static::bool\<rparr>"
+  "member" <= (type) "\<lparr>access::acc_modi,static::bool,\<dots>::'a\<rparr>"
+
+subsection {* Field *}
+
+record field = member +
+        type :: ty
+translations
+  "field" <= (type) "\<lparr>access::acc_modi, static::bool, type::ty\<rparr>"
+  "field" <= (type) "\<lparr>access::acc_modi, static::bool, type::ty,\<dots>::'a\<rparr>"
+
+types     
+        fdecl           (* field declaration, cf. 8.3 *)
+	= "vname \<times> field"
+
+
+translations
+  "fdecl" <= (type) "vname \<times> field"
+
+subsection  {* Method *}
+
+record mhead = member +     (* method head (excluding signature) *)
+        pars ::"vname list" (* parameter names *)
+        resT ::ty           (* result type *)
+
+record mbody =                      (* method body *)
+        lcls::  "(vname \<times> ty) list" (* local variables *)
+        stmt:: stmt                 (* the body statement *)
+
+record methd = mhead + (* method in a class *)
+        mbody::mbody
+
+types mdecl = "sig \<times> methd"  (* method declaration in a class *)
+
+
+translations
+  "mhead" <= (type) "\<lparr>access::acc_modi, static::bool, 
+                      pars::vname list, resT::ty\<rparr>"
+  "mhead" <= (type) "\<lparr>access::acc_modi, static::bool, 
+                      pars::vname list, resT::ty,\<dots>::'a\<rparr>"
+  "mbody" <= (type) "\<lparr>lcls::(vname \<times> ty) list,stmt::stmt\<rparr>"
+  "mbody" <= (type) "\<lparr>lcls::(vname \<times> ty) list,stmt::stmt,\<dots>::'a\<rparr>"      
+  "methd" <= (type) "\<lparr>access::acc_modi, static::bool, 
+                      pars::vname list, resT::ty,mbody::mbody\<rparr>"
+  "methd" <= (type) "\<lparr>access::acc_modi, static::bool, 
+                      pars::vname list, resT::ty,mbody::mbody,\<dots>::'a\<rparr>"
+  "mdecl" <= (type) "sig \<times> methd"
+
+
+constdefs 
+  mhead::"methd \<Rightarrow> mhead"
+  "mhead m \<equiv> \<lparr>access=access m, static=static m, pars=pars m, resT=resT m\<rparr>"
+
+lemma access_mhead [simp]:"access (mhead m) = access m"
+by (simp add: mhead_def)
+
+lemma static_mhead [simp]:"static (mhead m) = static m"
+by (simp add: mhead_def)
+
+lemma pars_mhead [simp]:"pars (mhead m) = pars m"
+by (simp add: mhead_def)
+
+lemma resT_mhead [simp]:"resT (mhead m) = resT m"
+by (simp add: mhead_def)
+
+text {* To be able to talk uniformaly about field and method declarations we
+introduce the notion of a member declaration (e.g. useful to define 
+accessiblity ) *}
+
+datatype memberdecl = fdecl fdecl | mdecl mdecl
+
+datatype memberid = fid vname | mid sig
+
+axclass has_memberid < "type"
+consts
+ memberid :: "'a::has_memberid \<Rightarrow> memberid"
+
+instance memberdecl::has_memberid
+by (intro_classes)
+
+defs (overloaded)
+memberdecl_memberid_def:
+  "memberid m \<equiv> (case m of
+                    fdecl (vn,f)  \<Rightarrow> fid vn
+                  | mdecl (sig,m) \<Rightarrow> mid sig)"
+
+lemma memberid_fdecl_simp[simp]: "memberid (fdecl (vn,f)) = fid vn"
+by (simp add: memberdecl_memberid_def)
+
+lemma memberid_fdecl_simp1: "memberid (fdecl f) = fid (fst f)"
+by (cases f) (simp add: memberdecl_memberid_def)
+
+lemma memberid_mdecl_simp[simp]: "memberid (mdecl (sig,m)) = mid sig"
+by (simp add: memberdecl_memberid_def)
+
+lemma memberid_mdecl_simp1: "memberid (mdecl m) = mid (fst m)"
+by (cases m) (simp add: memberdecl_memberid_def)
+
+instance * :: ("type","has_memberid") has_memberid
+by (intro_classes)
+
+defs (overloaded)
+pair_memberid_def:
+  "memberid p \<equiv> memberid (snd p)"
+
+lemma memberid_pair_simp[simp]: "memberid (c,m) = memberid m"
+by (simp add: pair_memberid_def)
+
+lemma memberid_pair_simp1: "memberid p  = memberid (snd p)"
+by (simp add: pair_memberid_def)
+
+constdefs is_field :: "qtname \<times> memberdecl \<Rightarrow> bool"
+"is_field m \<equiv> \<exists> declC f. m=(declC,fdecl f)"
+  
+lemma is_fieldD: "is_field m \<Longrightarrow> \<exists> declC f. m=(declC,fdecl f)"
+by (simp add: is_field_def)
+
+lemma is_fieldI: "is_field (C,fdecl f)"
+by (simp add: is_field_def)
+
+constdefs is_method :: "qtname \<times> memberdecl \<Rightarrow> bool"
+"is_method membr \<equiv> \<exists> declC m. membr=(declC,mdecl m)"
+  
+lemma is_methodD: "is_method membr \<Longrightarrow> \<exists> declC m. membr=(declC,mdecl m)"
+by (simp add: is_method_def)
+
+lemma is_methodI: "is_method (C,mdecl m)"
+by (simp add: is_method_def)
+
+
+subsection {* Interface *}
+
+
+record  ibody = decl +  (* interface body *)
+          imethods :: "(sig \<times> mhead) list" (* method heads *)
+
+record  iface = ibody + (* interface *)
+         isuperIfs:: "qtname list" (* superinterface list *)
+types	
+	idecl           (* interface declaration, cf. 9.1 *)
+	= "qtname \<times> iface"
+
+translations
+  "ibody" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list\<rparr>"
+  "ibody" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list,\<dots>::'a\<rparr>"
+  "iface" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list,
+                      isuperIfs::qtname list\<rparr>"
+  "iface" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list,
+                      isuperIfs::qtname list,\<dots>::'a\<rparr>"
+  "idecl" <= (type) "qtname \<times> iface"
+
+constdefs
+  ibody :: "iface \<Rightarrow> ibody"
+  "ibody i \<equiv> \<lparr>access=access i,imethods=imethods i\<rparr>"
+
+lemma access_ibody [simp]: "(access (ibody i)) = access i"
+by (simp add: ibody_def)
+
+lemma imethods_ibody [simp]: "(imethods (ibody i)) = imethods i"
+by (simp add: ibody_def)
+
+subsection  {* Class *}
+record cbody = decl +          (* class body *)
+         cfields:: "fdecl list" 
+         methods:: "mdecl list"
+         init   :: "stmt"       (* initializer *)
+
+record class = cbody +           (* class *)
+        super   :: "qtname"      (* superclass *)
+        superIfs:: "qtname list" (* implemented interfaces *)
+types	
+	cdecl           (* class declaration, cf. 8.1 *)
+	= "qtname \<times> class"
+
+translations
+  "cbody" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
+                      methods::mdecl list,init::stmt\<rparr>"
+  "cbody" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
+                      methods::mdecl list,init::stmt,\<dots>::'a\<rparr>"
+  "class" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
+                      methods::mdecl list,init::stmt,
+                      super::qtname,superIfs::qtname list\<rparr>"
+  "class" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
+                      methods::mdecl list,init::stmt,
+                      super::qtname,superIfs::qtname list,\<dots>::'a\<rparr>"
+  "cdecl" <= (type) "qtname \<times> class"
+
+constdefs
+  cbody :: "class \<Rightarrow> cbody"
+  "cbody c \<equiv> \<lparr>access=access c, cfields=cfields c,methods=methods c,init=init c\<rparr>"
+
+lemma access_cbody [simp]:"access (cbody c) = access c"
+by (simp add: cbody_def)
+
+lemma cfields_cbody [simp]:"cfields (cbody c) = cfields c"
+by (simp add: cbody_def)
+
+lemma methods_cbody [simp]:"methods (cbody c) = methods c"
+by (simp add: cbody_def)
+
+lemma init_cbody [simp]:"init (cbody c) = init c"
+by (simp add: cbody_def)
+
+
+section "standard classes"
+
+consts
+
+  Object_mdecls  ::  "mdecl list" (* methods of Object *)
+  SXcpt_mdecls   ::  "mdecl list" (* methods of SXcpts *)
+  ObjectC ::         "cdecl"      (* declaration  of root      class   *)
+  SXcptC  ::"xname \<Rightarrow> cdecl"      (* declarations of throwable classes *)
+
+defs 
+
+
+ObjectC_def:"ObjectC  \<equiv> (Object,\<lparr>access=Public,cfields=[],methods=Object_mdecls,
+                                  init=Skip,super=arbitrary,superIfs=[]\<rparr>)"
+SXcptC_def:"SXcptC xn\<equiv> (SXcpt xn,\<lparr>access=Public,cfields=[],methods=SXcpt_mdecls,
+                                   init=Skip,
+                                   super=if xn = Throwable then Object 
+                                                           else SXcpt Throwable,
+                                   superIfs=[]\<rparr>)"
+
+lemma ObjectC_neq_SXcptC [simp]: "ObjectC \<noteq> SXcptC xn"
+by (simp add: ObjectC_def SXcptC_def Object_def SXcpt_def)
+
+lemma SXcptC_inject [simp]: "(SXcptC xn = SXcptC xm) = (xn = xm)"
+apply (simp add: SXcptC_def)
+apply auto
+done
+
+constdefs standard_classes :: "cdecl list"
+         "standard_classes \<equiv> [ObjectC, SXcptC Throwable,
+		SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast,
+		SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]"
+
+
+section "programs"
+
+record prog =
+        ifaces ::"idecl list"
+        "classes"::"cdecl list"
+
+translations
+     "prog"<= (type) "\<lparr>ifaces::idecl list,classes::cdecl list\<rparr>"
+     "prog"<= (type) "\<lparr>ifaces::idecl list,classes::cdecl list,\<dots>::'a\<rparr>"
+
+syntax
+  iface     :: "prog  \<Rightarrow> (qtname, iface) table"
+  class     :: "prog  \<Rightarrow> (qtname, class) table"
+  is_iface  :: "prog  \<Rightarrow> qtname  \<Rightarrow> bool"
+  is_class  :: "prog  \<Rightarrow> qtname  \<Rightarrow> bool"
+
+translations
+	   "iface G I" == "table_of (ifaces G) I"
+	   "class G C" == "table_of (classes G) C"
+	"is_iface G I" == "iface G I \<noteq> None"
+	"is_class G C" == "class G C \<noteq> None"
+
+
+section "is type"
+
+consts
+  is_type :: "prog \<Rightarrow>     ty \<Rightarrow> bool"
+  isrtype :: "prog \<Rightarrow> ref_ty \<Rightarrow> bool"
+
+primrec	"is_type G (PrimT pt)  = True"
+	"is_type G (RefT  rt)  = isrtype G rt"
+	"isrtype G (NullT    ) = True"
+	"isrtype G (IfaceT tn) = is_iface G tn"
+	"isrtype G (ClassT tn) = is_class G tn"
+	"isrtype G (ArrayT T ) = is_type  G T"
+
+lemma type_is_iface: "is_type G (Iface I) \<Longrightarrow> is_iface G I"
+by auto
+
+lemma type_is_class: "is_type G (Class C) \<Longrightarrow>  is_class G C"
+by auto
+
+
+section "subinterface and subclass relation, in anticipation of TypeRel.thy"
+
+consts 
+  subint1  :: "prog \<Rightarrow> (qtname \<times> qtname) set"
+  subcls1  :: "prog \<Rightarrow> (qtname \<times> qtname) set"
+
+defs
+  subint1_def: "subint1 G \<equiv> {(I,J). \<exists>i\<in>iface G I: J\<in>set (isuperIfs i)}"
+  subcls1_def: "subcls1 G \<equiv> {(C,D). C\<noteq>Object \<and> (\<exists>c\<in>class G C: super c = D)}"
+
+syntax
+ "@subcls1" :: "prog => [qtname, qtname] => bool" ("_|-_<:C1_" [71,71,71] 70)
+ "@subclseq":: "prog => [qtname, qtname] => bool" ("_|-_<=:C _"[71,71,71] 70)
+ "@subcls"  :: "prog => [qtname, qtname] => bool" ("_|-_<:C _"[71,71,71] 70)
+
+syntax (xsymbols)
+  "@subcls1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>\<^sub>C\<^sub>1_"  [71,71,71] 70)
+  "@subclseq":: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>\<^sub>C _"  [71,71,71] 70)
+  "@subcls"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>\<^sub>C _"  [71,71,71] 70)
+
+translations
+        "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" == "(C,D) \<in> subcls1 G"
+	"G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" (* cf. 8.1.3 *)
+        "G\<turnstile>C \<prec>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^+"
+ 
+
+lemma subint1I: "\<lbrakk>iface G I = Some i; J \<in> set (isuperIfs i)\<rbrakk> 
+                 \<Longrightarrow> (I,J) \<in> subint1 G" 
+apply (simp add: subint1_def)
+done
+
+lemma subcls1I:"\<lbrakk>class G C = Some c; C \<noteq> Object\<rbrakk> \<Longrightarrow> (C,(super c)) \<in> subcls1 G"
+apply (simp add: subcls1_def)
+done
+
+
+lemma subint1D: "(I,J)\<in>subint1 G\<Longrightarrow> \<exists>i\<in>iface G I: J\<in>set (isuperIfs i)"
+by (simp add: subint1_def)
+
+lemma subcls1D: 
+  "(C,D)\<in>subcls1 G \<Longrightarrow> C\<noteq>Object \<and> (\<exists>c. class G C = Some c \<and> (super c = D))"
+apply (simp add: subcls1_def)
+apply auto
+done
+
+lemma subint1_def2:  
+   "subint1 G = (\<Sigma> I\<in>{I. is_iface G I}. set (isuperIfs (the (iface G I))))"
+apply (unfold subint1_def)
+apply auto
+done
+
+lemma subcls1_def2: 
+ "subcls1 G = (\<Sigma>C\<in>{C. is_class G C}. {D. C\<noteq>Object \<and> super (the(class G C))=D})"
+apply (unfold subcls1_def)
+apply auto
+done
+
+lemma subcls_is_class:
+"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D\<rbrakk> \<Longrightarrow> \<exists> c. class G C = Some c"
+by (auto simp add: subcls1_def dest: tranclD)
+
+lemma no_subcls1_Object:"G\<turnstile>Object\<prec>\<^sub>C\<^sub>1 D \<Longrightarrow> P"
+by (auto simp add: subcls1_def)
+
+lemma no_subcls_Object: "G\<turnstile>Object\<prec>\<^sub>C D \<Longrightarrow> P"
+apply (erule trancl_induct)
+apply (auto intro: no_subcls1_Object)
+done
+
+section "well-structured programs"
+
+constdefs
+  ws_idecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname list \<Rightarrow> bool"
+ "ws_idecl G I si \<equiv> \<forall>J\<in>set si.  is_iface G J   \<and> (J,I)\<notin>(subint1 G)^+"
+  
+  ws_cdecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+ "ws_cdecl G C sc \<equiv> C\<noteq>Object \<longrightarrow> is_class G sc \<and> (sc,C)\<notin>(subcls1 G)^+"
+  
+  ws_prog  :: "prog \<Rightarrow> bool"
+ "ws_prog G \<equiv> (\<forall>(I,i)\<in>set (ifaces  G). ws_idecl G I (isuperIfs i)) \<and> 
+	      (\<forall>(C,c)\<in>set (classes G). ws_cdecl G C (super c))"
+
+
+lemma ws_progI: 
+"\<lbrakk>\<forall>(I,i)\<in>set (ifaces G). \<forall>J\<in>set (isuperIfs i). is_iface G J \<and> 
+                                                (J,I) \<notin> (subint1 G)^+; 
+  \<forall>(C,c)\<in>set (classes G). C\<noteq>Object \<longrightarrow> is_class G (super c) \<and> 
+                                        ((super c),C) \<notin> (subcls1 G)^+  
+ \<rbrakk> \<Longrightarrow> ws_prog G"
+apply (unfold ws_prog_def ws_idecl_def ws_cdecl_def)
+apply (erule_tac conjI)
+apply blast
+done
+
+lemma ws_prog_ideclD: 
+"\<lbrakk>iface G I = Some i; J\<in>set (isuperIfs i); ws_prog G\<rbrakk> \<Longrightarrow>  
+  is_iface G J \<and> (J,I)\<notin>(subint1 G)^+"
+apply (unfold ws_prog_def ws_idecl_def)
+apply clarify
+apply (drule_tac map_of_SomeD)
+apply auto
+done
+
+lemma ws_prog_cdeclD: 
+"\<lbrakk>class G C = Some c; C\<noteq>Object; ws_prog G\<rbrakk> \<Longrightarrow>  
+  is_class G (super c) \<and> (super c,C)\<notin>(subcls1 G)^+"
+apply (unfold ws_prog_def ws_cdecl_def)
+apply clarify
+apply (drule_tac map_of_SomeD)
+apply auto
+done
+
+
+section "well-foundedness"
+
+lemma finite_is_iface: "finite {I. is_iface G I}"
+apply (fold dom_def)
+apply (rule_tac finite_dom_map_of)
+done
+
+lemma finite_is_class: "finite {C. is_class G C}"
+apply (fold dom_def)
+apply (rule_tac finite_dom_map_of)
+done
+
+lemma finite_subint1: "finite (subint1 G)"
+apply (subst subint1_def2)
+apply (rule finite_SigmaI)
+apply (rule finite_is_iface)
+apply (simp (no_asm))
+done
+
+lemma finite_subcls1: "finite (subcls1 G)"
+apply (subst subcls1_def2)
+apply (rule finite_SigmaI)
+apply (rule finite_is_class)
+apply (rule_tac B = "{super (the (class G C))}" in finite_subset)
+apply  auto
+done
+
+lemma subint1_irrefl_lemma1: 
+  "ws_prog G \<Longrightarrow> (subint1 G)^-1 \<inter> (subint1 G)^+ = {}"
+apply (force dest: subint1D ws_prog_ideclD conjunct2)
+done
+
+lemma subcls1_irrefl_lemma1: 
+  "ws_prog G \<Longrightarrow> (subcls1 G)^-1 \<inter> (subcls1 G)^+ = {}"
+apply (force dest: subcls1D ws_prog_cdeclD conjunct2)
+done
+
+lemmas subint1_irrefl_lemma2 = subint1_irrefl_lemma1 [THEN irrefl_tranclI']
+lemmas subcls1_irrefl_lemma2 = subcls1_irrefl_lemma1 [THEN irrefl_tranclI']
+
+lemma subint1_irrefl: "\<lbrakk>(x, y) \<in> subint1 G; ws_prog G\<rbrakk> \<Longrightarrow> x \<noteq> y"
+apply (rule irrefl_trancl_rD)
+apply (rule subint1_irrefl_lemma2)
+apply auto
+done
+
+lemma subcls1_irrefl: "\<lbrakk>(x, y) \<in> subcls1 G; ws_prog G\<rbrakk> \<Longrightarrow> x \<noteq> y"
+apply (rule irrefl_trancl_rD)
+apply (rule subcls1_irrefl_lemma2)
+apply auto
+done
+
+lemmas subint1_acyclic = subint1_irrefl_lemma2 [THEN acyclicI, standard]
+lemmas subcls1_acyclic = subcls1_irrefl_lemma2 [THEN acyclicI, standard]
+
+
+lemma wf_subint1: "ws_prog G \<Longrightarrow> wf ((subint1 G)\<inverse>)"
+by (auto intro: finite_acyclic_wf_converse finite_subint1 subint1_acyclic)
+
+lemma wf_subcls1: "ws_prog G \<Longrightarrow> wf ((subcls1 G)\<inverse>)"
+by (auto intro: finite_acyclic_wf_converse finite_subcls1 subcls1_acyclic)
+
+
+lemma subint1_induct: 
+  "\<lbrakk>ws_prog G; \<And>x. \<forall>y. (x, y) \<in> subint1 G \<longrightarrow> P y \<Longrightarrow> P x\<rbrakk> \<Longrightarrow> P a"
+apply (frule wf_subint1)
+apply (erule wf_induct)
+apply (simp (no_asm_use) only: converse_iff)
+apply blast
+done
+
+lemma subcls1_induct [consumes 1]:
+  "\<lbrakk>ws_prog G; \<And>x. \<forall>y. (x, y) \<in> subcls1 G \<longrightarrow> P y \<Longrightarrow> P x\<rbrakk> \<Longrightarrow> P a"
+apply (frule wf_subcls1)
+apply (erule wf_induct)
+apply (simp (no_asm_use) only: converse_iff)
+apply blast
+done
+
+lemma ws_subint1_induct: 
+ "\<lbrakk>is_iface G I; ws_prog G; \<And>I i. \<lbrakk>iface G I = Some i \<and> 
+   (\<forall>J \<in> set (isuperIfs i). (I,J)\<in>subint1 G \<and> P J \<and> is_iface G J)\<rbrakk> \<Longrightarrow> P I
+  \<rbrakk> \<Longrightarrow> P I"
+apply (erule make_imp)
+apply (rule subint1_induct)
+apply  assumption
+apply safe
+apply (fast dest: subint1I ws_prog_ideclD)
+done
+
+
+lemma ws_subcls1_induct: "\<lbrakk>is_class G C; ws_prog G;  
+  \<And>C c. \<lbrakk>class G C = Some c;  
+ (C \<noteq> Object \<longrightarrow> (C,(super c))\<in>subcls1 G \<and> 
+                  P (super c) \<and> is_class G (super c))\<rbrakk> \<Longrightarrow> P C
+ \<rbrakk> \<Longrightarrow> P C"
+apply (erule make_imp)
+apply (rule subcls1_induct)
+apply  assumption
+apply safe
+apply (fast dest: subcls1I ws_prog_cdeclD)
+done
+
+lemma ws_class_induct [consumes 2, case_names Object Subcls]:
+"\<lbrakk>class G C = Some c; ws_prog G; 
+  \<And> co. class G Object = Some co \<Longrightarrow> P Object; 
+  \<And>  C c. \<lbrakk>class G C = Some c; C \<noteq> Object; P (super c)\<rbrakk> \<Longrightarrow> P C
+ \<rbrakk> \<Longrightarrow> P C"
+proof -
+  assume clsC: "class G C = Some c"
+  and    init: "\<And> co. class G Object = Some co \<Longrightarrow> P Object"
+  and    step: "\<And>   C c. \<lbrakk>class G C = Some c; C \<noteq> Object; P (super c)\<rbrakk> \<Longrightarrow> P C"
+  assume ws: "ws_prog G"
+  then have "is_class G C \<Longrightarrow> P C"  
+  proof (induct rule: subcls1_induct)
+    fix C
+    assume   hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<longrightarrow> is_class G S \<longrightarrow> P S"
+       and iscls:"is_class G C"
+    show "P C"
+    proof (cases "C=Object")
+      case True with iscls init show "P C" by auto
+    next
+      case False with ws step hyp iscls 
+      show "P C" by (auto dest: subcls1I ws_prog_cdeclD)
+    qed
+  qed
+  with clsC show ?thesis by simp
+qed
+
+lemma ws_class_induct' [consumes 2, case_names Object Subcls]:
+"\<lbrakk>is_class G C; ws_prog G; 
+  \<And> co. class G Object = Some co \<Longrightarrow> P Object; 
+  \<And> C c. \<lbrakk>class G C = Some c; C \<noteq> Object; P (super c)\<rbrakk> \<Longrightarrow> P C
+ \<rbrakk> \<Longrightarrow> P C"
+by (blast intro: ws_class_induct)
+
+lemma ws_class_induct'' [consumes 2, case_names Object Subcls]:
+"\<lbrakk>class G C = Some c; ws_prog G; 
+  \<And> co. class G Object = Some co \<Longrightarrow> P Object co; 
+  \<And>  C c sc. \<lbrakk>class G C = Some c; class G (super c) = Some sc;
+            C \<noteq> Object; P (super c) sc\<rbrakk> \<Longrightarrow> P C c
+ \<rbrakk> \<Longrightarrow> P C c"
+proof -
+  assume clsC: "class G C = Some c"
+  and    init: "\<And> co. class G Object = Some co \<Longrightarrow> P Object co"
+  and    step: "\<And> C c sc . \<lbrakk>class G C = Some c; class G (super c) = Some sc;
+                             C \<noteq> Object; P (super c) sc\<rbrakk> \<Longrightarrow> P C c"
+  assume ws: "ws_prog G"
+  then have "\<And> c. class G C = Some c\<Longrightarrow> P C c"  
+  proof (induct rule: subcls1_induct)
+    fix C c
+    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)"
+       and iscls:"class G C = Some c"
+    show "P C c"
+    proof (cases "C=Object")
+      case True with iscls init show "P C c" by auto
+    next
+      case False
+      with ws iscls obtain sc where
+	sc: "class G (super c) = Some sc"
+	by (auto dest: ws_prog_cdeclD)
+      from iscls False have "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 (super c)" by (rule subcls1I)
+      with False ws step hyp iscls sc
+      show "P C c" 
+	by (auto)  
+    qed
+  qed
+  with clsC show "P C c" by auto
+qed
+
+lemma ws_interface_induct [consumes 2, case_names Step]:
+ (assumes is_if_I: "is_iface G I" and 
+               ws: "ws_prog G" and
+          hyp_sub: "\<And>I i. \<lbrakk>iface G I = Some i; 
+                            \<forall> J \<in> set (isuperIfs i).
+                                 (I,J)\<in>subint1 G \<and> P J \<and> is_iface G J\<rbrakk> \<Longrightarrow> P I"
+ ) "P I"
+proof -
+  from is_if_I ws 
+  show "P I"
+  proof (rule ws_subint1_induct)
+    fix I i
+    assume hyp: "iface G I = Some i \<and>
+                (\<forall>J\<in>set (isuperIfs i). (I,J) \<in>subint1 G \<and> P J \<and> is_iface G J)"
+    then have if_I: "iface G I = Some i"
+      by blast
+    show "P I"
+    proof (cases "isuperIfs i")
+      case Nil
+      with if_I hyp_sub 
+      show "P I" 
+	by auto
+    next
+      case (Cons hd tl)
+      with hyp if_I hyp_sub 
+      show "P I" 
+	by auto
+    qed
+  qed
+qed
+
+section "general recursion operators for the interface and class hiearchies"
+
+consts
+  iface_rec  :: "prog \<times> qtname \<Rightarrow>   \<spacespace>  (qtname \<Rightarrow> iface \<Rightarrow> 'a set \<Rightarrow> 'a) \<Rightarrow> 'a"
+  class_rec  :: "prog \<times> qtname \<Rightarrow> 'a \<Rightarrow> (qtname \<Rightarrow> class \<Rightarrow> 'a     \<Rightarrow> 'a) \<Rightarrow> 'a"
+
+recdef iface_rec "same_fst ws_prog (\<lambda>G. (subint1 G)^-1)" 
+"iface_rec (G,I) = 
+  (\<lambda>f. case iface G I of 
+         None \<Rightarrow> arbitrary 
+       | Some i \<Rightarrow> if ws_prog G 
+                      then f I i 
+                               ((\<lambda>J. iface_rec (G,J) f)`set (isuperIfs i))
+                      else arbitrary)"
+(hints recdef_wf: wf_subint1 intro: subint1I)
+declare iface_rec.simps [simp del]
+
+lemma iface_rec: 
+"\<lbrakk>iface G I = Some i; ws_prog G\<rbrakk> \<Longrightarrow> 
+ iface_rec (G,I) f = f I i ((\<lambda>J. iface_rec (G,J) f)`set (isuperIfs i))"
+apply (subst iface_rec.simps)
+apply simp
+done
+
+recdef class_rec "same_fst ws_prog (\<lambda>G. (subcls1 G)^-1)"
+"class_rec(G,C) = 
+  (\<lambda>t f. case class G C of 
+           None \<Rightarrow> arbitrary 
+         | Some c \<Rightarrow> if ws_prog G 
+                        then f C c 
+                                 (if C = Object then t 
+                                                else class_rec (G,super c) t f)
+                        else arbitrary)"
+(hints recdef_wf: wf_subcls1 intro: subcls1I)
+declare class_rec.simps [simp del]
+
+lemma class_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
+ class_rec (G,C) t f = 
+   f C c (if C = Object then t else class_rec (G,super c) t f)"
+apply (rule class_rec.simps [THEN trans [THEN fun_cong [THEN fun_cong]]])
+apply simp
+done
+(*
+lemma bar:
+ "[| P;  !!x.  P ==> Q x  |] ==> Q x"
+by simp
+
+lemma metaMP: "[| A ==> B; A |] ==> B"
+by blast
+
+lemma True
+proof- 
+  presume t: "C  ==> E"
+  thm metaMP [OF t]
+
+  presume r1: "\<And> B. P \<Longrightarrow> B"
+  presume r2: "\<And> C. C \<Longrightarrow> P"
+  thm r1 [OF r2]
+
+  thm metaMP [OF t]
+
+lemma ws_subcls1_induct4: "\<lbrakk>is_class G C; ws_prog G;  
+  \<And>C c. \<lbrakk>C \<noteq> Object\<longrightarrow> P (super c)\<rbrakk> \<Longrightarrow> P C
+ \<rbrakk> \<Longrightarrow> P C"
+proof -
+  assume cls_C: "is_class G C"
+  and       ws: "ws_prog G"
+  and      hyp: "\<And>C c. \<lbrakk>C \<noteq> Object\<longrightarrow> P (super c)\<rbrakk> \<Longrightarrow> P C"
+  thm ws_subcls1_induct [OF cls_C ws hyp]
+
+show
+(\<And>C c. class G C = Some c \<and>
+       (C \<noteq> Object \<longrightarrow> G\<turnstile>C\<prec>\<^sub>C\<^sub>1super c \<and> ?P (super c) \<and> is_class G (super c)) \<Longrightarrow>
+       ?P C) \<Longrightarrow>
+?P C
+  show ?thesis
+    thm "thm ws_subcls1_induct [OF cls_C ws hyp]"
+    apply (rule ws_subcls1_induct)
+  proof (rule ws_subcls1_induct)
+    fix C c
+    assume "class G C = Some c \<and>
+            (C \<noteq> Object \<longrightarrow>
+              G\<turnstile>C\<prec>\<^sub>C\<^sub>1super c \<and> P (super c) \<and> is_class G (super c))"
+    show "C \<noteq> Object \<longrightarrow> P (super  (?c C c))" 
+apply (erule ws_subcls1_induct)
+apply assumption
+apply (erule conjE)
+apply (case_tac "C=Object")
+apply blast
+apply (erule impE)
+apply assumption
+apply (erule conjE)+
+apply (rotate_tac 2)
+sorry
+
+*)
+
+
+constdefs
+imethds:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
+  (* methods of an interface, with overriding and inheritance, cf. 9.2 *)
+"imethds G I 
+  \<equiv> iface_rec (G,I)  
+              (\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus> 
+                        (o2s \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
+	
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/DeclConcepts.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,2540 @@
+header {* Advanced concepts on Java declarations like overriding, inheritance,
+dynamic method lookup*}
+
+theory DeclConcepts = TypeRel:
+
+section "access control (cf. 6.6), overriding and hiding (cf. 8.4.6.1)"
+
+constdefs
+is_public :: "prog \<Rightarrow> qtname \<Rightarrow> bool"
+"is_public G qn \<equiv> (case class G qn of
+                     None       \<Rightarrow> (case iface G qn of
+                                      None       \<Rightarrow> False
+                                    | Some iface \<Rightarrow> access iface = Public)
+                   | Some class \<Rightarrow> access class = Public)"
+
+subsection "accessibility of types (cf. 6.6.1)"
+text {* 
+Primitive types are always accessible, interfaces and classes are accessible
+in their package or if they are defined public, an array type is accessible if
+its element type is accessible *}
+ 
+consts accessible_in   :: "prog \<Rightarrow> ty     \<Rightarrow> pname \<Rightarrow> bool"  
+                                      ("_ \<turnstile> _ accessible'_in _" [61,61,61] 60)
+       rt_accessible_in:: "prog \<Rightarrow> ref_ty \<Rightarrow> pname \<Rightarrow> bool" 
+                                      ("_ \<turnstile> _ accessible'_in' _" [61,61,61] 60) 
+primrec
+"G\<turnstile>(PrimT p)   accessible_in pack  = True"
+accessible_in_RefT_simp: 
+"G\<turnstile>(RefT  r)   accessible_in pack  = G\<turnstile>r accessible_in' pack"
+
+"G\<turnstile>(NullT)     accessible_in' pack = True"
+"G\<turnstile>(IfaceT I)  accessible_in' pack = ((pid I = pack) \<or> is_public G I)"
+"G\<turnstile>(ClassT C)  accessible_in' pack = ((pid C = pack) \<or> is_public G C)"
+"G\<turnstile>(ArrayT ty) accessible_in' pack = G\<turnstile>ty accessible_in pack"
+
+declare accessible_in_RefT_simp [simp del]
+
+constdefs
+  is_acc_class :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
+    "is_acc_class G P C \<equiv> is_class G C \<and> G\<turnstile>(Class C) accessible_in P"
+  is_acc_iface :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
+    "is_acc_iface G P I \<equiv> is_iface G I \<and> G\<turnstile>(Iface I) accessible_in P"
+  is_acc_type  :: "prog \<Rightarrow> pname \<Rightarrow> ty     \<Rightarrow> bool"
+    "is_acc_type  G P T \<equiv> is_type G T  \<and> G\<turnstile>T accessible_in P"
+  is_acc_reftype  :: "prog \<Rightarrow> pname \<Rightarrow> ref_ty \<Rightarrow> bool"
+  "is_acc_reftype  G P T \<equiv> isrtype G T  \<and> G\<turnstile>T accessible_in' P"
+
+lemma is_acc_classD:
+ "is_acc_class G P C \<Longrightarrow>  is_class G C \<and> G\<turnstile>(Class C) accessible_in P"
+by (simp add: is_acc_class_def)
+
+lemma is_acc_class_is_class: "is_acc_class G P C \<Longrightarrow> is_class G C"
+by (auto simp add: is_acc_class_def)
+
+lemma is_acc_ifaceD:
+  "is_acc_iface G P I \<Longrightarrow> is_iface G I \<and> G\<turnstile>(Iface I) accessible_in P"
+by (simp add: is_acc_iface_def)
+
+lemma is_acc_typeD:
+ "is_acc_type  G P T \<equiv> is_type G T  \<and> G\<turnstile>T accessible_in P"
+by (simp add: is_acc_type_def)
+
+lemma is_acc_reftypeD:
+"is_acc_reftype  G P T \<Longrightarrow> isrtype G T  \<and> G\<turnstile>T accessible_in' P"
+by (simp add: is_acc_reftype_def)
+
+subsection "accessibility of members"
+text {*
+The accessibility of members is more involved as the accessibility of types.
+We have to distinguish several cases to model the different effects of 
+accessibility during inheritance, overriding and ordinary member access 
+*}
+
+subsubsection {* Various technical conversion and selection functions *}
+
+text {* overloaded selector @{text accmodi} to select the access modifier 
+out of various HOL types *}
+
+axclass has_accmodi < "type"
+consts accmodi:: "'a::has_accmodi \<Rightarrow> acc_modi"
+
+instance acc_modi::has_accmodi
+by (intro_classes)
+
+defs (overloaded)
+acc_modi_accmodi_def: "accmodi (a::acc_modi) \<equiv> a"
+
+lemma acc_modi_accmodi_simp[simp]: "accmodi (a::acc_modi) = a"
+by (simp add: acc_modi_accmodi_def)
+
+instance access_field_type:: ("type","type") has_accmodi
+by (intro_classes)
+
+defs (overloaded)
+decl_acc_modi_def: "accmodi (d::('a:: type) decl_scheme) \<equiv> access d"
+
+
+lemma decl_acc_modi_simp[simp]: "accmodi (d::('a::type) decl_scheme) = access d"
+by (simp add: decl_acc_modi_def)
+
+instance * :: ("type",has_accmodi) has_accmodi
+by (intro_classes)
+
+defs (overloaded)
+pair_acc_modi_def: "accmodi p \<equiv> (accmodi (snd p))"
+
+lemma pair_acc_modi_simp[simp]: "accmodi (x,a) = (accmodi a)"
+by (simp add: pair_acc_modi_def)
+
+instance memberdecl :: has_accmodi
+by (intro_classes)
+
+defs (overloaded)
+memberdecl_acc_modi_def: "accmodi m \<equiv> (case m of
+                                          fdecl f \<Rightarrow> accmodi f
+                                        | mdecl m \<Rightarrow> accmodi m)"
+
+lemma memberdecl_fdecl_acc_modi_simp[simp]:
+ "accmodi (fdecl m) = accmodi m"
+by (simp add: memberdecl_acc_modi_def)
+
+lemma memberdecl_mdecl_acc_modi_simp[simp]:
+ "accmodi (mdecl m) = accmodi m"
+by (simp add: memberdecl_acc_modi_def)
+
+text {* overloaded selector @{text declclass} to select the declaring class 
+out of various HOL types *}
+
+axclass has_declclass < "type"
+consts declclass:: "'a::has_declclass \<Rightarrow> qtname"
+
+instance pid_field_type::("type","type") has_declclass
+by (intro_classes)
+
+defs (overloaded)
+qtname_declclass_def: "declclass (q::qtname) \<equiv> q"
+
+lemma qtname_declclass_simp[simp]: "declclass (q::qtname) = q"
+by (simp add: qtname_declclass_def)
+
+instance * :: ("has_declclass","type") has_declclass
+by (intro_classes)
+
+defs (overloaded)
+pair_declclass_def: "declclass p \<equiv> declclass (fst p)"
+
+lemma pair_declclass_simp[simp]: "declclass (c,x) = declclass c" 
+by (simp add: pair_declclass_def)
+
+text {* overloaded selector @{text is_static} to select the static modifier 
+out of various HOL types *}
+
+
+axclass has_static < "type"
+consts is_static :: "'a::has_static \<Rightarrow> bool"
+
+(*
+consts is_static :: "'a \<Rightarrow> bool"
+*)
+
+instance access_field_type :: ("type","has_static") has_static
+by (intro_classes) 
+
+defs (overloaded)
+decl_is_static_def: 
+ "is_static (m::('a::has_static) decl_scheme) \<equiv> is_static (Decl.decl.more m)" 
+
+instance static_field_type :: ("type","type") has_static
+by (intro_classes)
+
+defs (overloaded)
+static_field_type_is_static_def: 
+ "is_static (m::(bool,'b::type) static_field_type) \<equiv> static_val m"
+
+lemma member_is_static_simp: "is_static (m::'a member_scheme) = static m"
+apply (cases m)
+apply (simp add: static_field_type_is_static_def 
+                 decl_is_static_def Decl.member.dest_convs)
+done
+
+instance * :: ("type","has_static") has_static
+by (intro_classes)
+
+defs (overloaded)
+pair_is_static_def: "is_static p \<equiv> is_static (snd p)"
+
+lemma pair_is_static_simp [simp]: "is_static (x,s) = is_static s"
+by (simp add: pair_is_static_def)
+
+lemma pair_is_static_simp1: "is_static p = is_static (snd p)"
+by (simp add: pair_is_static_def)
+
+instance memberdecl:: has_static
+by (intro_classes)
+
+defs (overloaded)
+memberdecl_is_static_def: 
+ "is_static m \<equiv> (case m of
+                    fdecl f \<Rightarrow> is_static f
+                  | mdecl m \<Rightarrow> is_static m)"
+
+lemma memberdecl_is_static_fdecl_simp[simp]:
+ "is_static (fdecl f) = is_static f"
+by (simp add: memberdecl_is_static_def)
+
+lemma memberdecl_is_static_mdecl_simp[simp]:
+ "is_static (mdecl m) = is_static m"
+by (simp add: memberdecl_is_static_def)
+
+lemma mhead_static_simp [simp]: "is_static (mhead m) = is_static m"
+by (cases m) (simp add: mhead_def member_is_static_simp)
+
+constdefs  (* some mnemotic selectors for (qtname \<times> ('a::more) decl_scheme) 
+            * the first component is a class or interface name
+            * the second component is a method, field or method head *)
+(* "declclass":: "(qtname \<times> ('a::more) decl_scheme) \<Rightarrow> qtname"*)
+(* "declclass \<equiv> fst" *)          (* get the class component *)
+
+ "decliface":: "(qtname \<times> ('a::type) decl_scheme) \<Rightarrow> qtname"
+ "decliface \<equiv> fst"          (* get the interface component *)
+
+(*
+ "member"::   "(qtname \<times> ('a::type) decl_scheme) \<Rightarrow> ('a::type) decl_scheme"
+*)
+ "mbr"::   "(qtname \<times> memberdecl) \<Rightarrow> memberdecl"
+ "mbr \<equiv> snd"            (* get the memberdecl component *)
+
+ "mthd"::   "('b \<times> 'a) \<Rightarrow> 'a"
+                           (* also used for mdecl,mhead *)
+ "mthd \<equiv> snd"              (* get the method component *)
+
+ "fld"::   "('b \<times> ('a::type) decl_scheme) \<Rightarrow> ('a::type) decl_scheme"
+              (* also used for ((vname \<times> qtname)\<times> field) *)
+ "fld \<equiv> snd"               (* get the field component *)
+
+(* "accmodi" :: "('b \<times> ('a::type) decl_scheme) \<Rightarrow> acc_modi"*)
+                           (* also used for mdecl *)
+(* "accmodi \<equiv> access \<circ> snd"*)  (* get the access modifier *) 
+(*
+ "is_static" ::"('b \<times> ('a::type) member_scheme) \<Rightarrow> bool" *)
+                            (* also defined for emhead cf. WellType *)
+ (*"is_static \<equiv> static \<circ> snd"*) (* get the static modifier *)
+
+constdefs (* some mnemotic selectors for (vname \<times> qtname) *)
+ fname:: "(vname \<times> 'a) \<Rightarrow> vname" (* also used for fdecl *)
+ "fname \<equiv> fst"
+  
+  declclassf:: "(vname \<times> qtname) \<Rightarrow> qtname"
+ "declclassf \<equiv> snd"
+
+(*
+lemma declclass_simp[simp]: "declclass (C,m) = C"
+by (simp add: declclass_def)
+*)
+
+lemma decliface_simp[simp]: "decliface (I,m) = I"
+by (simp add: decliface_def) 
+
+lemma mbr_simp[simp]: "mbr (C,m) = m"
+by (simp add: mbr_def)
+
+lemma access_mbr_simp [simp]: "(accmodi (mbr m)) = accmodi m"
+by (cases m) (simp add:  mbr_def) 
+
+lemma mthd_simp[simp]: "mthd (C,m) = m"
+by (simp add: mthd_def)
+
+lemma fld_simp[simp]: "fld (C,f) = f"
+by (simp add: fld_def)
+
+lemma accmodi_simp[simp]: "accmodi (C,m) = access m"
+by (simp )
+
+lemma access_mthd_simp [simp]: "(access (mthd m)) = accmodi m"
+by (cases m) (simp add: mthd_def) 
+
+lemma access_fld_simp [simp]: "(access (fld f)) = accmodi f"
+by (cases f) (simp add:  fld_def) 
+
+(*
+lemma is_static_simp[simp]: "is_static (C,m) = static m"
+by (simp add: is_static_def)
+*)
+
+lemma static_mthd_simp[simp]: "static (mthd m) = is_static m"
+by (cases m) (simp add:  mthd_def member_is_static_simp)
+
+lemma mthd_is_static_simp [simp]: "is_static (mthd m) = is_static m"
+by (cases m) simp
+
+lemma static_fld_simp[simp]: "static (fld f) = is_static f"
+by (cases f) (simp add:  fld_def member_is_static_simp)
+
+lemma ext_field_simp [simp]: "(declclass f,fld f) = f"
+by (cases f) (simp add:  fld_def)
+
+lemma ext_method_simp [simp]: "(declclass m,mthd m) = m"
+by (cases m) (simp add:  mthd_def)
+
+lemma ext_mbr_simp [simp]: "(declclass m,mbr m) = m"
+by (cases m) (simp add: mbr_def)
+
+lemma fname_simp[simp]:"fname (n,c) = n"
+by (simp add: fname_def)
+
+lemma declclassf_simp[simp]:"declclassf (n,c) = c"
+by (simp add: declclassf_def)
+
+constdefs  (* some mnemotic selectors for (vname \<times> qtname) *)
+  "fldname"  :: "(vname \<times> qtname) \<Rightarrow> vname" 
+  "fldname \<equiv> fst"
+
+  "fldclass" :: "(vname \<times> qtname) \<Rightarrow> qtname"
+  "fldclass \<equiv> snd"
+
+lemma fldname_simp[simp]: "fldname (n,c) = n"
+by (simp add: fldname_def)
+
+lemma fldclass_simp[simp]: "fldclass (n,c) = c"
+by (simp add: fldclass_def)
+
+lemma ext_fieldname_simp[simp]: "(fldname f,fldclass f) = f"
+by (simp add: fldname_def fldclass_def)
+
+text {* Convert a qualified method declaration (qualified with its declaring 
+class) to a qualified member declaration:  @{text methdMembr}  *}
+
+constdefs
+methdMembr :: "(qtname \<times> mdecl) \<Rightarrow> (qtname \<times> memberdecl)"
+ "methdMembr m \<equiv> (fst m,mdecl (snd m))"
+
+lemma methdMembr_simp[simp]: "methdMembr (c,m) = (c,mdecl m)"
+by (simp add: methdMembr_def)
+
+lemma accmodi_methdMembr_simp[simp]: "accmodi (methdMembr m) = accmodi m"
+by (cases m) (simp add: methdMembr_def)
+
+lemma is_static_methdMembr_simp[simp]: "is_static (methdMembr m) = is_static m"
+by (cases m) (simp add: methdMembr_def)
+
+lemma declclass_methdMembr_simp[simp]: "declclass (methdMembr m) = declclass m"
+by (cases m) (simp add: methdMembr_def)
+
+text {* Convert a qualified method (qualified with its declaring 
+class) to a qualified member declaration:  @{text method}  *}
+
+constdefs
+method :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> memberdecl)" 
+"method sig m \<equiv> (declclass m, mdecl (sig, mthd m))"
+
+lemma method_simp[simp]: "method sig (C,m) = (C,mdecl (sig,m))"
+by (simp add: method_def)
+
+lemma accmodi_method_simp[simp]: "accmodi (method sig m) = accmodi m"
+by (simp add: method_def)
+
+lemma declclass_method_simp[simp]: "declclass (method sig m) = declclass m"
+by (simp add: method_def)
+
+lemma is_static_method_simp[simp]: "is_static (method sig m) = is_static m"
+by (cases m) (simp add: method_def)
+
+lemma mbr_method_simp[simp]: "mbr (method sig m) = mdecl (sig,mthd m)"
+by (simp add: mbr_def method_def)
+
+lemma memberid_method_simp[simp]:  "memberid (method sig m) = mid sig"
+  by (simp add: method_def) 
+
+constdefs
+fieldm :: "vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> (qtname \<times> memberdecl)" 
+"fieldm n f \<equiv> (declclass f, fdecl (n, fld f))"
+
+lemma fieldm_simp[simp]: "fieldm n (C,f) = (C,fdecl (n,f))"
+by (simp add: fieldm_def)
+
+lemma accmodi_fieldm_simp[simp]: "accmodi (fieldm n f) = accmodi f"
+by (simp add: fieldm_def)
+
+lemma declclass_fieldm_simp[simp]: "declclass (fieldm n f) = declclass f"
+by (simp add: fieldm_def)
+
+lemma is_static_fieldm_simp[simp]: "is_static (fieldm n f) = is_static f"
+by (cases f) (simp add: fieldm_def)
+
+lemma mbr_fieldm_simp[simp]: "mbr (fieldm n f) = fdecl (n,fld f)"
+by (simp add: mbr_def fieldm_def)
+
+lemma memberid_fieldm_simp[simp]:  "memberid (fieldm n f) = fid n"
+by (simp add: fieldm_def) 
+
+text {* Select the signature out of a qualified method declaration:
+ @{text msig} *}
+
+constdefs msig:: "(qtname \<times> mdecl) \<Rightarrow> sig"
+"msig m \<equiv> fst (snd m)"
+
+lemma msig_simp[simp]: "msig (c,(s,m)) = s"
+by (simp add: msig_def)
+
+text {* Convert a qualified method (qualified with its declaring 
+class) to a qualified method declaration:  @{text qmdecl}  *}
+
+constdefs qmdecl :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> mdecl)"
+"qmdecl sig m \<equiv> (declclass m, (sig,mthd m))"
+
+lemma qmdecl_simp[simp]: "qmdecl sig (C,m) = (C,(sig,m))"
+by (simp add: qmdecl_def)
+
+lemma declclass_qmdecl_simp[simp]: "declclass (qmdecl sig m) = declclass m"
+by (simp add: qmdecl_def)
+
+lemma accmodi_qmdecl_simp[simp]: "accmodi (qmdecl sig m) = accmodi m"
+by (simp add: qmdecl_def)
+
+lemma is_static_qmdecl_simp[simp]: "is_static (qmdecl sig m) = is_static m"
+by (cases m) (simp add: qmdecl_def)
+
+lemma msig_qmdecl_simp[simp]: "msig (qmdecl sig m) = sig"
+by (simp add: qmdecl_def)
+
+lemma mdecl_qmdecl_simp[simp]:  
+ "mdecl (mthd (qmdecl sig new)) = mdecl (sig, mthd new)" 
+by (simp add: qmdecl_def) 
+
+lemma methdMembr_qmdecl_simp [simp]: 
+ "methdMembr (qmdecl sig old) = method sig old"
+by (simp add: methdMembr_def qmdecl_def method_def)
+
+text {* overloaded selector @{text resTy} to select the result type 
+out of various HOL types *}
+
+axclass has_resTy < "type"
+consts resTy:: "'a::has_resTy \<Rightarrow> ty"
+
+instance access_field_type :: ("type","has_resTy") has_resTy
+by (intro_classes)
+
+defs (overloaded)
+decl_resTy_def: 
+ "resTy (m::('a::has_resTy) decl_scheme) \<equiv> resTy (Decl.decl.more m)" 
+
+instance static_field_type :: ("type","has_resTy") has_resTy
+by (intro_classes)
+
+defs (overloaded)
+static_field_type_resTy_def: 
+ "resTy (m::(bool,'b::has_resTy) static_field_type) 
+  \<equiv> resTy (static_more m)" 
+
+instance pars_field_type :: ("type","has_resTy") has_resTy
+by (intro_classes)
+
+defs (overloaded)
+pars_field_type_resTy_def: 
+ "resTy (m::(vname list,'b::has_resTy) pars_field_type) 
+  \<equiv> resTy (pars_more m)" 
+
+instance resT_field_type :: ("type","type") has_resTy
+by (intro_classes)
+
+defs (overloaded)
+resT_field_type_resTy_def: 
+ "resTy (m::(ty,'b::type) resT_field_type) 
+  \<equiv> resT_val m" 
+
+lemma mhead_resTy_simp: "resTy (m::'a mhead_scheme) = resT m"
+apply (cases m)
+apply (simp add: decl_resTy_def static_field_type_resTy_def 
+                 pars_field_type_resTy_def resT_field_type_resTy_def
+                 member.dest_convs mhead.dest_convs)
+done
+
+lemma resTy_mhead [simp]:"resTy (mhead m) = resTy m"
+by (simp add: mhead_def mhead_resTy_simp)
+
+instance * :: ("type","has_resTy") has_resTy
+by (intro_classes)
+
+defs (overloaded)
+pair_resTy_def: "resTy p \<equiv> resTy (snd p)"
+
+lemma pair_resTy_simp[simp]: "resTy (x,m) = resTy m"
+by (simp add: pair_resTy_def)
+
+lemma qmdecl_resTy_simp [simp]: "resTy (qmdecl sig m) = resTy m"
+by (cases m) (simp)
+
+lemma resTy_mthd [simp]:"resTy (mthd m) = resTy m"
+by (cases m) (simp add: mthd_def )
+
+subsubsection "inheritable-in"
+text {*
+@{text "G\<turnstile>m inheritable_in P"}: m can be inherited by
+classes in package P if:
+\begin{itemize} 
+\item the declaration class of m is accessible in P and
+\item the member m is declared with protected or public access or if it is
+      declared with default (package) access, the package of the declaration 
+      class of m is also P. If the member m is declared with private access
+      it is not accessible for inheritance at all.
+\end{itemize}
+*}
+constdefs
+inheritable_in:: 
+ "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool"
+                  ("_ \<turnstile> _ inheritable'_in _" [61,61,61] 60)
+"G\<turnstile>membr inheritable_in pack 
+  \<equiv> (case (accmodi membr) of
+       Private   \<Rightarrow> False
+     | Package   \<Rightarrow> (pid (declclass membr)) = pack
+     | Protected \<Rightarrow> True
+     | Public    \<Rightarrow> True)"
+
+syntax
+Method_inheritable_in::
+ "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> pname \<Rightarrow> bool"
+                ("_ \<turnstile>Method _ inheritable'_in _ " [61,61,61] 60)
+
+translations
+"G\<turnstile>Method m inheritable_in p" == "G\<turnstile>methdMembr m inheritable_in p"
+
+syntax
+Methd_inheritable_in::
+ "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> pname \<Rightarrow> bool"
+                ("_ \<turnstile>Methd _ _ inheritable'_in _ " [61,61,61,61] 60)
+
+translations
+"G\<turnstile>Methd s m inheritable_in p" == "G\<turnstile>(method s m) inheritable_in p"
+
+subsubsection "declared-in/undeclared-in"
+
+constdefs cdeclaredmethd:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,methd) table"
+"cdeclaredmethd G C 
+  \<equiv> (case class G C of
+       None \<Rightarrow> \<lambda> sig. None
+     | Some c \<Rightarrow> table_of (methods c)
+    )"
+
+constdefs
+cdeclaredfield:: "prog \<Rightarrow> qtname \<Rightarrow> (vname,field) table"
+"cdeclaredfield G C 
+  \<equiv> (case class G C of
+       None \<Rightarrow> \<lambda> sig. None
+     | Some c \<Rightarrow> table_of (cfields c)
+    )"
+
+
+constdefs
+declared_in:: "prog  \<Rightarrow> memberdecl \<Rightarrow> qtname \<Rightarrow> bool"
+                                 ("_\<turnstile> _ declared'_in _" [61,61,61] 60)
+"G\<turnstile>m declared_in C \<equiv> (case m of
+                        fdecl (fn,f ) \<Rightarrow> cdeclaredfield G C fn  = Some f
+                      | mdecl (sig,m) \<Rightarrow> cdeclaredmethd G C sig = Some m)"
+
+syntax
+method_declared_in:: "prog  \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
+                                 ("_\<turnstile>Method _ declared'_in _" [61,61,61] 60)
+translations
+"G\<turnstile>Method m declared_in C" == "G\<turnstile>mdecl (mthd m) declared_in C"
+
+syntax
+methd_declared_in:: "prog  \<Rightarrow> sig  \<Rightarrow>(qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
+                               ("_\<turnstile>Methd _  _ declared'_in _" [61,61,61,61] 60)
+translations
+"G\<turnstile>Methd s m declared_in C" == "G\<turnstile>mdecl (s,mthd m) declared_in C"
+
+lemma declared_in_classD:
+ "G\<turnstile>m declared_in C \<Longrightarrow> is_class G C"
+by (cases m) 
+   (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)
+
+constdefs
+undeclared_in:: "prog  \<Rightarrow> memberid \<Rightarrow> qtname \<Rightarrow> bool"
+                                 ("_\<turnstile> _ undeclared'_in _" [61,61,61] 60)
+
+"G\<turnstile>m undeclared_in C \<equiv> (case m of
+                            fid fn  \<Rightarrow> cdeclaredfield G C fn  = None
+                          | mid sig \<Rightarrow> cdeclaredmethd G C sig = None)"
+
+lemma method_declared_inI:
+  "\<lbrakk>class G C = Some c; table_of (methods c) sig = Some m\<rbrakk> 
+   \<Longrightarrow> G\<turnstile>mdecl (sig,m) declared_in C"
+by (auto simp add: declared_in_def cdeclaredmethd_def)
+
+
+subsubsection "members"
+
+consts
+members:: "prog \<Rightarrow> (qtname \<times> (qtname \<times> memberdecl)) set"
+(* Can't just take a function: prog \<Rightarrow> qtname \<Rightarrow> memberdecl set because
+   the class qtname changes to the superclass in the inductive definition
+   below
+*)
+
+syntax
+member_of:: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
+                           ("_ \<turnstile> _ member'_of _" [61,61,61] 60)
+
+translations
+ "G\<turnstile>m member_of C" \<rightleftharpoons> "(C,m) \<in> members G"
+
+inductive "members G"  intros
+
+Immediate: "\<lbrakk>G\<turnstile>mbr m declared_in C;declclass m = C\<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
+Inherited: "\<lbrakk>G\<turnstile>m inheritable_in (pid C); G\<turnstile>memberid m undeclared_in C; 
+             G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S 
+            \<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
+text {* Note that in the case of an inherited member only the members of the
+direct superclass are concerned. If a member of a superclass of the direct
+superclass isn't inherited in the direct superclass (not member of the
+direct superclass) than it can't be a member of the class. E.g. If a
+member of a class A is defined with package access it isn't member of a 
+subclass S if S isn't in the same package as A. Any further subclasses 
+of S will not inherit the member, regardless if they are in the same
+package as A or not.*}
+
+syntax
+method_member_of:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
+                           ("_ \<turnstile>Method _ member'_of _" [61,61,61] 60)
+
+translations
+ "G\<turnstile>Method m member_of C" \<rightleftharpoons> "G\<turnstile>(methdMembr m) member_of C" 
+
+syntax
+methd_member_of:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
+                           ("_ \<turnstile>Methd _ _ member'_of _" [61,61,61,61] 60)
+
+translations
+ "G\<turnstile>Methd s m member_of C" \<rightleftharpoons> "G\<turnstile>(method s m) member_of C" 
+
+syntax
+fieldm_member_of:: "prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> bool"
+                           ("_ \<turnstile>Field _  _ member'_of _" [61,61,61] 60)
+
+translations
+ "G\<turnstile>Field n f member_of C" \<rightleftharpoons> "G\<turnstile>fieldm n f member_of C" 
+
+constdefs
+inherits:: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> bool"
+                           ("_ \<turnstile> _ inherits _" [61,61,61] 60)
+"G\<turnstile>C inherits m 
+  \<equiv> G\<turnstile>m inheritable_in (pid C) \<and> G\<turnstile>memberid m undeclared_in C \<and> 
+    (\<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)"
+
+lemma inherits_member: "G\<turnstile>C inherits m \<Longrightarrow> G\<turnstile>m member_of C"
+by (auto simp add: inherits_def intro: members.Inherited)
+
+
+constdefs member_in::"prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
+                           ("_ \<turnstile> _ member'_in _" [61,61,61] 60)
+"G\<turnstile>m member_in C \<equiv> \<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC"
+text {* A member is in a class if it is member of the class or a superclass.
+If a member is in a class we can select this member. This additional notion
+is necessary since not all members are inherited to subclasses. So such
+members are not member-of the subclass but member-in the subclass.*}
+
+syntax
+method_member_in:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
+                           ("_ \<turnstile>Method _ member'_in _" [61,61,61] 60)
+
+translations
+ "G\<turnstile>Method m member_in C" \<rightleftharpoons> "G\<turnstile>(methdMembr m) member_in C" 
+
+syntax
+methd_member_in:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
+                           ("_ \<turnstile>Methd _ _ member'_in _" [61,61,61,61] 60)
+
+translations
+ "G\<turnstile>Methd s m member_in C" \<rightleftharpoons> "G\<turnstile>(method s m) member_in C" 
+
+consts stat_overridesR::
+  "prog  \<Rightarrow> ((qtname \<times> mdecl) \<times> (qtname \<times> mdecl)) set"
+
+lemma member_inD: "G\<turnstile>m member_in C 
+ \<Longrightarrow> \<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC"
+by (auto simp add: member_in_def)
+
+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"
+by (auto simp add: member_in_def)
+
+lemma member_of_to_member_in: "G \<turnstile> m member_of C \<Longrightarrow> G \<turnstile>m member_in C"
+by (auto intro: member_inI)
+
+
+subsubsection "overriding"
+
+text {* Unfortunately the static notion of overriding (used during the
+typecheck of the compiler) and the dynamic notion of overriding (used during
+execution in the JVM) are not exactly the same. 
+*}
+
+text {* Static overriding (used during the typecheck of the compiler) *}
+syntax
+stat_overrides:: "prog  \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
+                                  ("_ \<turnstile> _ overrides\<^sub>S _" [61,61,61] 60)
+translations
+ "G\<turnstile>new overrides\<^sub>S  old" == "(new,old) \<in> stat_overridesR G "
+
+inductive "stat_overridesR G" intros
+
+Direct: "\<lbrakk>\<not> is_static new; msig new = msig old; 
+         G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
+         G\<turnstile>Method new declared_in (declclass new);  
+         G\<turnstile>Method old declared_in (declclass old); 
+         G\<turnstile>Method old inheritable_in pid (declclass new);
+         G\<turnstile>(declclass new) \<prec>\<^sub>C\<^sub>1 superNew;
+         G \<turnstile>Method old member_of superNew
+         \<rbrakk> \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
+
+Indirect: "\<lbrakk>G\<turnstile>new overrides\<^sub>S inter; G\<turnstile>inter overrides\<^sub>S old\<rbrakk>
+           \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
+
+text {* Dynamic overriding (used during the typecheck of the compiler) *}
+consts overridesR::
+  "prog  \<Rightarrow> ((qtname \<times> mdecl) \<times> (qtname \<times> mdecl)) set"
+
+
+overrides:: "prog  \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
+                                  ("_ \<turnstile> _ overrides _" [61,61,61] 60)
+translations
+ "G\<turnstile>new overrides old" == "(new,old) \<in> overridesR G "
+
+inductive "overridesR G" intros
+
+Direct: "\<lbrakk>\<not> is_static new; \<not> is_static old; accmodi new \<noteq> Private;
+         msig new = msig old; 
+         G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
+         G\<turnstile>Method new declared_in (declclass new);  
+         G\<turnstile>Method old declared_in (declclass old);    
+         G\<turnstile>Method old inheritable_in pid (declclass new);
+         G\<turnstile>resTy new \<preceq> resTy old
+         \<rbrakk> \<Longrightarrow> G\<turnstile>new overrides old"
+
+Indirect: "\<lbrakk>G\<turnstile>new overrides inter; G\<turnstile>inter overrides old\<rbrakk>
+           \<Longrightarrow> G\<turnstile>new overrides old"
+
+syntax
+sig_stat_overrides:: 
+ "prog  \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool" 
+                                  ("_,_\<turnstile> _ overrides\<^sub>S _" [61,61,61,61] 60)
+translations
+ "G,s\<turnstile>new overrides\<^sub>S old" \<rightharpoonup> "G\<turnstile>(qmdecl s new) overrides\<^sub>S (qmdecl s old)" 
+
+syntax
+sig_overrides:: "prog  \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool" 
+                                  ("_,_\<turnstile> _ overrides _" [61,61,61,61] 60)
+translations
+ "G,s\<turnstile>new overrides old" \<rightharpoonup> "G\<turnstile>(qmdecl s new) overrides (qmdecl s old)" 
+
+subsubsection "Hiding"
+
+constdefs hides::
+"prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
+                                ("_\<turnstile> _ hides _" [61,61,61] 60)
+"G\<turnstile>new hides old
+  \<equiv> is_static new \<and> msig new = msig old \<and>
+    G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
+    G\<turnstile>Method new declared_in (declclass new) \<and>
+    G\<turnstile>Method old declared_in (declclass old) \<and> 
+    G\<turnstile>Method old inheritable_in pid (declclass new)"
+
+syntax
+sig_hides:: "prog  \<Rightarrow> sig \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
+                                  ("_,_\<turnstile> _ hides _" [61,61,61,61] 60)
+translations
+ "G,s\<turnstile>new hides old" \<rightharpoonup> "G\<turnstile>(qmdecl s new) hides (qmdecl s old)" 
+
+lemma hidesI:
+"\<lbrakk>is_static new; msig new = msig old;
+  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
+  G\<turnstile>Method new declared_in (declclass new);
+  G\<turnstile>Method old declared_in (declclass old);
+  G\<turnstile>Method old inheritable_in pid (declclass new)
+ \<rbrakk> \<Longrightarrow> G\<turnstile>new hides old"
+by (auto simp add: hides_def)
+
+lemma hidesD:
+"\<lbrakk>G\<turnstile>new hides old\<rbrakk> \<Longrightarrow>  
+  declclass new \<noteq> Object \<and> is_static new \<and> msig new = msig old \<and> 
+  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
+  G\<turnstile>Method new declared_in (declclass new) \<and>   
+  G\<turnstile>Method old declared_in (declclass old)"
+by (auto simp add: hides_def)
+
+lemma overrides_commonD:
+"\<lbrakk>G\<turnstile>new overrides old\<rbrakk> \<Longrightarrow>  
+  declclass new \<noteq> Object \<and> \<not> is_static new \<and> \<not> is_static old \<and>
+  accmodi new \<noteq> Private \<and> 
+  msig new = msig old  \<and>
+  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
+  G\<turnstile>Method new declared_in (declclass new) \<and>   
+  G\<turnstile>Method old declared_in (declclass old)"
+by (induct set: overridesR) (auto intro: trancl_trans)
+
+lemma ws_overrides_commonD:
+"\<lbrakk>G\<turnstile>new overrides old;ws_prog G\<rbrakk> \<Longrightarrow>  
+  declclass new \<noteq> Object \<and> \<not> is_static new \<and> \<not> is_static old \<and>
+  accmodi new \<noteq> Private \<and> G\<turnstile>resTy new \<preceq> resTy old \<and>
+  msig new = msig old  \<and>
+  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
+  G\<turnstile>Method new declared_in (declclass new) \<and>   
+  G\<turnstile>Method old declared_in (declclass old)"
+by (induct set: overridesR) (auto intro: trancl_trans ws_widen_trans)
+
+lemma stat_overrides_commonD:
+"\<lbrakk>G\<turnstile>new overrides\<^sub>S old\<rbrakk> \<Longrightarrow>  
+  declclass new \<noteq> Object \<and> \<not> is_static new \<and> msig new = msig old \<and> 
+  G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
+  G\<turnstile>Method new declared_in (declclass new) \<and>   
+  G\<turnstile>Method old declared_in (declclass old)"
+by (induct set: stat_overridesR) (auto intro: trancl_trans)
+
+lemma overrides_eq_sigD: 
+ "\<lbrakk>G\<turnstile>new overrides old\<rbrakk> \<Longrightarrow> msig old=msig new"
+by (auto dest: overrides_commonD)
+
+lemma hides_eq_sigD: 
+ "\<lbrakk>G\<turnstile>new hides old\<rbrakk> \<Longrightarrow> msig old=msig new"
+by (auto simp add: hides_def)
+
+subsubsection "permits access" 
+constdefs 
+permits_acc:: 
+ "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+                   ("_ \<turnstile> _ in _ permits'_acc'_to _" [61,61,61,61] 60)
+
+"G\<turnstile>membr in class permits_acc_to accclass 
+  \<equiv> (case (accmodi membr) of
+       Private   \<Rightarrow> (declclass membr = accclass)
+     | Package   \<Rightarrow> (pid (declclass membr) = pid accclass)
+     | Protected \<Rightarrow> (pid (declclass membr) = pid accclass)
+                    \<or>
+                    (G\<turnstile>accclass \<prec>\<^sub>C declclass membr \<and> G\<turnstile>class \<preceq>\<^sub>C accclass) 
+     | Public    \<Rightarrow> True)"
+text {*
+The subcondition of the @{term "Protected"} case: 
+@{term "G\<turnstile>accclass \<prec>\<^sub>C declclass membr"} could also be relaxed to:
+@{term "G\<turnstile>accclass \<preceq>\<^sub>C declclass membr"} since in case both classes are the
+same the other condition @{term "(pid (declclass membr) = pid accclass)"}
+holds anyway.
+*} 
+
+text {* Like in case of overriding, the static and dynamic accessibility 
+of members is not uniform.
+\begin{itemize}
+\item Statically the class/interface of the member must be accessible for the
+      member to be accessible. During runtime this is not necessary. For
+      Example, if a class is accessible and we are allowed to access a member
+      of this class (statically) we expect that we can access this member in 
+      an arbitrary subclass (during runtime). It's not intended to restrict
+      the access to accessible subclasses during runtime.
+\item Statically the member we want to access must be "member of" the class.
+      Dynamically it must only be "member in" the class.
+\end{itemize} 
+*} 
+
+
+consts
+accessible_fromR:: 
+ "prog \<Rightarrow> qtname \<Rightarrow> ((qtname \<times> memberdecl) \<times>  qtname) set"
+
+syntax 
+accessible_from:: 
+ "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+                   ("_ \<turnstile> _ of _ accessible'_from _" [61,61,61,61] 60)
+
+translations
+"G\<turnstile>membr of cls accessible_from accclass"  
+ \<rightleftharpoons> "(membr,cls) \<in> accessible_fromR G accclass"
+
+syntax 
+method_accessible_from:: 
+ "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+                   ("_ \<turnstile>Method _ of _ accessible'_from _" [61,61,61,61] 60)
+
+translations
+"G\<turnstile>Method m of cls accessible_from accclass"  
+ \<rightleftharpoons> "G\<turnstile>methdMembr m of cls accessible_from accclass"  
+
+syntax 
+methd_accessible_from:: 
+ "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+                 ("_ \<turnstile>Methd _ _ of _ accessible'_from _" [61,61,61,61,61] 60)
+
+translations
+"G\<turnstile>Methd s m of cls accessible_from accclass"  
+ \<rightleftharpoons> "G\<turnstile>(method s m) of cls accessible_from accclass"  
+
+syntax 
+field_accessible_from:: 
+ "prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+                 ("_ \<turnstile>Field _  _ of _ accessible'_from _" [61,61,61,61,61] 60)
+
+translations
+"G\<turnstile>Field fn f of C accessible_from accclass"  
+ \<rightleftharpoons> "G\<turnstile>(fieldm fn f) of C accessible_from accclass" 
+
+inductive "accessible_fromR G accclass" intros
+immediate:  "\<lbrakk>G\<turnstile>membr member_of class;
+              G\<turnstile>(Class class) accessible_in (pid accclass);
+              G\<turnstile>membr in class permits_acc_to accclass 
+             \<rbrakk> \<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
+
+overriding: "\<lbrakk>G\<turnstile>membr member_of class;
+              G\<turnstile>(Class class) accessible_in (pid accclass);
+              membr=(C,mdecl new);
+              G\<turnstile>(C,new) overrides\<^sub>S old; 
+              G\<turnstile>class \<prec>\<^sub>C sup;
+              G\<turnstile>Method old of sup accessible_from accclass
+             \<rbrakk>\<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
+
+consts
+dyn_accessible_fromR:: 
+ "prog \<Rightarrow> qtname \<Rightarrow> ((qtname \<times> memberdecl) \<times>  qtname) set"
+
+syntax 
+dyn_accessible_from:: 
+ "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+                   ("_ \<turnstile> _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
+
+translations
+"G\<turnstile>membr in C dyn_accessible_from accC"  
+ \<rightleftharpoons> "(membr,C) \<in> dyn_accessible_fromR G accC"
+
+syntax 
+method_dyn_accessible_from:: 
+ "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+                 ("_ \<turnstile>Method _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
+
+translations
+"G\<turnstile>Method m in C dyn_accessible_from accC"  
+ \<rightleftharpoons> "G\<turnstile>methdMembr m in C dyn_accessible_from accC"  
+
+syntax 
+methd_dyn_accessible_from:: 
+ "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+             ("_ \<turnstile>Methd _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60)
+
+translations
+"G\<turnstile>Methd s m in C dyn_accessible_from accC"  
+ \<rightleftharpoons> "G\<turnstile>(method s m) in C dyn_accessible_from accC"  
+
+syntax 
+field_dyn_accessible_from:: 
+ "prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+         ("_ \<turnstile>Field _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60)
+
+translations
+"G\<turnstile>Field fn f in dynC dyn_accessible_from accC"  
+ \<rightleftharpoons> "G\<turnstile>(fieldm fn f) in dynC dyn_accessible_from accC"
+  
+(* #### Testet JVM noch über den Bytecodeverifier hinaus ob der
+ statische Typ accessible ist bevor es den Zugriff erlaubt 
+ \<longrightarrow> Test mit Reflektion\<dots>
+*)
+inductive "dyn_accessible_fromR G accclass" intros
+immediate:  "\<lbrakk>G\<turnstile>membr member_in class;
+              G\<turnstile>membr in class permits_acc_to accclass 
+             \<rbrakk> \<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
+
+overriding: "\<lbrakk>G\<turnstile>membr member_in class;
+              membr=(C,mdecl new);
+              G\<turnstile>(C,new) overrides old; 
+              G\<turnstile>class \<prec>\<^sub>C sup;
+              G\<turnstile>Method old in sup dyn_accessible_from accclass
+             \<rbrakk>\<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
+
+
+lemma accessible_from_commonD: "G\<turnstile>m of C accessible_from S
+ \<Longrightarrow> G\<turnstile>m member_of C \<and> G\<turnstile>(Class C) accessible_in (pid S)"
+by (auto elim: accessible_fromR.induct)
+
+lemma declared_not_undeclared:
+  "G\<turnstile>m declared_in C \<Longrightarrow> \<not> G\<turnstile> memberid m undeclared_in C"
+by (cases m) (auto simp add: declared_in_def undeclared_in_def)
+
+lemma not_undeclared_declared:
+  "\<not> G\<turnstile> membr_id undeclared_in C \<Longrightarrow> (\<exists> m. G\<turnstile>m declared_in C \<and> 
+                                           membr_id = memberid m)"
+proof -
+  assume not_undecl:"\<not> G\<turnstile> membr_id undeclared_in C"
+  show ?thesis (is "?P membr_id")
+  proof (cases membr_id)
+    case (fid vname)
+    with not_undecl
+    obtain fld where
+      "G\<turnstile>fdecl (vname,fld) declared_in C" 
+      by (auto simp add: undeclared_in_def declared_in_def
+                         cdeclaredfield_def)
+    with fid show ?thesis 
+      by auto
+  next
+    case (mid sig)
+    with not_undecl
+    obtain mthd where
+      "G\<turnstile>mdecl (sig,mthd) declared_in C" 
+      by (auto simp add: undeclared_in_def declared_in_def
+                         cdeclaredmethd_def)
+    with mid show ?thesis 
+      by auto
+  qed
+qed
+
+lemma unique_declared_in:
+ "\<lbrakk>G\<turnstile>m declared_in C; G\<turnstile>n declared_in C; memberid m = memberid n\<rbrakk>
+ \<Longrightarrow> m = n"
+by (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def
+            split: memberdecl.splits)
+
+lemma unique_member_of: 
+ (assumes n: "G\<turnstile>n member_of C" and  
+          m: "G\<turnstile>m member_of C" and
+       eqid: "memberid n = memberid m"
+ ) "n=m"
+proof -
+  from n m eqid  
+  show "n=m"
+  proof (induct)
+    case (Immediate C n)
+    assume member_n: "G\<turnstile> mbr n declared_in C" "declclass n = C"
+    assume eqid: "memberid n = memberid m"
+    assume "G \<turnstile> m member_of C"
+    then show "n=m"
+    proof (cases)
+      case (Immediate _ m')
+      with eqid 
+      have "m=m'"
+           "memberid n = memberid m" 
+           "G\<turnstile> mbr m declared_in C" 
+           "declclass m = C"
+	by auto
+      with member_n   
+      show ?thesis
+	by (cases n, cases m) 
+           (auto simp add: declared_in_def 
+                           cdeclaredmethd_def cdeclaredfield_def
+                    split: memberdecl.splits)
+    next
+      case (Inherited _ _ m')
+      then have "G\<turnstile> memberid m undeclared_in C"
+	by simp
+      with eqid member_n
+      show ?thesis
+	by (cases n) (auto dest: declared_not_undeclared)
+    qed
+  next
+    case (Inherited C S n)
+    assume undecl: "G\<turnstile> memberid n undeclared_in C"
+    assume  super: "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S"
+    assume    hyp: "\<lbrakk>G \<turnstile> m member_of S; memberid n = memberid m\<rbrakk> \<Longrightarrow> n = m"
+    assume   eqid: "memberid n = memberid m"
+    assume "G \<turnstile> m member_of C"
+    then show "n=m"
+    proof (cases)
+      case Immediate
+      then have "G\<turnstile> mbr m declared_in C" by simp
+      with eqid undecl
+      show ?thesis
+	by (cases m) (auto dest: declared_not_undeclared)
+    next
+      case Inherited 
+      with super have "G \<turnstile> m member_of S"
+	by (auto dest!: subcls1D)
+      with eqid hyp
+      show ?thesis 
+	by blast
+    qed
+  qed
+qed
+
+lemma member_of_is_classD: "G\<turnstile>m member_of C \<Longrightarrow> is_class G C"
+proof (induct set: members)
+  case (Immediate C m)
+  assume "G\<turnstile> mbr m declared_in C"
+  then show "is_class G C"
+    by (cases "mbr m")
+       (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)
+next
+  case (Inherited C S m)  
+  assume "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S" and "is_class G S"
+  then show "is_class G C"
+    by - (rule subcls_is_class2,auto)
+qed
+
+lemma member_of_declC: 
+ "G\<turnstile>m member_of C 
+  \<Longrightarrow> G\<turnstile>mbr m declared_in (declclass m)"
+by (induct set: members) auto
+
+lemma member_of_member_of_declC:
+ "G\<turnstile>m member_of C 
+  \<Longrightarrow> G\<turnstile>m member_of (declclass m)"
+by (auto dest: member_of_declC intro: members.Immediate)
+
+lemma member_of_class_relation:
+  "G\<turnstile>m member_of C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
+proof (induct set: members)
+  case (Immediate C m)
+  then show "G\<turnstile>C \<preceq>\<^sub>C declclass m" by simp
+next
+  case (Inherited C S m)
+  then show "G\<turnstile>C \<preceq>\<^sub>C declclass m" 
+    by (auto dest: r_into_rtrancl intro: rtrancl_trans)
+qed
+
+lemma member_in_class_relation:
+  "G\<turnstile>m member_in C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
+by (auto dest: member_inD member_of_class_relation
+        intro: rtrancl_trans)
+
+lemma member_of_Package: 
+ "\<lbrakk>G\<turnstile>m member_of C; accmodi m = Package\<rbrakk> 
+  \<Longrightarrow> pid (declclass m) = pid C" 
+proof -
+  assume   member: "G\<turnstile>m member_of C"
+  then show " accmodi m = Package \<Longrightarrow> ?thesis" (is "PROP ?P m C")
+  proof (induct rule: members.induct)
+    fix C m
+    assume     C: "declclass m = C"
+    then show "pid (declclass m) = pid C"
+      by simp
+  next
+    fix C S m  
+    assume inheritable: "G \<turnstile> m inheritable_in pid C"
+    assume         hyp: "PROP ?P m S" and
+           package_acc: "accmodi m = Package" 
+    with inheritable package_acc hyp
+    show "pid (declclass m) = pid C" 
+      by (auto simp add: inheritable_in_def)
+  qed
+qed
+
+lemma dyn_accessible_from_commonD: "G\<turnstile>m in C dyn_accessible_from S
+ \<Longrightarrow> G\<turnstile>m member_in C"
+by (auto elim: dyn_accessible_fromR.induct)
+
+(* ### Gilt nicht für wf_progs!dynmaisches Override, 
+  da die accmodi Bedingung nur für stat override gilt! *)
+(*
+lemma override_Package:
+ "\<lbrakk>G\<turnstile>new overrides old; 
+  \<And> new old. G\<turnstile>new overrides old \<Longrightarrow> accmodi old \<le> accmodi new;
+  accmodi old = Package; accmodi new = Package\<rbrakk>
+  \<Longrightarrow> pid (declclass old) = pid (declclass new)"
+proof - 
+  assume wf: "\<And> new old. G\<turnstile>new overrides old \<Longrightarrow> accmodi old \<le> accmodi new"
+  assume ovverride: "G\<turnstile>new overrides old"
+  then show "\<lbrakk>accmodi old = Package;accmodi new = Package\<rbrakk> \<Longrightarrow> ?thesis"
+    (is "?Pack old \<Longrightarrow> ?Pack new \<Longrightarrow> ?EqPid old new")
+  proof (induct rule: overridesR.induct)
+    case Direct
+    fix new old
+    assume "accmodi old = Package"
+           "G \<turnstile> methdMembr old inheritable_in pid (declclass new)"
+    then show "pid (declclass old) =  pid (declclass new)"
+      by (auto simp add: inheritable_in_def)
+  next
+    case (Indirect inter new old)
+    assume accmodi_old: "accmodi old = Package" and
+           accmodi_new: "accmodi new = Package" 
+    assume "G \<turnstile> new overrides inter"
+    with wf have le_inter_new: "accmodi inter \<le> accmodi new"
+      by blast
+    assume "G \<turnstile> inter overrides old"
+    with wf have le_old_inter: "accmodi old \<le> accmodi inter"
+      by blast
+    from accmodi_old accmodi_new le_inter_new le_old_inter
+    have "accmodi inter = Package"
+      by(auto simp add: le_acc_def less_acc_def)
+    with Indirect accmodi_old accmodi_new
+    show "?EqPid old new"
+      by auto
+  qed
+qed
+
+lemma stat_override_Package:
+ "\<lbrakk>G\<turnstile>new overrides\<^sub>S old; 
+  \<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new;
+  accmodi old = Package; accmodi new = Package\<rbrakk>
+  \<Longrightarrow> pid (declclass old) = pid (declclass new)"
+proof - 
+  assume wf: "\<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new"
+  assume ovverride: "G\<turnstile>new overrides\<^sub>S old"
+  then show "\<lbrakk>accmodi old = Package;accmodi new = Package\<rbrakk> \<Longrightarrow> ?thesis"
+    (is "?Pack old \<Longrightarrow> ?Pack new \<Longrightarrow> ?EqPid old new")
+  proof (induct rule: stat_overridesR.induct)
+    case Direct
+    fix new old
+    assume "accmodi old = Package"
+           "G \<turnstile> methdMembr old inheritable_in pid (declclass new)"
+    then show "pid (declclass old) =  pid (declclass new)"
+      by (auto simp add: inheritable_in_def)
+  next
+    case (Indirect inter new old)
+    assume accmodi_old: "accmodi old = Package" and
+           accmodi_new: "accmodi new = Package" 
+    assume "G \<turnstile> new overrides\<^sub>S inter"
+    with wf have le_inter_new: "accmodi inter \<le> accmodi new"
+      by blast
+    assume "G \<turnstile> inter overrides\<^sub>S old"
+    with wf have le_old_inter: "accmodi old \<le> accmodi inter"
+      by blast
+    from accmodi_old accmodi_new le_inter_new le_old_inter
+    have "accmodi inter = Package"
+      by(auto simp add: le_acc_def less_acc_def)
+    with Indirect accmodi_old accmodi_new
+    show "?EqPid old new"
+      by auto
+  qed
+qed
+
+*)
+lemma no_Private_stat_override: 
+ "\<lbrakk>G\<turnstile>new overrides\<^sub>S old\<rbrakk> \<Longrightarrow> accmodi old \<noteq> Private"
+by (induct set:  stat_overridesR) (auto simp add: inheritable_in_def)
+
+lemma no_Private_override: "\<lbrakk>G\<turnstile>new overrides old\<rbrakk> \<Longrightarrow> accmodi old \<noteq> Private"
+by (induct set: overridesR) (auto simp add: inheritable_in_def)
+
+lemma permits_acc_inheritance:
+ "\<lbrakk>G\<turnstile>m in statC permits_acc_to accC; G\<turnstile>dynC \<preceq>\<^sub>C statC
+  \<rbrakk> \<Longrightarrow> G\<turnstile>m in dynC permits_acc_to accC"
+by (cases "accmodi m")
+   (auto simp add: permits_acc_def
+            intro: subclseq_trans) 
+
+lemma field_accessible_fromD:
+ "\<lbrakk>G\<turnstile>membr of C accessible_from accC;is_field membr\<rbrakk> 
+  \<Longrightarrow> G\<turnstile>membr member_of C \<and>
+      G\<turnstile>(Class C) accessible_in (pid accC) \<and>
+      G\<turnstile>membr in C permits_acc_to accC"
+by (cases set: accessible_fromR)
+   (auto simp add: is_field_def split: memberdecl.splits) 
+
+lemma field_accessible_from_permits_acc_inheritance:
+"\<lbrakk>G\<turnstile>membr of statC accessible_from accC; is_field membr; G \<turnstile> dynC \<preceq>\<^sub>C statC\<rbrakk>
+\<Longrightarrow> G\<turnstile>membr in dynC permits_acc_to accC"
+by (auto dest: field_accessible_fromD intro: permits_acc_inheritance)
+
+
+(*
+lemma accessible_Package:
+ "\<lbrakk>G \<turnstile> m of C accessible_from S;accmodi m = Package;
+   \<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new\<rbrakk>
+  \<Longrightarrow> pid S = pid C \<and> pid C = pid (declclass m)"
+proof -
+  assume wf: "\<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new"
+  assume "G \<turnstile> m of C accessible_from S"
+  then show "accmodi m = Package \<Longrightarrow> pid S = pid C \<and> pid C = pid (declclass m)"
+    (is "?Pack m \<Longrightarrow> ?P C m")
+  proof (induct rule: accessible_fromR.induct)
+    fix C m
+    assume "G\<turnstile>m member_of C"
+           "G \<turnstile> m in C permits_acc_to S"
+           "accmodi m = Package"      
+    then show "?P C m"
+      by (auto dest: member_of_Package simp add: permits_acc_def)
+  next
+    fix declC C new newm old Sup
+    assume member_new: "G \<turnstile> new member_of C" and 
+                acc_C: "G \<turnstile> Class C accessible_in pid S" and
+                  new: "new = (declC, mdecl newm)" and
+             override: "G \<turnstile> (declC, newm) overrides\<^sub>S old" and
+         subcls_C_Sup: "G\<turnstile>C \<prec>\<^sub>C Sup" and
+              acc_old: "G \<turnstile> methdMembr old of Sup accessible_from S" and
+                  hyp: "?Pack (methdMembr old) \<Longrightarrow> ?P Sup (methdMembr old)" and
+          accmodi_new: "accmodi new = Package"
+    from override wf 
+    have accmodi_weaken: "accmodi old \<le> accmodi newm"
+      by (cases old,cases newm) auto
+    from override new
+    have "accmodi old \<noteq> Private"
+      by (simp add: no_Private_stat_override)
+    with accmodi_weaken accmodi_new new
+    have accmodi_old: "accmodi old = Package"
+      by (cases "accmodi old") (auto simp add: le_acc_def less_acc_def) 
+    with hyp 
+    have P_sup: "?P Sup (methdMembr old)"
+      by (simp)
+    from wf override new accmodi_old accmodi_new
+    have eq_pid_new_old: "pid (declclass new) = pid (declclass old)"
+      by (auto dest: stat_override_Package) 
+    from member_new accmodi_new
+    have "pid (declclass new) = pid C"
+      by (auto dest: member_of_Package)
+    with eq_pid_new_old P_sup show "?P C new"
+      by auto
+  qed
+qed
+*)
+lemma accessible_fieldD: 
+ "\<lbrakk>G\<turnstile>membr of C accessible_from accC; is_field membr\<rbrakk>
+ \<Longrightarrow> G\<turnstile>membr member_of C \<and>
+     G\<turnstile>(Class C) accessible_in (pid accC) \<and>
+     G\<turnstile>membr in C permits_acc_to accC"
+by (induct rule: accessible_fromR.induct) (auto dest: is_fieldD)
+      
+(* lemmata:
+ Wegen  G\<turnstile>Super accessible_in (pid C) folgt:
+  G\<turnstile>m declared_in C; G\<turnstile>m member_of D; accmodi m = Package (G\<turnstile>D \<preceq>\<^sub>C C)
+  \<Longrightarrow> pid C = pid D 
+
+  C package
+  m public in C
+  für alle anderen D: G\<turnstile>m undeclared_in C
+  m wird in alle subklassen vererbt, auch aus dem Package heraus!
+
+  G\<turnstile>m member_of C \<Longrightarrow> \<exists> D. G\<turnstile>C \<preceq>\<^sub>C D \<and> G\<turnstile>m declared_in D
+*)
+
+(* Begriff (C,m) overrides (D,m)
+    3 Fälle: Direkt,
+             Indirekt über eine Zwischenklasse (ohne weiteres override)
+             Indirekt über override
+*)
+   
+(*
+"G\<turnstile>m member_of C \<equiv> 
+constdefs declares_method:: "prog \<Rightarrow> sig \<Rightarrow> qtname \<Rightarrow> methd \<Rightarrow> bool"
+                                 ("_,_\<turnstile> _ declares'_method _" [61,61,61,61] 60)
+"G,sig\<turnstile>C declares_method m \<equiv> cdeclaredmethd G C sig = Some m" 
+
+constdefs is_declared:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool"
+"is_declared G sig em \<equiv> G,sig\<turnstile>declclass em declares_method mthd em"
+*)
+
+lemma member_of_Private:
+"\<lbrakk>G\<turnstile>m member_of C; accmodi m = Private\<rbrakk> \<Longrightarrow> declclass m = C"
+by (induct set: members) (auto simp add: inheritable_in_def)
+
+lemma member_of_subclseq_declC:
+  "G\<turnstile>m member_of C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
+by (induct set: members) (auto dest: r_into_rtrancl intro: rtrancl_trans)
+
+lemma member_of_inheritance:
+  (assumes    m: "G\<turnstile>m member_of D" and 
+     subclseq_D_C: "G\<turnstile>D \<preceq>\<^sub>C C" and
+     subclseq_C_m: "G\<turnstile>C \<preceq>\<^sub>C declclass m" and
+               ws: "ws_prog G" 
+  ) "G\<turnstile>m member_of C"
+proof -
+  from m subclseq_D_C subclseq_C_m
+  show ?thesis
+  proof (induct)
+    case (Immediate D m)
+    assume "declclass m = D" and
+           "G\<turnstile>D\<preceq>\<^sub>C C" and "G\<turnstile>C\<preceq>\<^sub>C declclass m"
+    with ws have "D=C" 
+      by (auto intro: subclseq_acyclic)
+    with Immediate 
+    show "G\<turnstile>m member_of C"
+      by (auto intro: members.Immediate)
+  next
+    case (Inherited D S m)
+    assume member_of_D_props: 
+            "G \<turnstile> m inheritable_in pid D" 
+            "G\<turnstile> memberid m undeclared_in D"  
+            "G \<turnstile> Class S accessible_in pid D" 
+            "G \<turnstile> m member_of S"
+    assume super: "G\<turnstile>D\<prec>\<^sub>C\<^sub>1S"
+    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"
+    assume subclseq_C_m: "G\<turnstile>C\<preceq>\<^sub>C declclass m"
+    assume "G\<turnstile>D\<preceq>\<^sub>C C"
+    then show "G\<turnstile>m member_of C"
+    proof (cases rule: subclseq_cases)
+      case Eq
+      assume "D=C" 
+      with super member_of_D_props 
+      show ?thesis
+	by (auto intro: members.Inherited)
+    next
+      case Subcls
+      assume "G\<turnstile>D\<prec>\<^sub>C C"
+      with super 
+      have "G\<turnstile>S\<preceq>\<^sub>C C"
+	by (auto dest: subcls1D subcls_superD)
+      with subclseq_C_m hyp show ?thesis
+	by blast
+    qed
+  qed
+qed
+
+lemma member_of_subcls:
+  (assumes    old: "G\<turnstile>old member_of C" and 
+              new: "G\<turnstile>new member_of D" and
+             eqid: "memberid new = memberid old" and
+     subclseq_D_C: "G\<turnstile>D \<preceq>\<^sub>C C" and 
+   subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old" and
+               ws: "ws_prog G"
+  ) "G\<turnstile>D \<prec>\<^sub>C C"
+proof -
+  from old 
+  have subclseq_C_old: "G\<turnstile>C \<preceq>\<^sub>C declclass old"
+    by (auto dest: member_of_subclseq_declC)
+  from new 
+  have subclseq_D_new: "G\<turnstile>D \<preceq>\<^sub>C declclass new"
+    by (auto dest: member_of_subclseq_declC)
+  from subcls_new_old ws
+  have neq_new_old: "new\<noteq>old"
+    by (cases new,cases old) (auto dest: subcls_irrefl)
+  from subclseq_D_new subclseq_D_C
+  have "G\<turnstile>(declclass new) \<preceq>\<^sub>C C \<or> G\<turnstile>C \<preceq>\<^sub>C (declclass new)" 
+    by (rule subcls_compareable)
+  then have "G\<turnstile>(declclass new) \<preceq>\<^sub>C C"
+  proof
+    assume "G\<turnstile>declclass new\<preceq>\<^sub>C C" then show ?thesis .
+  next
+    assume "G\<turnstile>C \<preceq>\<^sub>C (declclass new)"
+    with new subclseq_D_C ws 
+    have "G\<turnstile>new member_of C"
+      by (blast intro: member_of_inheritance)
+    with eqid old 
+    have "new=old"
+      by (blast intro: unique_member_of)
+    with neq_new_old 
+    show ?thesis
+      by contradiction
+  qed
+  then show ?thesis
+  proof (cases rule: subclseq_cases)
+    case Eq
+    assume "declclass new = C"
+    with new have "G\<turnstile>new member_of C"
+      by (auto dest: member_of_member_of_declC)
+    with eqid old 
+    have "new=old"
+      by (blast intro: unique_member_of)
+    with neq_new_old 
+    show ?thesis
+      by contradiction
+  next
+    case Subcls
+    assume "G\<turnstile>declclass new\<prec>\<^sub>C C"
+    with subclseq_D_new
+    show "G\<turnstile>D\<prec>\<^sub>C C"
+      by (rule rtrancl_trancl_trancl)
+  qed
+qed
+
+corollary member_of_overrides_subcls:
+ "\<lbrakk>G\<turnstile>Methd sig old member_of C; G\<turnstile>Methd sig new member_of D;G\<turnstile>D \<preceq>\<^sub>C C;
+   G,sig\<turnstile>new overrides old; ws_prog G\<rbrakk>
+ \<Longrightarrow> G\<turnstile>D \<prec>\<^sub>C C"
+by (drule overrides_commonD) (auto intro: member_of_subcls)    
+
+corollary member_of_stat_overrides_subcls:
+ "\<lbrakk>G\<turnstile>Methd sig old member_of C; G\<turnstile>Methd sig new member_of D;G\<turnstile>D \<preceq>\<^sub>C C;
+   G,sig\<turnstile>new overrides\<^sub>S old; ws_prog G\<rbrakk>
+ \<Longrightarrow> G\<turnstile>D \<prec>\<^sub>C C"
+by (drule stat_overrides_commonD) (auto intro: member_of_subcls)    
+
+
+
+lemma inherited_field_access: 
+ (assumes stat_acc: "G\<turnstile>membr of statC accessible_from accC" and
+          is_field: "is_field membr" and 
+          subclseq: "G \<turnstile> dynC \<preceq>\<^sub>C statC"
+ ) "G\<turnstile>membr in dynC dyn_accessible_from accC"
+proof -
+  from stat_acc is_field subclseq 
+  show ?thesis
+    by (auto dest: accessible_fieldD 
+            intro: dyn_accessible_fromR.immediate 
+                   member_inI 
+                   permits_acc_inheritance)
+qed
+
+lemma accessible_inheritance:
+ (assumes stat_acc: "G\<turnstile>m of statC accessible_from accC" and
+          subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
+       member_dynC: "G\<turnstile>m member_of dynC" and
+          dynC_acc: "G\<turnstile>(Class dynC) accessible_in (pid accC)"
+ ) "G\<turnstile>m of dynC accessible_from accC"
+proof -
+  from stat_acc
+  have member_statC: "G\<turnstile>m member_of statC" 
+    by (auto dest: accessible_from_commonD)
+  from stat_acc
+  show ?thesis
+  proof (cases)
+    case immediate
+    with member_dynC member_statC subclseq dynC_acc
+    show ?thesis
+      by (auto intro: accessible_fromR.immediate permits_acc_inheritance)
+  next
+    case overriding
+    with member_dynC subclseq dynC_acc
+    show ?thesis
+      by (auto intro: accessible_fromR.overriding rtrancl_trancl_trancl)
+  qed
+qed
+
+section "fields and methods"
+
+
+types
+  fspec = "vname \<times> qtname"
+
+translations 
+  "fspec" <= (type) "vname \<times> qtname" 
+
+constdefs
+imethds:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
+"imethds G I 
+  \<equiv> iface_rec (G,I)  
+              (\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus> 
+                        (o2s \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
+text {* methods of an interface, with overriding and inheritance, cf. 9.2 *}
+
+constdefs
+accimethds:: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
+"accimethds G pack I
+  \<equiv> if G\<turnstile>Iface I accessible_in pack 
+       then imethds G I
+       else \<lambda> k. {}"
+text {* only returns imethds if the interface is accessible *}
+
+constdefs
+methd:: "prog \<Rightarrow> qtname  \<Rightarrow> (sig,qtname \<times> methd) table"
+
+"methd G C 
+ \<equiv> class_rec (G,C) empty
+             (\<lambda>C c subcls_mthds. 
+               filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)
+                          subcls_mthds 
+               ++ 
+               table_of (map (\<lambda>(s,m). (s,C,m)) (methods c)))"
+text {* @{term "methd G C"}: methods of a class C (statically visible from C), 
+     with inheritance and hiding cf. 8.4.6;
+     Overriding is captured by @{text dynmethd}.
+     Every new method with the same signature coalesces the
+     method of a superclass. *}
+
+constdefs                      
+accmethd:: "prog \<Rightarrow> qtname \<Rightarrow> qtname  \<Rightarrow> (sig,qtname \<times> methd) table"
+"accmethd G S C 
+ \<equiv> filter_tab (\<lambda>sig m. G\<turnstile>method sig m of C accessible_from S) 
+              (methd G C)"
+text {* @{term "accmethd G S C"}: only those methods of @{term "methd G C"}, 
+        accessible from S *}
+
+text {* Note the class component in the accessibility filter. The class where
+    method @{term m} is declared (@{term declC}) isn't necessarily accessible 
+    from the current scope @{term S}. The method can be made accessible 
+    through inheritance, too.
+    So we must test accessibility of method @{term m} of class @{term C} 
+    (not @{term "declclass m"}) *}
+
+constdefs 
+dynmethd:: "prog  \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+"dynmethd G statC dynC  
+  \<equiv> \<lambda> sig. 
+     (if G\<turnstile>dynC \<preceq>\<^sub>C statC
+        then (case methd G statC sig of
+                None \<Rightarrow> None
+              | Some statM 
+                  \<Rightarrow> (class_rec (G,dynC) empty
+                       (\<lambda>C c subcls_mthds. 
+                          subcls_mthds
+                          ++
+                          (filter_tab 
+                            (\<lambda> _ dynM. G,sig\<turnstile>dynM overrides statM \<or> dynM=statM)
+                            (methd G C) ))
+                      ) sig
+              )
+        else None)"
+(*
+"dynmethd G statC dynC  
+  \<equiv> \<lambda> sig. 
+     (if G\<turnstile>dynC \<preceq>\<^sub>C statC
+        then (case methd G statC sig of
+                None \<Rightarrow> None
+              | Some statM 
+                    \<Rightarrow> (class_rec (G,statC) empty
+                         (\<lambda>C c subcls_mthds. 
+                            subcls_mthds
+                            ++
+                            (filter_tab 
+                              (\<lambda> _ dynM. G,sig\<turnstile>dynM overrides statM)  
+                              (table_of (map (\<lambda>(s,m). (s,C,m)) (methods c)))))
+                        ) sig
+              )
+        else None)"*)
+text {* @{term "dynmethd G statC dynC"}: dynamic method lookup of a reference 
+        with dynamic class @{term dynC} and static class @{term statC} *}
+text {* Note some kind of duality between @{term methd} and @{term dynmethd} 
+        in the @{term class_rec} arguments. Whereas @{term methd} filters the 
+        subclass methods (to get only the inherited ones), @{term dynmethd} 
+        filters the new methods (to get only those methods which actually
+        override the methods of the static class) *}
+
+constdefs 
+dynimethd:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+"dynimethd G I dynC 
+  \<equiv> \<lambda> sig. if imethds G I sig \<noteq> {}
+               then methd G dynC sig
+               else dynmethd G Object dynC sig"
+text {* @{term "dynimethd G I dynC"}: dynamic method lookup of a reference with 
+        dynamic class dynC and static interface type I *}
+text {* 
+   When calling an interface method, we must distinguish if the method signature
+   was defined in the interface or if it must be an Object method in the other
+   case. If it was an interface method we search the class hierarchy
+   starting at the dynamic class of the object up to Object to find the 
+   first matching method (@{term methd}). Since all interface methods have 
+   public access the method can't be coalesced due to some odd visibility 
+   effects like in case of dynmethd. The method will be inherited or 
+   overridden in all classes from the first class implementing the interface 
+   down to the actual dynamic class.
+ *}
+
+constdefs
+dynlookup::"prog  \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+"dynlookup G statT dynC
+  \<equiv> (case statT of
+       NullT        \<Rightarrow> empty
+     | IfaceT I     \<Rightarrow> dynimethd G I      dynC
+     | ClassT statC \<Rightarrow> dynmethd  G statC  dynC
+     | ArrayT ty    \<Rightarrow> dynmethd  G Object dynC)"
+text {* @{term "dynlookup G statT dynC"}: dynamic lookup of a method within the 
+    static reference type statT and the dynamic class dynC. 
+    In a wellformd context statT will not be NullT and in case
+    statT is an array type, dynC=Object *}
+
+constdefs
+fields:: "prog \<Rightarrow> qtname \<Rightarrow> ((vname \<times> qtname) \<times> field) list"
+"fields G C 
+  \<equiv> class_rec (G,C) [] (\<lambda>C c ts. map (\<lambda>(n,t). ((n,C),t)) (cfields c) @ ts)"
+text {* @{term "fields G C"} 
+     list of fields of a class, including all the fields of the superclasses
+     (private, inherited and hidden ones) not only the accessible ones
+     (an instance of a object allocates all these fields *}
+
+constdefs
+accfield:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (vname, qtname  \<times>  field) table"
+"accfield G S C
+  \<equiv> let field_tab = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (fields G C))
+    in filter_tab (\<lambda>n (declC,f). G\<turnstile> (declC,fdecl (n,f)) of C accessible_from S)
+                  field_tab"
+text  {* @{term "accfield G C S"}: fields of a class @{term C} which are 
+         accessible from scope of class
+         @{term S} with inheritance and hiding, cf. 8.3 *}
+text {* note the class component in the accessibility filter (see also 
+        @{term methd}).
+   The class declaring field @{term f} (@{term declC}) isn't necessarily 
+   accessible from scope @{term S}. The field can be made visible through 
+   inheritance, too. So we must test accessibility of field @{term f} of class 
+   @{term C} (not @{term "declclass f"}) *} 
+
+constdefs
+
+  is_methd :: "prog \<Rightarrow> qtname  \<Rightarrow> sig \<Rightarrow> bool"
+ "is_methd G \<equiv> \<lambda>C sig. is_class G C \<and> methd G C sig \<noteq> None"
+
+constdefs efname:: "((vname \<times> qtname) \<times> field) \<Rightarrow> (vname \<times> qtname)"
+"efname \<equiv> fst"
+
+lemma efname_simp[simp]:"efname (n,f) = n"
+by (simp add: efname_def) 
+
+
+subsection "imethds"
+
+lemma imethds_rec: "\<lbrakk>iface G I = Some i; ws_prog G\<rbrakk> \<Longrightarrow>  
+  imethds G I = Un_tables ((\<lambda>J. imethds  G J)`set (isuperIfs i)) \<oplus>\<oplus>  
+                      (o2s \<circ> table_of (map (\<lambda>(s,mh). (s,I,mh)) (imethods i)))"
+apply (unfold imethds_def)
+apply (rule iface_rec [THEN trans])
+apply auto
+done
+
+
+(* local lemma *)
+lemma imethds_norec:
+  "\<lbrakk>iface G md = Some i; ws_prog G; table_of (imethods i) sig = Some mh\<rbrakk> \<Longrightarrow>  
+  (md, mh) \<in> imethds G md sig"
+apply (subst imethds_rec)
+apply assumption+
+apply (rule iffD2)
+apply (rule overrides_t_Some_iff)
+apply (rule disjI1)
+apply (auto elim: table_of_map_SomeI)
+done
+
+lemma imethds_declI: "\<lbrakk>m \<in> imethds G I sig; ws_prog G; is_iface G I\<rbrakk> \<Longrightarrow> 
+  (\<exists>i. iface G (decliface m) = Some i \<and> 
+  table_of (imethods i) sig = Some (mthd m)) \<and>  
+  (I,decliface m) \<in> (subint1 G)^* \<and> m \<in> imethds G (decliface m) sig"
+apply (erule make_imp)
+apply (rule ws_subint1_induct, assumption, assumption)
+apply (subst imethds_rec, erule conjunct1, assumption)
+apply (force elim: imethds_norec intro: rtrancl_into_rtrancl2)
+done
+
+lemma imethds_cases [consumes 3, case_names NewMethod InheritedMethod]:
+ (assumes im: "im \<in> imethds G I sig" and  
+         ifI: "iface G I = Some i" and
+          ws: "ws_prog G" and
+     hyp_new:  "table_of (map (\<lambda>(s, mh). (s, I, mh)) (imethods i)) sig 
+                = Some im \<Longrightarrow> P" and
+     hyp_inh: "\<And> J. \<lbrakk>J \<in> set (isuperIfs i); im \<in> imethds G J sig\<rbrakk> \<Longrightarrow> P"
+  ) "P"
+proof -
+  from ifI ws im hyp_new hyp_inh
+  show "P"
+    by (auto simp add: imethds_rec)
+qed
+
+subsection "accimethd"
+
+lemma accimethds_simp [simp]: 
+"G\<turnstile>Iface I accessible_in pack \<Longrightarrow> accimethds G pack I = imethds G I"
+by (simp add: accimethds_def)
+
+lemma accimethdsD:
+ "im \<in> accimethds G pack I sig 
+  \<Longrightarrow> im \<in> imethds G I sig \<and> G\<turnstile>Iface I accessible_in pack"
+by (auto simp add: accimethds_def)
+
+lemma accimethdsI: 
+"\<lbrakk>im \<in> imethds G I sig;G\<turnstile>Iface I accessible_in pack\<rbrakk>
+ \<Longrightarrow> im \<in> accimethds G pack I sig"
+by simp
+
+subsection "methd"
+
+lemma methd_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
+  methd G C 
+    = (if C = Object 
+          then empty 
+          else filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)
+                          (methd G (super c))) 
+      ++ table_of (map (\<lambda>(s,m). (s,C,m)) (methods c))"
+apply (unfold methd_def)
+apply (erule class_rec [THEN trans], assumption)
+apply (simp)
+done
+
+(* local lemma *)
+lemma methd_norec: 
+ "\<lbrakk>class G declC = Some c; ws_prog G;table_of (methods c) sig = Some m\<rbrakk> 
+  \<Longrightarrow> methd G declC sig = Some (declC, m)"
+apply (simp only: methd_rec)
+apply (rule disjI1 [THEN override_Some_iff [THEN iffD2]])
+apply (auto elim: table_of_map_SomeI)
+done
+
+
+lemma methd_declC: 
+"\<lbrakk>methd G C sig = Some m; ws_prog G;is_class G C\<rbrakk> \<Longrightarrow>
+ (\<exists>d. class G (declclass m)=Some d \<and> table_of (methods d) sig=Some (mthd m)) \<and> 
+ G\<turnstile>C \<preceq>\<^sub>C (declclass m) \<and> methd G (declclass m) sig = Some m"   
+apply (erule make_imp)
+apply (rule ws_subcls1_induct, assumption, assumption)
+apply (subst methd_rec, assumption)
+apply (case_tac "Ca=Object")
+apply   (force elim: methd_norec )
+
+apply   simp
+apply   (case_tac "table_of (map (\<lambda>(s, m). (s, Ca, m)) (methods c)) sig")
+apply     (force intro: rtrancl_into_rtrancl2)
+
+apply     (auto intro: methd_norec)
+done
+
+lemma methd_inheritedD:
+  "\<lbrakk>class G C = Some c; ws_prog G;methd G C sig = Some m\<rbrakk>
+  \<Longrightarrow>  (declclass m \<noteq> C \<longrightarrow> G \<turnstile>C inherits method sig m)"
+by (auto simp add: methd_rec)
+
+lemma methd_diff_cls:
+"\<lbrakk>ws_prog G; is_class G C; is_class G D;
+ methd G C sig = m; methd G D sig = n; m\<noteq>n
+\<rbrakk> \<Longrightarrow> C\<noteq>D"
+by (auto simp add: methd_rec)
+
+lemma method_declared_inI: 
+ "\<lbrakk>table_of (methods c) sig = Some m; class G C = Some c\<rbrakk>
+  \<Longrightarrow> G\<turnstile>mdecl (sig,m) declared_in C"
+by (auto simp add: cdeclaredmethd_def declared_in_def)
+
+lemma methd_declared_in_declclass: 
+ "\<lbrakk>methd G C sig = Some m; ws_prog G;is_class G C\<rbrakk> 
+ \<Longrightarrow> G\<turnstile>Methd sig m declared_in (declclass m)"
+by (auto dest: methd_declC method_declared_inI)
+
+lemma member_methd:
+ (assumes member_of: "G\<turnstile>Methd sig m member_of C" and
+                 ws: "ws_prog G"
+ ) "methd G C sig = Some m"
+proof -
+  from member_of 
+  have iscls_C: "is_class G C" 
+    by (rule member_of_is_classD)
+  from iscls_C ws member_of
+  show ?thesis (is "?Methd C")
+  proof (induct rule: ws_class_induct')
+    case (Object co)
+    assume "G \<turnstile>Methd sig m member_of Object"
+    then have "G\<turnstile>Methd sig m declared_in Object \<and> declclass m = Object"
+      by (cases set: members) (cases m, auto dest: subcls1D)
+    with ws Object 
+    show "?Methd Object"
+      by (cases m)
+         (auto simp add: declared_in_def cdeclaredmethd_def methd_rec
+                  intro:  table_of_mapconst_SomeI)
+  next
+    case (Subcls C c)
+    assume clsC: "class G C = Some c" and
+      neq_C_Obj: "C \<noteq> Object" and
+            hyp: "G \<turnstile>Methd sig m member_of super c \<Longrightarrow> ?Methd (super c)" and
+      member_of: "G \<turnstile>Methd sig m member_of C"
+    from member_of
+    show "?Methd C"
+    proof (cases)
+      case (Immediate Ca membr)
+      then have "Ca=C" "membr = method sig m" and 
+                "G\<turnstile>Methd sig m declared_in C" "declclass m = C"
+	by (cases m,auto)
+      with clsC 
+      have "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = Some m"
+	by (cases m)
+	   (auto simp add: declared_in_def cdeclaredmethd_def
+	            intro: table_of_mapconst_SomeI)
+      with clsC neq_C_Obj ws 
+      show ?thesis
+	by (simp add: methd_rec)
+    next
+      case (Inherited Ca S membr)
+      with clsC
+      have eq_Ca_C: "Ca=C" and
+            undecl: "G\<turnstile>mid sig undeclared_in C" and
+             super: "G \<turnstile>Methd sig m member_of (super c)"
+	by (auto dest: subcls1D)
+      from eq_Ca_C clsC undecl 
+      have "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = None"
+	by (auto simp add: undeclared_in_def cdeclaredmethd_def
+	            intro: table_of_mapconst_NoneI)
+      moreover
+      from Inherited have "G \<turnstile> C inherits (method sig m)"
+	by (auto simp add: inherits_def)
+      moreover
+      note clsC neq_C_Obj ws super hyp 
+      ultimately
+      show ?thesis
+	by (auto simp add: methd_rec intro: filter_tab_SomeI)
+    qed
+  qed
+qed
+
+(*unused*)
+lemma finite_methd:"ws_prog G \<Longrightarrow> finite {methd G C sig |sig C. is_class G C}"
+apply (rule finite_is_class [THEN finite_SetCompr2])
+apply (intro strip)
+apply (erule_tac ws_subcls1_induct, assumption)
+apply (subst methd_rec)
+apply (assumption)
+apply (auto intro!: finite_range_map_of finite_range_filter_tab finite_range_map_of_override)
+done
+
+lemma finite_dom_methd:
+ "\<lbrakk>ws_prog G; is_class G C\<rbrakk> \<Longrightarrow> finite (dom (methd G C))"
+apply (erule_tac ws_subcls1_induct)
+apply assumption
+apply (subst methd_rec)
+apply (assumption)
+apply (auto intro!: finite_dom_map_of finite_dom_filter_tab)
+done
+
+
+subsection "accmethd"
+
+lemma accmethd_SomeD:
+"accmethd G S C sig = Some m
+ \<Longrightarrow> methd G C sig = Some m \<and> G\<turnstile>method sig m of C accessible_from S"
+by (auto simp add: accmethd_def dest: filter_tab_SomeD)
+
+lemma accmethd_SomeI:
+"\<lbrakk>methd G C sig = Some m; G\<turnstile>method sig m of C accessible_from S\<rbrakk> 
+ \<Longrightarrow> accmethd G S C sig = Some m"
+by (auto simp add: accmethd_def intro: filter_tab_SomeI)
+
+lemma accmethd_declC: 
+"\<lbrakk>accmethd G S C sig = Some m; ws_prog G; is_class G C\<rbrakk> \<Longrightarrow>
+ (\<exists>d. class G (declclass m)=Some d \<and> 
+  table_of (methods d) sig=Some (mthd m)) \<and> 
+ G\<turnstile>C \<preceq>\<^sub>C (declclass m) \<and> methd G (declclass m) sig = Some m \<and> 
+ G\<turnstile>method sig m of C accessible_from S"
+by (auto dest: accmethd_SomeD methd_declC accmethd_SomeI)
+
+
+lemma finite_dom_accmethd:
+ "\<lbrakk>ws_prog G; is_class G C\<rbrakk> \<Longrightarrow> finite (dom (accmethd G S C))"
+by (auto simp add: accmethd_def intro: finite_dom_filter_tab finite_dom_methd)
+
+
+subsection "dynmethd"
+
+lemma dynmethd_rec:
+"\<lbrakk>class G dynC = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
+ dynmethd G statC dynC sig
+   = (if G\<turnstile>dynC \<preceq>\<^sub>C statC
+        then (case methd G statC sig of
+                None \<Rightarrow> None
+              | Some statM 
+                  \<Rightarrow> (case methd G dynC sig of
+                        None \<Rightarrow> dynmethd G statC (super c) sig
+                      | Some dynM \<Rightarrow> 
+                          (if G,sig\<turnstile> dynM overrides statM \<or> dynM = statM 
+                              then Some dynM
+                              else (dynmethd G statC (super c) sig)
+                      )))
+         else None)" 
+   (is "_ \<Longrightarrow> _ \<Longrightarrow> ?Dynmethd_def dynC sig  = ?Dynmethd_rec dynC c sig") 
+proof -
+  assume clsDynC: "class G dynC = Some c" and 
+              ws: "ws_prog G"
+  then show "?Dynmethd_def dynC sig  = ?Dynmethd_rec dynC c sig" 
+  proof (induct rule: ws_class_induct'')
+    case (Object co)
+    show "?Dynmethd_def Object sig = ?Dynmethd_rec Object co sig"
+    proof (cases "G\<turnstile>Object \<preceq>\<^sub>C statC") 
+      case False
+      then show ?thesis by (simp add: dynmethd_def)
+    next
+      case True
+      then have eq_statC_Obj: "statC = Object" ..
+      show ?thesis 
+      proof (cases "methd G statC sig")
+	case None then show ?thesis by (simp add: dynmethd_def)
+      next
+	case Some
+	with True Object ws eq_statC_Obj 
+	show ?thesis
+	  by (auto simp add: dynmethd_def class_rec
+                      intro: filter_tab_SomeI)
+      qed
+    qed
+  next  
+    case (Subcls dynC c sc)
+    show "?Dynmethd_def dynC sig = ?Dynmethd_rec dynC c sig"
+    proof (cases "G\<turnstile>dynC \<preceq>\<^sub>C statC") 
+      case False
+      then show ?thesis by (simp add: dynmethd_def)
+    next
+      case True
+      note subclseq_dynC_statC = True
+      show ?thesis
+      proof (cases "methd G statC sig")
+	case None then show ?thesis by (simp add: dynmethd_def)
+      next
+	case (Some statM)
+	note statM = Some
+	let "?filter C" = 
+              "filter_tab
+                (\<lambda>_ dynM. G,sig \<turnstile> dynM overrides statM \<or> dynM = statM)
+                (methd G C)"
+        let "?class_rec C" =
+              "(class_rec (G, C) empty
+                           (\<lambda>C c subcls_mthds. subcls_mthds ++ (?filter C)))"
+	from statM Subcls ws subclseq_dynC_statC
+	have dynmethd_dynC_def:
+             "?Dynmethd_def dynC sig =
+                ((?class_rec (super c)) 
+                 ++
+                (?filter dynC)) sig"
+         by (simp (no_asm_simp) only: dynmethd_def class_rec)
+	    auto
+       show ?thesis
+       proof (cases "dynC = statC")
+	 case True
+	 with subclseq_dynC_statC statM dynmethd_dynC_def
+	 have "?Dynmethd_def dynC sig = Some statM"
+	   by (auto intro: override_find_right filter_tab_SomeI)
+	 with subclseq_dynC_statC True Some 
+	 show ?thesis
+	   by auto
+       next
+	 case False
+	 with subclseq_dynC_statC Subcls 
+	 have subclseq_super_statC: "G\<turnstile>(super c) \<preceq>\<^sub>C statC"
+	   by (blast dest: subclseq_superD)
+	 show ?thesis
+	 proof (cases "methd G dynC sig") 
+	   case None
+	   then have "?filter dynC sig = None"
+	     by (rule filter_tab_None)
+           then have "?Dynmethd_def dynC sig=?class_rec (super c) sig"
+	     by (simp add: dynmethd_dynC_def)
+	   with  subclseq_super_statC statM None
+	   have "?Dynmethd_def dynC sig = ?Dynmethd_def (super c) sig"
+	     by (auto simp add: empty_def dynmethd_def)
+           with None subclseq_dynC_statC statM
+	   show ?thesis 
+	     by simp
+	 next
+	   case (Some dynM)
+	   note dynM = Some
+	   let ?Termination = "G \<turnstile> qmdecl sig dynM overrides qmdecl sig statM \<or>
+                               dynM = statM"
+	   show ?thesis
+	   proof (cases "?filter dynC sig")
+	     case None
+	     with dynM 
+	     have no_termination: "\<not> ?Termination"
+	       by (simp add: filter_tab_def)
+	     from None 
+	     have "?Dynmethd_def dynC sig=?class_rec (super c) sig"
+	       by (simp add: dynmethd_dynC_def)
+	     with subclseq_super_statC statM dynM no_termination
+	     show ?thesis
+	       by (auto simp add: empty_def dynmethd_def)
+	   next
+	     case Some
+	     with dynM
+	     have termination: "?Termination"
+	       by (auto)
+	     with Some dynM
+	     have "?Dynmethd_def dynC sig=Some dynM"
+	      by (auto simp add: dynmethd_dynC_def)
+	     with subclseq_super_statC statM dynM termination
+	     show ?thesis
+	       by (auto simp add: dynmethd_def)
+	   qed
+	 qed
+       qed
+     qed
+   qed
+ qed
+qed
+	       
+lemma dynmethd_C_C:"\<lbrakk>is_class G C; ws_prog G\<rbrakk> 
+\<Longrightarrow> dynmethd G C C sig = methd G C sig"          
+apply (auto simp add: dynmethd_rec)
+done
+ 
+lemma dynmethdSomeD: 
+ "\<lbrakk>dynmethd G statC dynC sig = Some dynM; is_class G dynC; ws_prog G\<rbrakk> 
+  \<Longrightarrow> G\<turnstile>dynC \<preceq>\<^sub>C statC \<and> (\<exists> statM. methd G statC sig = Some statM)"
+apply clarify
+apply rotate_tac
+by (auto simp add: dynmethd_rec)
+ 
+lemma dynmethd_Some_cases [consumes 3, case_names Static Overrides]:
+  (assumes dynM: "dynmethd G statC dynC sig = Some dynM" and
+        is_cls_dynC: "is_class G dynC" and
+                 ws: "ws_prog G" and
+         hyp_static: "methd G statC sig = Some dynM \<Longrightarrow> P" and
+       hyp_override: "\<And> statM. \<lbrakk>methd G statC sig = Some statM;dynM\<noteq>statM;
+                       G,sig\<turnstile>dynM overrides statM\<rbrakk> \<Longrightarrow> P"
+   ) "P"
+proof -
+  from is_cls_dynC obtain dc where clsDynC: "class G dynC = Some dc" by blast
+  from clsDynC ws dynM hyp_static hyp_override
+  show "P"
+  proof (induct rule: ws_class_induct)
+    case (Object co)
+    with ws  have "statC = Object" 
+      by (auto simp add: dynmethd_rec)
+    with ws Object show ?thesis by (auto simp add: dynmethd_C_C)
+  next
+    case (Subcls C c)
+    with ws show ?thesis
+      by (auto simp add: dynmethd_rec)
+  qed
+qed
+
+lemma no_override_in_Object:
+  (assumes     dynM: "dynmethd G statC dynC sig = Some dynM" and
+            is_cls_dynC: "is_class G dynC" and
+                     ws: "ws_prog G" and
+                  statM: "methd G statC sig = Some statM" and
+         neq_dynM_statM: "dynM\<noteq>statM"
+   )
+   "dynC \<noteq> Object"
+proof -
+  from is_cls_dynC obtain dc where clsDynC: "class G dynC = Some dc" by blast
+  from clsDynC ws dynM statM neq_dynM_statM 
+  show ?thesis (is "?P dynC")
+  proof (induct rule: ws_class_induct)
+    case (Object co)
+    with ws  have "statC = Object" 
+      by (auto simp add: dynmethd_rec)
+    with ws Object show "?P Object" by (auto simp add: dynmethd_C_C)
+  next
+    case (Subcls dynC c)
+    with ws show "?P dynC"
+      by (auto simp add: dynmethd_rec)
+  qed
+qed
+
+
+lemma dynmethd_Some_rec_cases [consumes 3, 
+                               case_names Static Override  Recursion]:
+(assumes     dynM: "dynmethd G statC dynC sig = Some dynM" and
+                clsDynC: "class G dynC = Some c" and
+                     ws: "ws_prog G" and
+             hyp_static: "methd G statC sig = Some dynM \<Longrightarrow> P" and
+           hyp_override: "\<And> statM. \<lbrakk>methd G statC sig = Some statM;
+                                     methd G dynC sig = Some dynM; statM\<noteq>dynM;
+                                     G,sig\<turnstile> dynM overrides statM\<rbrakk> \<Longrightarrow> P" and
+
+          hyp_recursion: "\<lbrakk>dynC\<noteq>Object;
+                           dynmethd G statC (super c) sig = Some dynM\<rbrakk> \<Longrightarrow> P" 
+  ) "P"
+proof -
+  from clsDynC have "is_class G dynC" by simp
+  note no_override_in_Object' = no_override_in_Object [OF dynM this ws]
+  from ws clsDynC dynM hyp_static hyp_override hyp_recursion
+  show ?thesis
+    by (auto simp add: dynmethd_rec dest: no_override_in_Object')
+qed
+
+lemma dynmethd_declC: 
+"\<lbrakk>dynmethd G statC dynC sig = Some m;
+  is_class G statC;ws_prog G
+ \<rbrakk> \<Longrightarrow>
+  (\<exists>d. class G (declclass m)=Some d \<and> table_of (methods d) sig=Some (mthd m)) \<and>
+  G\<turnstile>dynC \<preceq>\<^sub>C (declclass m) \<and> methd G (declclass m) sig = Some m"
+proof - 
+  assume  is_cls_statC: "is_class G statC" 
+  assume            ws: "ws_prog G"  
+  assume             m: "dynmethd G statC dynC sig = Some m"
+  from m 
+  have "G\<turnstile>dynC \<preceq>\<^sub>C statC" by (auto simp add: dynmethd_def)
+  from this is_cls_statC 
+  have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2)
+  from is_cls_dynC ws m  
+  show ?thesis (is "?P dynC") 
+  proof (induct rule: ws_class_induct')
+    case (Object co)
+    with ws have "statC=Object" by (auto simp add: dynmethd_rec)
+    with ws Object  
+    show "?P Object" 
+      by (auto simp add: dynmethd_C_C dest: methd_declC)
+  next
+    case (Subcls dynC c)
+    assume   hyp: "dynmethd G statC (super c) sig = Some m \<Longrightarrow> ?P (super c)" and
+         clsDynC: "class G dynC = Some c"  and
+              m': "dynmethd G statC dynC sig = Some m" and
+    neq_dynC_Obj: "dynC \<noteq> Object"
+    from ws this obtain statM where
+      subclseq_dynC_statC: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and 
+                     statM: "methd G statC sig = Some statM"
+      by (blast dest: dynmethdSomeD)
+    from clsDynC neq_dynC_Obj 
+    have subclseq_dynC_super: "G\<turnstile>dynC \<preceq>\<^sub>C (super c)" 
+      by (auto intro: subcls1I) 
+    from m' clsDynC ws 
+    show "?P dynC"
+    proof (cases rule: dynmethd_Some_rec_cases) 
+      case Static
+      with is_cls_statC ws subclseq_dynC_statC 
+      show ?thesis 
+	by (auto intro: rtrancl_trans dest: methd_declC)
+    next
+      case Override
+      with clsDynC ws 
+      show ?thesis 
+	by (auto dest: methd_declC)
+    next
+      case Recursion
+      with hyp subclseq_dynC_super 
+      show ?thesis 
+	by (auto intro: rtrancl_trans) 
+    qed
+  qed
+qed
+
+lemma methd_Some_dynmethd_Some:
+  (assumes    statM: "methd G statC sig = Some statM" and 
+           subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
+       is_cls_statC: "is_class G statC" and
+                 ws: "ws_prog G"
+   ) "\<exists> dynM. dynmethd G statC dynC sig = Some dynM"
+   (is "?P dynC")
+proof -
+  from subclseq is_cls_statC 
+  have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2)
+  then obtain dc where
+    clsDynC: "class G dynC = Some dc" by blast
+  from clsDynC ws subclseq 
+  show "?thesis"
+  proof (induct rule: ws_class_induct)
+    case (Object co)
+    with ws  have "statC = Object" 
+      by (auto)
+    with ws Object statM
+    show "?P Object"  
+      by (auto simp add: dynmethd_C_C)
+  next
+    case (Subcls dynC dc)
+    assume clsDynC': "class G dynC = Some dc"
+    assume neq_dynC_Obj: "dynC \<noteq> Object"
+    assume hyp: "G\<turnstile>super dc\<preceq>\<^sub>C statC \<Longrightarrow> ?P (super dc)"
+    assume subclseq': "G\<turnstile>dynC\<preceq>\<^sub>C statC"
+    then
+    show "?P dynC"
+    proof (cases rule: subclseq_cases)
+      case Eq
+      with ws statM clsDynC' 
+      show ?thesis
+	by (auto simp add: dynmethd_rec)
+    next
+      case Subcls
+      assume "G\<turnstile>dynC\<prec>\<^sub>C statC"
+      from this clsDynC' 
+      have "G\<turnstile>super dc\<preceq>\<^sub>C statC" by (rule subcls_superD)
+      with hyp ws clsDynC' subclseq' statM
+      show ?thesis
+	by (auto simp add: dynmethd_rec)
+    qed
+  qed
+qed
+
+lemma dynmethd_cases [consumes 4, case_names Static Overrides]:
+  (assumes    statM: "methd G statC sig = Some statM" and 
+           subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
+       is_cls_statC: "is_class G statC" and
+                 ws: "ws_prog G" and
+         hyp_static: "dynmethd G statC dynC sig = Some statM \<Longrightarrow> P" and
+       hyp_override: "\<And> dynM. \<lbrakk>dynmethd G statC dynC sig = Some dynM;
+                                 dynM\<noteq>statM;
+                           G,sig\<turnstile>dynM overrides statM\<rbrakk> \<Longrightarrow> P"
+   ) "P"
+proof -
+  from subclseq is_cls_statC 
+  have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2)
+  then obtain dc where
+    clsDynC: "class G dynC = Some dc" by blast
+  from statM subclseq is_cls_statC ws 
+  obtain dynM
+    where dynM: "dynmethd G statC dynC sig = Some dynM"
+    by (blast dest: methd_Some_dynmethd_Some)
+  from dynM is_cls_dynC ws 
+  show ?thesis
+  proof (cases rule: dynmethd_Some_cases)
+    case Static
+    with hyp_static dynM statM show ?thesis by simp
+  next
+    case Overrides
+    with hyp_override dynM statM show ?thesis by simp
+  qed
+qed
+
+lemma ws_dynmethd:
+  (assumes statM: "methd G statC sig = Some statM" and
+        subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
+    is_cls_statC: "is_class G statC" and
+              ws: "ws_prog G"
+   )
+   "\<exists> dynM. dynmethd G statC dynC sig = Some dynM \<and>
+            is_static dynM = is_static statM \<and> G\<turnstile>resTy dynM\<preceq>resTy statM"
+proof - 
+  from statM subclseq is_cls_statC ws
+  show ?thesis
+  proof (cases rule: dynmethd_cases)
+    case Static
+    with statM 
+    show ?thesis
+      by simp
+  next
+    case Overrides
+    with ws
+    show ?thesis
+      by (auto dest: ws_overrides_commonD)
+  qed
+qed
+
+(*
+lemma dom_dynmethd: 
+  "dom (dynmethd G statC dynC) \<subseteq> dom (methd G statC) \<union> dom (methd G dynC)"
+by (auto simp add: dynmethd_def dom_def)
+
+lemma finite_dom_dynmethd:
+ "\<lbrakk>ws_prog G; is_class G statC; is_class G dynC\<rbrakk> 
+  \<Longrightarrow> finite (dom (dynmethd G statC dynC))"
+apply (rule_tac B="dom (methd G statC) \<union> dom (methd G dynC)" in finite_subset)
+apply (rule dom_dynmethd)
+apply (rule finite_UnI)
+apply (drule (2) finite_dom_methd)+
+done
+*)
+(*
+lemma dynmethd_SomeD: 
+"\<lbrakk>ws_prog G; is_class G statC; is_class G dynC;
+ methd G statC sig = Some sm; dynmethd G statC dynC sig = Some dm; sm \<noteq> dm
+ \<rbrakk> \<Longrightarrow> G\<turnstile>dynC \<prec>\<^sub>C statC \<and> 
+       (declclass dm \<noteq> dynC \<longrightarrow> G \<turnstile> dm accessible_through_inheritance_in dynC)"
+by (auto simp add: dynmethd_def 
+         dest: methd_inheritedD methd_diff_cls
+         intro: rtrancl_into_trancl3)
+*)
+
+subsection "dynlookup"
+
+lemma dynlookup_cases [consumes 1, case_names NullT IfaceT ClassT ArrayT]:
+"\<lbrakk>dynlookup G statT dynC sig = x;
+           \<lbrakk>statT = NullT       ; empty sig = x                  \<rbrakk> \<Longrightarrow> P;
+  \<And> I.    \<lbrakk>statT = IfaceT I    ; dynimethd G I      dynC sig = x\<rbrakk> \<Longrightarrow> P;
+  \<And> statC.\<lbrakk>statT = ClassT statC; dynmethd  G statC  dynC sig = x\<rbrakk> \<Longrightarrow> P;
+  \<And> ty.   \<lbrakk>statT = ArrayT ty   ; dynmethd  G Object dynC sig = x\<rbrakk> \<Longrightarrow> P
+ \<rbrakk> \<Longrightarrow> P"
+by (cases statT) (auto simp add: dynlookup_def)
+
+subsection "fields"
+
+lemma fields_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
+  fields G C = map (\<lambda>(fn,ft). ((fn,C),ft)) (cfields c) @  
+  (if C = Object then [] else fields G (super c))"
+apply (simp only: fields_def)
+apply (erule class_rec [THEN trans])
+apply assumption
+apply clarsimp
+done
+
+(* local lemma *)
+lemma fields_norec: 
+"\<lbrakk>class G fd = Some c; ws_prog G;  table_of (cfields c) fn = Some f\<rbrakk> 
+ \<Longrightarrow> table_of (fields G fd) (fn,fd) = Some f"
+apply (subst fields_rec)
+apply assumption+
+apply (subst map_of_override [symmetric])
+apply (rule disjI1 [THEN override_Some_iff [THEN iffD2]])
+apply (auto elim: table_of_map2_SomeI)
+done
+
+(* local lemma *)
+lemma table_of_fieldsD:
+"table_of (map (\<lambda>(fn,ft). ((fn,C),ft)) (cfields c)) efn = Some f
+\<Longrightarrow> (declclassf efn) = C \<and> table_of (cfields c) (fname efn) = Some f"
+apply (case_tac "efn")
+by auto
+
+lemma fields_declC: 
+ "\<lbrakk>table_of (fields G C) efn = Some f; ws_prog G; is_class G C\<rbrakk> \<Longrightarrow>  
+  (\<exists>d. class G (declclassf efn) = Some d \<and> 
+                    table_of (cfields d) (fname efn)=Some f) \<and> 
+  G\<turnstile>C \<preceq>\<^sub>C (declclassf efn)  \<and> table_of (fields G (declclassf efn)) efn = Some f"
+apply (erule make_imp)
+apply (rule ws_subcls1_induct, assumption, assumption)
+apply (subst fields_rec, assumption)
+apply clarify
+apply (simp only: map_of_override [symmetric] del: map_of_override)
+apply (case_tac "table_of (map (split (\<lambda>fn. Pair (fn, Ca))) (cfields c)) efn") 
+apply   (force intro:rtrancl_into_rtrancl2 simp add: override_def)
+
+apply   (frule_tac fd="Ca" in fields_norec)
+apply     assumption
+apply     blast
+apply   (frule table_of_fieldsD)  
+apply   (frule_tac n="table_of (map (split (\<lambda>fn. Pair (fn, Ca))) (cfields c))"
+              and  m="table_of (if Ca = Object then [] else fields G (super c))"
+         in override_find_right)
+apply   (case_tac "efn")
+apply   (simp)
+done
+
+lemma fields_emptyI: "\<And>y. \<lbrakk>ws_prog G; class G C = Some c;cfields c = [];  
+  C \<noteq> Object \<longrightarrow> class G (super c) = Some y \<and> fields G (super c) = []\<rbrakk> \<Longrightarrow>  
+  fields G C = []"
+apply (subst fields_rec)
+apply assumption
+apply auto
+done
+
+(* easier than with table_of *)
+lemma fields_mono_lemma: 
+"\<lbrakk>x \<in> set (fields G C); G\<turnstile>D \<preceq>\<^sub>C C; ws_prog G\<rbrakk> 
+ \<Longrightarrow> x \<in> set (fields G D)"
+apply (erule make_imp)
+apply (erule converse_rtrancl_induct)
+apply  fast
+apply (drule subcls1D)
+apply clarsimp
+apply (subst fields_rec)
+apply   auto
+done
+
+
+lemma ws_unique_fields_lemma: 
+ "\<lbrakk>(efn,fd)  \<in> set (fields G (super c)); fc \<in> set (cfields c); ws_prog G;  
+   fname efn = fname fc; declclassf efn = C;
+   class G C = Some c; C \<noteq> Object; class G (super c) = Some d\<rbrakk> \<Longrightarrow> R"
+apply (frule_tac ws_prog_cdeclD [THEN conjunct2], assumption, assumption)
+apply (drule_tac weak_map_of_SomeI)
+apply (frule_tac subcls1I [THEN subcls1_irrefl], assumption, assumption)
+apply (auto dest: fields_declC [THEN conjunct2 [THEN conjunct1[THEN rtranclD]]])
+done
+
+lemma ws_unique_fields: "\<lbrakk>is_class G C; ws_prog G; 
+       \<And>C c. \<lbrakk>class G C = Some c\<rbrakk> \<Longrightarrow> unique (cfields c) \<rbrakk> \<Longrightarrow>
+      unique (fields G C)" 
+apply (rule ws_subcls1_induct, assumption, assumption)
+apply (subst fields_rec, assumption)            
+apply (auto intro!: unique_map_inj injI 
+            elim!: unique_append ws_unique_fields_lemma fields_norec
+            )
+done
+
+
+subsection "accfield"
+
+lemma accfield_fields: 
+ "accfield G S C fn = Some f 
+  \<Longrightarrow> table_of (fields G C) (fn, declclass f) = Some (fld f)"
+apply (simp only: accfield_def Let_def)
+apply (rule table_of_remap_SomeD)
+apply (auto dest: filter_tab_SomeD)
+done
+
+
+lemma accfield_declC_is_class: 
+ "\<lbrakk>is_class G C; accfield G S C en = Some (fd, f); ws_prog G\<rbrakk> \<Longrightarrow>  
+   is_class G fd"
+apply (drule accfield_fields)
+apply (drule fields_declC [THEN conjunct1], assumption)
+apply auto
+done
+
+lemma accfield_accessibleD: 
+  "accfield G S C fn = Some f \<Longrightarrow> G\<turnstile>Field fn f of C accessible_from S"
+by (auto simp add: accfield_def Let_def)
+
+subsection "is methd"
+
+lemma is_methdI: 
+"\<lbrakk>class G C = Some y; methd G C sig = Some b\<rbrakk> \<Longrightarrow> is_methd G C sig"
+apply (unfold is_methd_def)
+apply auto
+done
+
+lemma is_methdD: 
+"is_methd G C sig \<Longrightarrow> class G C \<noteq> None \<and> methd G C sig \<noteq> None"
+apply (unfold is_methd_def)
+apply auto
+done
+
+lemma finite_is_methd: 
+ "ws_prog G \<Longrightarrow> finite (Collect (split (is_methd G)))"
+apply (unfold is_methd_def)
+apply (subst SetCompr_Sigma_eq)
+apply (rule finite_is_class [THEN finite_SigmaI])
+apply (simp only: mem_Collect_eq)
+apply (fold dom_def)
+apply (erule finite_dom_methd)
+apply assumption
+done
+
+section "calculation of the superclasses of a class"
+
+constdefs 
+ superclasses:: "prog \<Rightarrow> qtname \<Rightarrow> qtname set"
+ "superclasses G C \<equiv> class_rec (G,C) {} 
+                       (\<lambda> C c superclss. (if C=Object 
+                                            then {} 
+                                            else insert (super c) superclss))"
+   
+lemma superclasses_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
+ superclasses G C 
+ = (if (C=Object) 
+       then {}
+       else insert (super c) (superclasses G (super c)))"
+apply (unfold superclasses_def)
+apply (erule class_rec [THEN trans], assumption)
+apply (simp)
+done
+
+lemma superclasses_mono:
+"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D;ws_prog G; class G C = Some c;
+  \<And> C c. \<lbrakk>class G C = Some c;C\<noteq>Object\<rbrakk> \<Longrightarrow> \<exists> sc. class G (super c) = Some sc;
+  x\<in>superclasses G D 
+\<rbrakk> \<Longrightarrow> x\<in>superclasses G C" 
+proof -
+  
+  assume     ws: "ws_prog G"          and 
+          cls_C: "class G C = Some c" and
+             wf: "\<And>C c. \<lbrakk>class G C = Some c; C \<noteq> Object\<rbrakk>
+                  \<Longrightarrow> \<exists>sc. class G (super c) = Some sc"
+  assume clsrel: "G\<turnstile>C\<prec>\<^sub>C D"           
+  thus "\<And> c. \<lbrakk>class G C = Some c; x\<in>superclasses G D\<rbrakk>\<Longrightarrow>
+        x\<in>superclasses G C" (is "PROP ?P C"  
+                             is "\<And> c. ?CLS C c \<Longrightarrow> ?SUP D \<Longrightarrow> ?SUP C")
+  proof (induct ?P C  rule: converse_trancl_induct)
+    fix C c
+    assume "G\<turnstile>C\<prec>\<^sub>C\<^sub>1D" "class G C = Some c" "x \<in> superclasses G D"
+    with wf ws show "?SUP C"
+      by (auto    intro: no_subcls1_Object 
+               simp add: superclasses_rec subcls1_def)
+  next
+    fix C S c
+    assume clsrel': "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" "G\<turnstile>S \<prec>\<^sub>C D"
+       and    hyp : "\<And> s. \<lbrakk>class G S = Some s; x \<in> superclasses G D\<rbrakk>
+                           \<Longrightarrow> x \<in> superclasses G S"
+       and  cls_C': "class G C = Some c"
+       and       x: "x \<in> superclasses G D"
+    moreover note wf ws
+    moreover from calculation 
+    have "?SUP S" 
+      by (force intro: no_subcls1_Object simp add: subcls1_def)
+    moreover from calculation 
+    have "super c = S" 
+      by (auto intro: no_subcls1_Object simp add: subcls1_def)
+    ultimately show "?SUP C" 
+      by (auto intro: no_subcls1_Object simp add: superclasses_rec)
+  qed
+qed
+
+lemma subclsEval:
+"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D;ws_prog G; class G C = Some c;
+  \<And> C c. \<lbrakk>class G C = Some c;C\<noteq>Object\<rbrakk> \<Longrightarrow> \<exists> sc. class G (super c) = Some sc 
+ \<rbrakk> \<Longrightarrow> D\<in>superclasses G C" 
+proof -
+  note converse_trancl_induct 
+       = converse_trancl_induct [consumes 1,case_names Single Step]
+  assume 
+             ws: "ws_prog G"          and 
+          cls_C: "class G C = Some c" and
+             wf: "\<And>C c. \<lbrakk>class G C = Some c; C \<noteq> Object\<rbrakk>
+                  \<Longrightarrow> \<exists>sc. class G (super c) = Some sc"
+  assume clsrel: "G\<turnstile>C\<prec>\<^sub>C D"           
+  thus "\<And> c. class G C = Some c\<Longrightarrow> D\<in>superclasses G C" 
+    (is "PROP ?P C"  is "\<And> c. ?CLS C c  \<Longrightarrow> ?SUP C")
+  proof (induct ?P C  rule: converse_trancl_induct)
+    fix C c
+    assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" "class G C = Some c"
+    with ws wf show "?SUP C"
+      by (auto intro: no_subcls1_Object simp add: superclasses_rec subcls1_def)
+  next
+    fix C S c
+    assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" "G\<turnstile>S\<prec>\<^sub>C D" 
+           "\<And>s. class G S = Some s \<Longrightarrow> D \<in> superclasses G S"
+           "class G C = Some c" 
+    with ws wf show "?SUP C"
+      by - (rule superclasses_mono,
+            auto dest: no_subcls1_Object simp add: subcls1_def )
+  qed
+qed
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Eval.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,993 @@
+(*  Title:      isabelle/Bali/Eval.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* Operational evaluation (big-step) semantics of Java expressions and 
+          statements
+*}
+
+theory Eval = State + DeclConcepts:
+
+text {*
+
+improvements over Java Specification 1.0:
+\begin{itemize}
+\item dynamic method lookup does not need to consider the return type 
+      (cf.15.11.4.4)
+\item throw raises a NullPointer exception if a null reference is given, and 
+      each throw of a standard exception yield a fresh exception object 
+      (was not specified)
+\item if there is not enough memory even to allocate an OutOfMemory exception,
+  evaluation/execution fails, i.e. simply stops (was not specified)
+\item array assignment checks lhs (and may throw exceptions) before evaluating 
+      rhs
+\item fixed exact positions of class initializations 
+      (immediate at first active use)
+\end{itemize}
+
+design issues:
+\begin{itemize}
+\item evaluation vs. (single-step) transition semantics
+  evaluation semantics chosen, because:
+  \begin{itemize} 
+  \item[++] less verbose and therefore easier to read (and to handle in proofs)
+  \item[+]  more abstract
+  \item[+]  intermediate values (appearing in recursive rules) need not be 
+     stored explicitly, e.g. no call body construct or stack of invocation 
+     frames containing local variables and return addresses for method calls 
+     needed
+  \item[+]  convenient rule induction for subject reduction theorem
+  \item[-]  no interleaving (for parallelism) can be described
+  \item[-]  stating a property of infinite executions requires the meta-level 
+     argument that this property holds for any finite prefixes of it 
+     (e.g. stopped using a counter that is decremented to zero and then 
+     throwing an exception)
+  \end{itemize}
+\item unified evaluation for variables, expressions, expression lists, 
+      statements
+\item the value entry in statement rules is redundant 
+\item the value entry in rules is irrelevant in case of exceptions, but its full
+  inclusion helps to make the rule structure independent of exception occurence.
+\item as irrelevant value entries are ignored, it does not matter if they are 
+      unique.
+  For simplicity, (fixed) arbitrary values are preferred over "free" values.
+\item the rule format is such that the start state may contain an exception.
+  \begin{itemize}
+  \item[++] faciliates exception handling
+  \item[+]  symmetry
+  \end{itemize}
+\item the rules are defined carefully in order to be applicable even in not
+  type-correct situations (yielding undefined values),
+  e.g. @{text "the_Addr (Val (Bool b)) = arbitrary"}.
+  \begin{itemize}
+  \item[++] fewer rules 
+  \item[-]  less readable because of auxiliary functions like @{text the_Addr}
+  \end{itemize}
+  Alternative: "defensive" evaluation throwing some InternalError exception
+               in case of (impossible, for correct programs) type mismatches
+\item there is exactly one rule per syntactic construct
+  \begin{itemize}
+  \item[+] no redundancy in case distinctions
+  \end{itemize}
+\item halloc fails iff there is no free heap address. When there is
+  only one free heap address left, it returns an OutOfMemory exception.
+  In this way it is guaranteed that when an OutOfMemory exception is thrown for
+  the first time, there is a free location on the heap to allocate it.
+\item the allocation of objects that represent standard exceptions is deferred 
+      until execution of any enclosing catch clause, which is transparent to 
+      the program.
+  \begin{itemize}
+  \item[-]  requires an auxiliary execution relation
+  \item[++] avoids copies of allocation code and awkward case distinctions 
+           (whether there is enough memory to allocate the exception) in 
+            evaluation rules
+  \end{itemize}
+\item unfortunately @{text new_Addr} is not directly executable because of 
+      Hilbert operator.
+\end{itemize}
+simplifications:
+\begin{itemize}
+\item local variables are initialized with default values 
+      (no definite assignment)
+\item garbage collection not considered, therefore also no finalizers
+\item stack overflow and memory overflow during class initialization not 
+      modelled
+\item exceptions in initializations not replaced by ExceptionInInitializerError
+\end{itemize}
+*}
+
+types vvar  =         "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
+      vals  =        "(val, vvar, val list) sum3"
+translations
+     "vvar" <= (type) "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
+     "vals" <= (type)"(val, vvar, val list) sum3"
+
+syntax (xsymbols)
+  dummy_res :: "vals" ("\<diamondsuit>")
+translations
+  "\<diamondsuit>" == "In1 Unit"
+
+constdefs
+  arbitrary3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals"
+ "arbitrary3 \<equiv> sum3_case (In1 \<circ> sum_case (\<lambda>x. arbitrary) (\<lambda>x. Unit))
+                     (\<lambda>x. In2 arbitrary) (\<lambda>x. In3 arbitrary)"
+
+lemma [simp]: "arbitrary3 (In1l x) = In1 arbitrary"
+by (simp add: arbitrary3_def)
+
+lemma [simp]: "arbitrary3 (In1r x) = \<diamondsuit>"
+by (simp add: arbitrary3_def)
+
+lemma [simp]: "arbitrary3 (In2  x) = In2 arbitrary"
+by (simp add: arbitrary3_def)
+
+lemma [simp]: "arbitrary3 (In3  x) = In3 arbitrary"
+by (simp add: arbitrary3_def)
+
+
+section "exception throwing and catching"
+
+constdefs
+  throw :: "val \<Rightarrow> abopt \<Rightarrow> abopt"
+ "throw a' x \<equiv> abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)"
+
+lemma throw_def2: 
+ "throw a' x = abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)"
+apply (unfold throw_def)
+apply (simp (no_asm))
+done
+
+constdefs
+  fits    :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60)
+ "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"
+
+lemma fits_Null [simp]: "G,s\<turnstile>Null fits T"
+by (simp add: fits_def)
+
+
+lemma fits_Addr_RefT [simp]:
+  "G,s\<turnstile>Addr a fits RefT t = G\<turnstile>obj_ty (the (heap s a))\<preceq>RefT t"
+by (simp add: fits_def)
+
+lemma fitsD: "\<And>X. G,s\<turnstile>a' fits T \<Longrightarrow> (\<exists>pt. T = PrimT pt) \<or>  
+  (\<exists>t. T = RefT t) \<and> a' = Null \<or>  
+  (\<exists>t. T = RefT t) \<and> a' \<noteq> Null \<and>  G\<turnstile>obj_ty (lookup_obj s a')\<preceq>T"
+apply (unfold fits_def)
+apply (case_tac "\<exists>pt. T = PrimT pt")
+apply  simp_all
+apply (case_tac "T")
+defer 
+apply (case_tac "a' = Null")
+apply  simp_all
+done
+
+constdefs
+  catch ::"prog \<Rightarrow> state \<Rightarrow> qtname \<Rightarrow> bool"      ("_,_\<turnstile>catch _"[61,61,61]60)
+ "G,s\<turnstile>catch C\<equiv>\<exists>xc. abrupt s=Some (Xcpt xc) \<and> 
+                    G,store s\<turnstile>Addr (the_Loc xc) fits Class C"
+
+lemma catch_Norm [simp]: "\<not>G,Norm s\<turnstile>catch tn"
+apply (unfold catch_def)
+apply (simp (no_asm))
+done
+
+lemma catch_XcptLoc [simp]: 
+  "G,(Some (Xcpt (Loc a)),s)\<turnstile>catch C = G,s\<turnstile>Addr a fits Class C"
+apply (unfold catch_def)
+apply (simp (no_asm))
+done
+
+constdefs
+  new_xcpt_var :: "vname \<Rightarrow> state \<Rightarrow> state"
+ "new_xcpt_var vn \<equiv> 
+     \<lambda>(x,s). Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)"
+
+lemma new_xcpt_var_def2 [simp]: 
+ "new_xcpt_var vn (x,s) = 
+    Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)"
+apply (unfold new_xcpt_var_def)
+apply (simp (no_asm))
+done
+
+
+
+section "misc"
+
+constdefs
+
+  assign     :: "('a \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> 'a \<Rightarrow> state \<Rightarrow> state"
+ "assign f v \<equiv> \<lambda>(x,s). let (x',s') = (if x = None then f v else id) (x,s)
+		   in  (x',if x' = None then s' else s)"
+
+(*
+lemma assign_Norm_Norm [simp]: 
+"f v \<lparr>abrupt=None,store=s\<rparr> = \<lparr>abrupt=None,store=s'\<rparr> 
+ \<Longrightarrow> assign f v \<lparr>abrupt=None,store=s\<rparr> = \<lparr>abrupt=None,store=s'\<rparr>"
+by (simp add: assign_def Let_def)
+*)
+
+lemma assign_Norm_Norm [simp]: 
+"f v (Norm s) = Norm s' \<Longrightarrow> assign f v (Norm s) = Norm s'"
+by (simp add: assign_def Let_def)
+
+(*
+lemma assign_Norm_Some [simp]: 
+  "\<lbrakk>abrupt (f v \<lparr>abrupt=None,store=s\<rparr>) = Some y\<rbrakk> 
+   \<Longrightarrow> assign f v \<lparr>abrupt=None,store=s\<rparr> = \<lparr>abrupt=Some y,store =s\<rparr>"
+by (simp add: assign_def Let_def split_beta)
+*)
+
+lemma assign_Norm_Some [simp]: 
+  "\<lbrakk>abrupt (f v (Norm s)) = Some y\<rbrakk> 
+   \<Longrightarrow> assign f v (Norm s) = (Some y,s)"
+by (simp add: assign_def Let_def split_beta)
+
+
+lemma assign_Some [simp]: 
+"assign f v (Some x,s) = (Some x,s)" 
+by (simp add: assign_def Let_def split_beta)
+
+lemma assign_supd [simp]: 
+"assign (\<lambda>v. supd (f v)) v (x,s)  
+  = (x, if x = None then f v s else s)"
+apply auto
+done
+
+lemma assign_raise_if [simp]: 
+  "assign (\<lambda>v (x,s). ((raise_if (b s v) xcpt) x, f v s)) v (x, s) =  
+  (raise_if (b s v) xcpt x, if x=None \<and> \<not>b s v then f v s else s)"
+apply (case_tac "x = None")
+apply auto
+done
+
+(*
+lemma assign_raise_if [simp]: 
+  "assign (\<lambda>v s. \<lparr>abrupt=(raise_if (b (store s) v) xcpt) (abrupt s),
+                  store = f v (store s)\<rparr>) v s =  
+  \<lparr>abrupt=raise_if (b (store s) v) xcpt (abrupt s),
+   store= if (abrupt s)=None \<and> \<not>b (store s) v 
+             then f v (store s) else (store s)\<rparr>"
+apply (case_tac "abrupt s = None")
+apply auto
+done
+*)
+
+constdefs
+
+  init_comp_ty :: "ty \<Rightarrow> stmt"
+ "init_comp_ty T \<equiv> if (\<exists>C. T = Class C) then Init (the_Class T) else Skip"
+
+lemma init_comp_ty_PrimT [simp]: "init_comp_ty (PrimT pt) = Skip"
+apply (unfold init_comp_ty_def)
+apply (simp (no_asm))
+done
+
+constdefs
+
+(*
+  target  :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname"
+ "target m s a' t 
+    \<equiv> if m = IntVir
+	 then obj_class (lookup_obj s a') 
+         else the_Class (RefT t)"
+*)
+
+ invocation_class  :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname"
+ "invocation_class m s a' statT 
+    \<equiv> (case m of
+         Static \<Rightarrow> if (\<exists> statC. statT = ClassT statC) 
+                      then the_Class (RefT statT) 
+                      else Object
+       | SuperM \<Rightarrow> the_Class (RefT statT)
+       | IntVir \<Rightarrow> obj_class (lookup_obj s a'))"
+
+invocation_declclass::"prog \<Rightarrow> inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> qtname"
+"invocation_declclass G m s a' statT sig 
+   \<equiv> declclass (the (dynlookup G statT 
+                                (invocation_class m s a' statT)
+                                sig))" 
+  
+lemma invocation_class_IntVir [simp]: 
+"invocation_class IntVir s a' statT = obj_class (lookup_obj s a')"
+by (simp add: invocation_class_def)
+
+lemma dynclass_SuperM [simp]: 
+ "invocation_class SuperM s a' statT = the_Class (RefT statT)"
+by (simp add: invocation_class_def)
+(*
+lemma invocation_class_notIntVir [simp]: 
+ "m \<noteq> IntVir \<Longrightarrow> invocation_class m s a' statT = the_Class (RefT statT)"
+by (simp add: invocation_class_def)
+*)
+
+lemma invocation_class_Static [simp]: 
+  "invocation_class Static s a' statT = (if (\<exists> statC. statT = ClassT statC) 
+                                            then the_Class (RefT statT) 
+                                            else Object)"
+by (simp add: invocation_class_def)
+
+constdefs
+  init_lvars :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> inv_mode \<Rightarrow> val \<Rightarrow> val list \<Rightarrow>
+		   state \<Rightarrow> state"
+ "init_lvars G C sig mode a' pvs 
+   \<equiv> \<lambda> (x,s). 
+      let m = mthd (the (methd G C sig));
+          l = \<lambda> k. 
+              (case k of
+                 EName e 
+                   \<Rightarrow> (case e of 
+                         VNam v \<Rightarrow> (init_vals (table_of (lcls (mbody m)))
+                                                     ((pars m)[\<mapsto>]pvs)) v
+                       | Res    \<Rightarrow> Some (default_val (resTy m)))
+               | This 
+                   \<Rightarrow> (if mode=Static then None else Some a'))
+      in set_lvars l (if mode = Static then x else np a' x,s)"
+
+
+
+lemma init_lvars_def2: "init_lvars G C sig mode a' pvs (x,s) =  
+  set_lvars 
+    (\<lambda> k. 
+       (case k of
+          EName e 
+            \<Rightarrow> (case e of 
+                  VNam v 
+                  \<Rightarrow> (init_vals 
+                       (table_of (lcls (mbody (mthd (the (methd G C sig))))))
+                                 ((pars (mthd (the (methd G C sig))))[\<mapsto>]pvs)) v
+               | Res \<Rightarrow> Some (default_val (resTy (mthd (the (methd G C sig))))))
+        | This 
+            \<Rightarrow> (if mode=Static then None else Some a')))
+    (if mode = Static then x else np a' x,s)"
+apply (unfold init_lvars_def)
+apply (simp (no_asm) add: Let_def)
+done
+
+constdefs
+  body :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> expr"
+ "body G C sig \<equiv> let m = the (methd G C sig) 
+                 in Body (declclass m) (stmt (mbody (mthd m)))"
+
+lemma body_def2: 
+"body G C sig = Body  (declclass (the (methd G C sig))) 
+                      (stmt (mbody (mthd (the (methd G C sig)))))"
+apply (unfold body_def Let_def)
+apply auto
+done
+
+section "variables"
+
+constdefs
+
+  lvar :: "lname \<Rightarrow> st \<Rightarrow> vvar"
+ "lvar vn s \<equiv> (the (locals s vn), \<lambda>v. supd (lupd(vn\<mapsto>v)))"
+
+  fvar :: "qtname \<Rightarrow> bool \<Rightarrow> vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
+ "fvar C stat fn a' s 
+    \<equiv> let (oref,xf) = if stat then (Stat C,id)
+                              else (Heap (the_Addr a'),np a');
+	          n = Inl (fn,C); 
+                  f = (\<lambda>v. supd (upd_gobj oref n v)) 
+      in ((the (values (the (globs (store s) oref)) n),f),abupd xf s)"
+
+  avar :: "prog \<Rightarrow> val \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
+ "avar G i' a' s 
+    \<equiv> let   oref = Heap (the_Addr a'); 
+               i = the_Intg i'; 
+               n = Inr i;
+        (T,k,cs) = the_Arr (globs (store s) oref); 
+               f = (\<lambda>v (x,s). (raise_if (\<not>G,s\<turnstile>v fits T) 
+                                           ArrStore x
+                              ,upd_gobj oref n v s)) 
+      in ((the (cs n),f)
+         ,abupd (raise_if (\<not>i in_bounds k) IndOutBound \<circ> np a') s)"
+
+lemma fvar_def2: "fvar C stat fn a' s =  
+  ((the 
+     (values 
+      (the (globs (store s) (if stat then Stat C else Heap (the_Addr a')))) 
+      (Inl (fn,C)))
+   ,(\<lambda>v. supd (upd_gobj (if stat then Stat C else Heap (the_Addr a')) 
+                        (Inl (fn,C)) 
+                        v)))
+  ,abupd (if stat then id else np a') s)
+  "
+apply (unfold fvar_def)
+apply (simp (no_asm) add: Let_def split_beta)
+done
+
+lemma avar_def2: "avar G i' a' s =  
+  ((the ((snd(snd(the_Arr (globs (store s) (Heap (the_Addr a')))))) 
+           (Inr (the_Intg i')))
+   ,(\<lambda>v (x,s').  (raise_if (\<not>G,s'\<turnstile>v fits (fst(the_Arr (globs (store s)
+                                                   (Heap (the_Addr a')))))) 
+                            ArrStore x
+                 ,upd_gobj  (Heap (the_Addr a')) 
+                               (Inr (the_Intg i')) v s')))
+  ,abupd (raise_if (\<not>(the_Intg i') in_bounds (fst(snd(the_Arr (globs (store s) 
+                   (Heap (the_Addr a'))))))) IndOutBound \<circ> np a')
+          s)"
+apply (unfold avar_def)
+apply (simp (no_asm) add: Let_def split_beta)
+done
+
+
+section "evaluation judgments"
+
+consts
+  eval   :: "prog \<Rightarrow> (state \<times> term    \<times> vals \<times> state) set"
+  halloc::  "prog \<Rightarrow> (state \<times> obj_tag \<times> loc  \<times> state) set"
+  sxalloc:: "prog \<Rightarrow> (state                  \<times> state) set"
+
+
+syntax
+eval ::"[prog,state,term,vals*state]=>bool"("_|-_ -_>-> _"  [61,61,80,   61]60)
+exec ::"[prog,state,stmt      ,state]=>bool"("_|-_ -_-> _"   [61,61,65,   61]60)
+evar ::"[prog,state,var  ,vvar,state]=>bool"("_|-_ -_=>_-> _"[61,61,90,61,61]60)
+eval_::"[prog,state,expr ,val, state]=>bool"("_|-_ -_->_-> _"[61,61,80,61,61]60)
+evals::"[prog,state,expr list ,
+		    val  list ,state]=>bool"("_|-_ -_#>_-> _"[61,61,61,61,61]60)
+hallo::"[prog,state,obj_tag,
+	             loc,state]=>bool"("_|-_ -halloc _>_-> _"[61,61,61,61,61]60)
+sallo::"[prog,state        ,state]=>bool"("_|-_ -sxalloc-> _"[61,61,      61]60)
+
+syntax (xsymbols)
+eval ::"[prog,state,term,vals\<times>state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>_\<succ>\<rightarrow> _"  [61,61,80,   61]60)
+exec ::"[prog,state,stmt      ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<rightarrow> _"   [61,61,65,   61]60)
+evar ::"[prog,state,var  ,vvar,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_=\<succ>_\<rightarrow> _"[61,61,90,61,61]60)
+eval_::"[prog,state,expr ,val ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_-\<succ>_\<rightarrow> _"[61,61,80,61,61]60)
+evals::"[prog,state,expr list ,
+		    val  list ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
+hallo::"[prog,state,obj_tag,
+	             loc,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>halloc _\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
+sallo::"[prog,state,        state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>sxalloc\<rightarrow> _"[61,61,      61]60)
+
+translations
+  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow>  w___s' " == "(s,t,w___s') \<in> eval G"
+  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow> (w,  s')" <= "(s,t,w,  s') \<in> eval G"
+  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow> (w,x,s')" <= "(s,t,w,x,s') \<in> eval G"
+  "G\<turnstile>s \<midarrow>c    \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In1r c\<succ>\<rightarrow> (\<diamondsuit>,x,s')"
+  "G\<turnstile>s \<midarrow>c    \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In1r c\<succ>\<rightarrow> (\<diamondsuit>  ,  s')"
+  "G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In1l e\<succ>\<rightarrow> (In1 v ,x,s')"
+  "G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In1l e\<succ>\<rightarrow> (In1 v ,  s')"
+  "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In2  e\<succ>\<rightarrow> (In2 vf,x,s')"
+  "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In2  e\<succ>\<rightarrow> (In2 vf,  s')"
+  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In3  e\<succ>\<rightarrow> (In3 v ,x,s')"
+  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In3  e\<succ>\<rightarrow> (In3 v ,  s')"
+  "G\<turnstile>s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,s')" <= "(s,oi,a,x,s') \<in> halloc G"
+  "G\<turnstile>s \<midarrow>halloc oi\<succ>a\<rightarrow>    s' " == "(s,oi,a,  s') \<in> halloc G"
+  "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>     (x,s')" <= "(s     ,x,s') \<in> sxalloc G"
+  "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>        s' " == "(s     ,  s') \<in> sxalloc G"
+
+inductive "halloc G" intros (* allocating objects on the heap, cf. 12.5 *)
+
+  Abrupt: 
+  "G\<turnstile>(Some x,s) \<midarrow>halloc oi\<succ>arbitrary\<rightarrow> (Some x,s)"
+
+  New:	  "\<lbrakk>new_Addr (heap s) = Some a; 
+	    (x,oi') = (if atleast_free (heap s) (Suc (Suc 0)) then (None,oi)
+		       else (Some (Xcpt (Loc a)),CInst (SXcpt OutOfMemory)))\<rbrakk>
+            \<Longrightarrow>
+	    G\<turnstile>Norm s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,init_obj G oi' (Heap a) s)"
+
+inductive "sxalloc G" intros (* allocating exception objects for
+	 	 	      standard exceptions (other than OutOfMemory) *)
+
+  Norm:	 "G\<turnstile> Norm              s   \<midarrow>sxalloc\<rightarrow>  Norm             s"
+
+  XcptL: "G\<turnstile>(Some (Xcpt (Loc a) ),s)  \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s)"
+
+  SXcpt: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>halloc (CInst (SXcpt xn))\<succ>a\<rightarrow> (x,s1)\<rbrakk> \<Longrightarrow>
+	  G\<turnstile>(Some (Xcpt (Std xn)),s0) \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s1)"
+
+
+inductive "eval G" intros
+
+(* propagation of abrupt completion *)
+
+  (* cf. 14.1, 15.5 *)
+  Abrupt: 
+   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (arbitrary3 t,(Some xc,s))"
+
+
+(* execution of statements *)
+
+  (* cf. 14.5 *)
+  Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<rightarrow> Norm s"
+
+  (* cf. 14.7 *)
+  Expr:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+				  G\<turnstile>Norm s0 \<midarrow>Expr e\<rightarrow> s1"
+
+  Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<rightarrow> s1\<rbrakk> \<Longrightarrow>
+                                G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<rightarrow> abupd (absorb (Break l)) s1"
+  (* cf. 14.2 *)
+  Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<rightarrow> s1;
+	  G\<turnstile>     s1 \<midarrow>c2 \<rightarrow> s2\<rbrakk> \<Longrightarrow>
+				 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<rightarrow> s2"
+
+
+  (* cf. 14.8.2 *)
+  If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
+	  G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<rightarrow> s2"
+
+  (* cf. 14.10, 14.10.1 *)
+  (*      G\<turnstile>Norm s0 \<midarrow>If(e) (c;; While(e) c) Else Skip\<rightarrow> s3 *)
+  (* A "continue jump" from the while body c is handled by 
+     this rule. If a continue jump with the proper label was invoked inside c
+     this label (Cont l) is deleted out of the abrupt component of the state 
+     before the iterative evaluation of the while statement.
+     A "break jump" is handled by the Lab Statement (Lab l (while\<dots>).
+  *)
+  Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
+	  if normal s1 \<and> the_Bool b 
+             then (G\<turnstile>s1 \<midarrow>c\<rightarrow> s2 \<and> 
+                   G\<turnstile>(abupd (absorb (Cont l)) s2) \<midarrow>l\<bullet> While(e) c\<rightarrow> s3)
+	     else s3 = s1\<rbrakk> \<Longrightarrow>
+			      G\<turnstile>Norm s0 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
+
+  Do: "G\<turnstile>Norm s \<midarrow>Do j\<rightarrow> (Some (Jump j), s)"
+   
+  (* cf. 14.16 *)
+  Throw: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+				 G\<turnstile>Norm s0 \<midarrow>Throw e\<rightarrow> abupd (throw a') s1"
+
+  (* cf. 14.18.1 *)
+  Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; 
+	  if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2\<rbrakk> \<Longrightarrow>
+		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(C vn) c2\<rightarrow> s3"
+
+  (* cf. 14.18.2 *)
+  Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1,s1);
+	  G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+               G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<rightarrow> abupd (abrupt_if (x1\<noteq>None) x1) s2"
+
+  (* cf. 12.4.2, 8.5 *)
+  Init:	"\<lbrakk>the (class G C) = c;
+	  if inited C (globs s0) then s3 = Norm s0
+	  else (G\<turnstile>Norm (init_class_obj G C s0) 
+		  \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1 \<and>
+	       G\<turnstile>set_lvars empty s1 \<midarrow>init c\<rightarrow> s2 \<and> s3 = restore_lvars s1 s2)\<rbrakk> 
+              \<Longrightarrow>
+		 G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s3"
+
+(* evaluation of expressions *)
+
+  (* cf. 15.8.1, 12.4.1 *)
+  NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1;
+	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
+
+  (* cf. 15.9.1, 12.4.1 *)
+  NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<rightarrow> s2; 
+	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
+	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<rightarrow> s3"
+
+  (* cf. 15.15 *)
+  Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
+	  s2 = abupd (raise_if (\<not>G,store s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
+			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<rightarrow> s2"
+
+  (* cf. 15.19.2 *)
+  Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
+	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
+			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<rightarrow> s1"
+
+  (* cf. 15.7.1 *)
+  Lit:	"G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<rightarrow> Norm s"
+
+
+
+  (* cf. 15.10.2 *)
+  Super: "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<rightarrow> Norm s"
+
+  (* cf. 15.2 *)
+  Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<rightarrow> s1"
+
+  (* cf. 15.25.1 *)
+  Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<rightarrow> s1;
+          G\<turnstile>     s1 \<midarrow>e-\<succ>v  \<rightarrow> s2\<rbrakk> \<Longrightarrow>
+				   G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v\<rightarrow> assign f v s2"
+
+  (* cf. 15.24 *)
+  Cond:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1;
+          G\<turnstile>     s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+			    G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<rightarrow> s2"
+
+
+  (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5 *)
+  Call:	
+  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1; G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2;
+    D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>;
+    G\<turnstile>init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs s2 
+         \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<rightarrow> s3\<rbrakk>
+   \<Longrightarrow>
+       G\<turnstile>Norm s0 \<midarrow>{statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<rightarrow> (restore_lvars s2 s3)"
+
+  Methd:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+				G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<rightarrow> s1"
+
+  (* cf. 14.15, 12.4.1 *)
+  Body:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+ G\<turnstile>Norm s0 \<midarrow>Body D c -\<succ>the (locals (store s2) Result)\<rightarrow>abupd (absorb Ret) s2"
+
+(* evaluation of variables *)
+
+  (* cf. 15.13.1, 15.7.2 *)
+  LVar:	"G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<rightarrow> Norm s"
+
+  (* cf. 15.10.1, 12.4.1 *)
+  FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2;
+	  (v,s2') = fvar C stat fn a s2\<rbrakk> \<Longrightarrow>
+	  G\<turnstile>Norm s0 \<midarrow>{C,stat}e..fn=\<succ>v\<rightarrow> s2'"
+
+  (* cf. 15.12.1, 15.25.1 *)
+  AVar:	"\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2;
+	  (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
+	              G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<rightarrow> s2'"
+
+
+(* evaluation of expression lists *)
+
+  (* cf. 15.11.4.2 *)
+  Nil:
+				    "G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<rightarrow> Norm s0"
+
+  (* cf. 15.6.4 *)
+  Cons:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<rightarrow> s1;
+          G\<turnstile>     s1 \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+				   G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<rightarrow> s2"
+
+(* Rearrangement of premisses:
+[0,1(Abrupt),2(Skip),8(Do),4(Lab),28(Nil),29(Cons),25(LVar),15(Cast),16(Inst),
+ 17(Lit),18(Super),19(Acc),3(Expr),5(Comp),23(Methd),24(Body),21(Cond),6(If),
+ 7(Loop),11(Fin),9(Throw),13(NewC),14(NewA),12(Init),20(Ass),10(Try),26(FVar),
+ 27(AVar),22(Call)]
+*)
+ML {*
+bind_thm ("eval_induct_", rearrange_prems 
+[0,1,2,8,4,28,29,25,15,16,
+ 17,18,19,3,5,23,24,21,6,
+ 7,11,9,13,14,12,20,10,26,
+ 27,22] (thm "eval.induct"))
+*}
+
+lemmas eval_induct = eval_induct_ [split_format and and and and and and and and
+   and and and and s1 (* Acc *) and and s2 (* Comp *) and and and and and and 
+   s2 (* Fin *) and and s2 (* NewC *)] 
+
+declare split_if     [split del] split_if_asm     [split del] 
+        option.split [split del] option.split_asm [split del]
+inductive_cases halloc_elim_cases: 
+  "G\<turnstile>(Some xc,s) \<midarrow>halloc oi\<succ>a\<rightarrow> s'"
+  "G\<turnstile>(Norm    s) \<midarrow>halloc oi\<succ>a\<rightarrow> s'"
+
+inductive_cases sxalloc_elim_cases:
+ 	"G\<turnstile> Norm                 s  \<midarrow>sxalloc\<rightarrow> s'"
+ 	"G\<turnstile>(Some (Xcpt (Loc a )),s) \<midarrow>sxalloc\<rightarrow> s'"
+ 	"G\<turnstile>(Some (Xcpt (Std xn)),s) \<midarrow>sxalloc\<rightarrow> s'"
+inductive_cases sxalloc_cases: "G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'"
+
+lemma sxalloc_elim_cases2: "\<lbrakk>G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s';  
+       \<And>s.   \<lbrakk>s' = Norm s\<rbrakk> \<Longrightarrow> P;  
+       \<And>a s. \<lbrakk>s' = (Some (Xcpt (Loc a)),s)\<rbrakk> \<Longrightarrow> P  
+      \<rbrakk> \<Longrightarrow> P"
+apply cut_tac 
+apply (erule sxalloc_cases)
+apply blast+
+done
+
+declare not_None_eq [simp del] (* IntDef.Zero_def [simp del] *)
+declare split_paired_All [simp del] split_paired_Ex [simp del]
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac"
+*}
+inductive_cases eval_cases: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> vs'"
+
+inductive_cases eval_elim_cases:
+        "G\<turnstile>(Some xc,s) \<midarrow>t                         \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r Skip                      \<succ>\<rightarrow> xs'"
+        "G\<turnstile>Norm s \<midarrow>In1r (Do j)                    \<succ>\<rightarrow> xs'"
+        "G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> c)                    \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In3  ([])                      \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In3  (e#es)                    \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Lit w)                   \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In2  (LVar vn)                 \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Cast T e)                \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (e InstOf T)              \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Super)                   \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Acc va)                  \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (Expr e)                  \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (c1;; c2)                 \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Methd C sig)             \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Body D c)                \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2)            \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2)        \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> While(e) c)           \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2)           \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (Throw e)                 \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (NewC C)                  \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (New T[e])                \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Ass va e)                \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2)  \<succ>\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In2  ({C,stat}e..fn)           \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                 \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l ({statT,mode}e\<cdot>mn({pT}p)) \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (Init C)                  \<succ>\<rightarrow> xs'"
+declare not_None_eq [simp]  (* IntDef.Zero_def [simp] *)
+declare split_paired_All [simp] split_paired_Ex [simp]
+ML_setup {*
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
+*}
+declare split_if     [split] split_if_asm     [split] 
+        option.split [split] option.split_asm [split]
+
+lemma eval_Inj_elim: 
+ "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') 
+ \<Longrightarrow> case t of 
+       In1 ec \<Rightarrow> (case ec of 
+                    Inl e \<Rightarrow> (\<exists>v. w = In1 v) 
+                  | Inr c \<Rightarrow> w = \<diamondsuit>)  
+     | In2 e \<Rightarrow> (\<exists>v. w = In2 v) 
+     | In3 e \<Rightarrow> (\<exists>v. w = In3 v)"
+apply (erule eval_cases)
+apply auto
+apply (induct_tac "t")
+apply (induct_tac "a")
+apply auto
+done
+
+ML_setup {*
+fun eval_fun nam inj rhs =
+let
+  val name = "eval_" ^ nam ^ "_eq"
+  val lhs = "G\<turnstile>s \<midarrow>" ^ inj ^ " t\<succ>\<rightarrow> (w, s')"
+  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
+	(K [Auto_tac, ALLGOALS (ftac (thm "eval_Inj_elim")) THEN Auto_tac])
+  fun is_Inj (Const (inj,_) $ _) = true
+    | is_Inj _                   = false
+  fun pred (_ $ (Const ("Pair",_) $ _ $ 
+      (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ x $ _ ))) $ _ ) = is_Inj x
+in
+  make_simproc name lhs pred (thm name)
+end
+
+val eval_expr_proc =eval_fun "expr" "In1l" "\<exists>v.  w=In1 v   \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<rightarrow> s'"
+val eval_var_proc  =eval_fun "var"  "In2"  "\<exists>vf. w=In2 vf  \<and> G\<turnstile>s \<midarrow>t=\<succ>vf\<rightarrow> s'"
+val eval_exprs_proc=eval_fun "exprs""In3"  "\<exists>vs. w=In3 vs  \<and> G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs\<rightarrow> s'"
+val eval_stmt_proc =eval_fun "stmt" "In1r" "     w=\<diamondsuit> \<and> G\<turnstile>s \<midarrow>t    \<rightarrow> s'";
+Addsimprocs [eval_expr_proc,eval_var_proc,eval_exprs_proc,eval_stmt_proc];
+bind_thms ("AbruptIs", sum3_instantiate (thm "eval.Abrupt"))
+*}
+
+declare halloc.Abrupt [intro!] eval.Abrupt [intro!]  AbruptIs [intro!] 
+
+
+lemma eval_no_abrupt_lemma: 
+   "\<And>s s'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<Longrightarrow> normal s' \<longrightarrow> normal s"
+by (erule eval_cases, auto)
+
+lemma eval_no_abrupt: 
+  "G\<turnstile>(x,s) \<midarrow>t\<succ>\<rightarrow> (w,Norm s') = 
+        (x = None \<and> G\<turnstile>Norm s \<midarrow>t\<succ>\<rightarrow> (w,Norm s'))"
+apply auto
+apply (frule eval_no_abrupt_lemma, auto)+
+done
+
+ML {*
+local
+  fun is_None (Const ("Option.option.None",_)) = true
+    | is_None _ = false
+  fun pred (t as (_ $ (Const ("Pair",_) $
+     (Const ("Pair", _) $ x $ _) $ _ ) $ _)) = is_None x
+in
+  val eval_no_abrupt_proc = 
+  make_simproc "eval_no_abrupt" "G\<turnstile>(x,s) \<midarrow>e\<succ>\<rightarrow> (w,Norm s')" pred 
+          (thm "eval_no_abrupt")
+end;
+Addsimprocs [eval_no_abrupt_proc]
+*}
+
+
+lemma eval_abrupt_lemma: 
+  "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s') \<Longrightarrow> abrupt s=Some xc \<longrightarrow> s'= s \<and> v = arbitrary3 t"
+by (erule eval_cases, auto)
+
+lemma eval_abrupt: 
+ " G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (w,s') =  
+     (s'=(Some xc,s) \<and> w=arbitrary3 t \<and> 
+     G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (arbitrary3 t,(Some xc,s)))"
+apply auto
+apply (frule eval_abrupt_lemma, auto)+
+done
+
+ML {*
+local
+  fun is_Some (Const ("Pair",_) $ (Const ("Option.option.Some",_) $ _)$ _) =true
+    | is_Some _ = false
+  fun pred (_ $ (Const ("Pair",_) $
+     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
+       x))) $ _ ) = is_Some x
+in
+  val eval_abrupt_proc = 
+  make_simproc "eval_abrupt" 
+               "G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<rightarrow> (w,s')" pred (thm "eval_abrupt")
+end;
+Addsimprocs [eval_abrupt_proc]
+*}
+
+
+lemma LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else arbitrary)\<rightarrow> s"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: eval.Lit)
+
+lemma SkipI [intro!]: "G\<turnstile>s \<midarrow>Skip\<rightarrow> s"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: eval.Skip)
+
+lemma ExprI: "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>Expr e\<rightarrow> s'"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: eval.Expr)
+
+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"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: eval.Comp)
+
+lemma CondI: 
+  "\<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> 
+         G\<turnstile>s \<midarrow>e ? e1 : e2-\<succ>(if normal s1 then v else arbitrary)\<rightarrow> s2"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: eval.Cond)
+
+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>
+                 \<Longrightarrow> G\<turnstile>s \<midarrow>If(e) c1 Else c2\<rightarrow> s2"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: eval.If)
+
+lemma MethdI: "G\<turnstile>s \<midarrow>body G C sig-\<succ>v\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>Methd C sig-\<succ>v\<rightarrow> s'"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: eval.Methd)
+
+lemma eval_Call: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1; G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>pvs\<rightarrow> s2;  
+       D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>; 
+       G\<turnstile>init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' pvs s2 
+        \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> v\<rightarrow> s3; 
+       s3' = restore_lvars s2 s3\<rbrakk> \<Longrightarrow>  
+       G\<turnstile>Norm s0 \<midarrow>{statT,mode}e\<cdot>mn({pTs}ps)-\<succ>v\<rightarrow> s3'"
+apply (drule eval.Call, assumption)
+apply (rule HOL.refl)
+apply simp+
+done
+
+lemma eval_Init: 
+"\<lbrakk>if inited C (globs s0) then s3 = Norm s0 
+  else G\<turnstile>Norm (init_class_obj G C s0)  
+         \<midarrow>(if C = Object then Skip else Init (super (the (class G C))))\<rightarrow> s1 \<and>
+       G\<turnstile>set_lvars empty s1 \<midarrow>(init (the (class G C)))\<rightarrow> s2 \<and> 
+      s3 = restore_lvars s1 s2\<rbrakk> \<Longrightarrow>  
+  G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s3"
+apply (rule eval.Init)
+apply auto
+done
+
+lemma init_done: "initd C s \<Longrightarrow> G\<turnstile>s \<midarrow>Init C\<rightarrow> s"
+apply (case_tac "s", simp)
+apply (case_tac "a")
+apply  safe
+apply (rule eval_Init)
+apply   auto
+done
+
+lemma eval_StatRef: 
+"G\<turnstile>s \<midarrow>StatRef rt-\<succ>(if abrupt s=None then Null else arbitrary)\<rightarrow> s"
+apply (case_tac "s", simp)
+apply (case_tac "a = None")
+apply (auto del: eval.Abrupt intro!: eval.intros)
+done
+
+
+lemma SkipD [dest!]: "G\<turnstile>s \<midarrow>Skip\<rightarrow> s' \<Longrightarrow> s' = s" 
+apply (erule eval_cases)
+by auto
+
+lemma Skip_eq [simp]: "G\<turnstile>s \<midarrow>Skip\<rightarrow> s' = (s = s')"
+by auto
+
+(*unused*)
+lemma init_retains_locals [rule_format (no_asm)]: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<Longrightarrow>  
+  (\<forall>C. t=In1r (Init C) \<longrightarrow> locals (store s) = locals (store s'))"
+apply (erule eval.induct)
+apply (simp (no_asm_use) split del: split_if_asm option.split_asm)+
+apply auto
+done
+
+lemma halloc_xcpt [dest!]: 
+  "\<And>s'. G\<turnstile>(Some xc,s) \<midarrow>halloc oi\<succ>a\<rightarrow> s' \<Longrightarrow> s'=(Some xc,s)"
+apply (erule_tac halloc_elim_cases)
+by auto
+
+(*
+G\<turnstile>(x,(h,l)) \<midarrow>e\<succ>v\<rightarrow> (x',(h',l'))) \<Longrightarrow> l This = l' This"
+G\<turnstile>(x,(h,l)) \<midarrow>s  \<rightarrow> (x',(h',l'))) \<Longrightarrow> l This = l' This"
+*)
+
+lemma eval_Methd: 
+  "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')"
+apply (case_tac "s")
+apply (case_tac "a")
+apply clarsimp+
+apply (erule eval.Methd)
+apply (drule eval_abrupt_lemma)
+apply force
+done
+
+
+section "single valued"
+
+lemma unique_halloc [rule_format (no_asm)]: 
+  "\<And>s as as'. (s,oi,as)\<in>halloc G \<Longrightarrow> (s,oi,as')\<in>halloc G \<longrightarrow> as'=as"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (erule halloc.induct)
+apply  (auto elim!: halloc_elim_cases split del: split_if split_if_asm)
+apply (drule trans [THEN sym], erule sym) 
+defer
+apply (drule trans [THEN sym], erule sym)
+apply auto
+done
+
+
+lemma single_valued_halloc: 
+  "single_valued {((s,oi),(a,s')). G\<turnstile>s \<midarrow>halloc oi\<succ>a \<rightarrow> s'}"
+apply (unfold single_valued_def)
+by (clarsimp, drule (1) unique_halloc, auto)
+
+
+lemma unique_sxalloc [rule_format (no_asm)]: 
+  "\<And>s s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'' \<longrightarrow> s'' = s'"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (erule sxalloc.induct)
+apply   (auto dest: unique_halloc elim!: sxalloc_elim_cases 
+              split del: split_if split_if_asm)
+done
+
+lemma single_valued_sxalloc: "single_valued {(s,s'). G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'}"
+apply (unfold single_valued_def)
+apply (blast dest: unique_sxalloc)
+done
+
+lemma split_pairD: "(x,y) = p \<Longrightarrow> x = fst p & y = snd p"
+by auto
+
+lemma unique_eval [rule_format (no_asm)]: 
+  "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws \<Longrightarrow> (\<forall>ws'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws' \<longrightarrow> ws' = ws)"
+apply (case_tac "ws")
+apply (simp only:)
+apply (erule thin_rl)
+apply (erule eval_induct)
+apply (tactic {* ALLGOALS (EVERY'
+      [strip_tac, rotate_tac ~1, eresolve_tac (thms "eval_elim_cases")]) *})
+(* 29 subgoals *)
+prefer 26 (* Try *) 
+apply (simp (no_asm_use) only: split add: split_if_asm)
+(* 32 subgoals *)
+prefer 28 (* Init *)
+apply (case_tac "inited C (globs s0)", (simp only: if_True if_False)+)
+prefer 24 (* While *)
+apply (simp (no_asm_use) only: split add: split_if_asm, blast)
+apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp)
+apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp)
+apply blast
+(* 31 subgoals *)
+apply (blast dest: unique_sxalloc unique_halloc split_pairD)+
+done
+
+(* unused *)
+lemma single_valued_eval: 
+ "single_valued {((s,t),vs'). G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> vs'}"
+apply (unfold single_valued_def)
+by (clarify, drule (1) unique_eval, auto)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Evaln.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,373 @@
+(*  Title:      isabelle/Bali/Evaln.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1999 Technische Universitaet Muenchen
+*)
+header {* Operational evaluation (big-step) semantics of Java expressions and 
+          statements
+*}
+
+theory Evaln = Eval:
+
+text {*
+Variant of eval relation with counter for bounded recursive depth
+Evaln could completely replace Eval.
+*}
+
+consts
+
+  evaln	:: "prog \<Rightarrow> (state \<times> term \<times> nat \<times> vals \<times> state) set"
+
+syntax
+
+  evaln	:: "[prog, state, term,        nat, vals * state] => bool"
+				("_|-_ -_>-_-> _"   [61,61,80,   61,61] 60)
+  evarn	:: "[prog, state, var  , vvar        , nat, state] => bool"
+				("_|-_ -_=>_-_-> _" [61,61,90,61,61,61] 60)
+  eval_n:: "[prog, state, expr , val         , nat, state] => bool"
+				("_|-_ -_->_-_-> _" [61,61,80,61,61,61] 60)
+  evalsn:: "[prog, state, expr list, val list, nat, state] => bool"
+				("_|-_ -_#>_-_-> _" [61,61,61,61,61,61] 60)
+  execn	:: "[prog, state, stmt ,               nat, state] => bool"
+				("_|-_ -_-_-> _"    [61,61,65,   61,61] 60)
+
+syntax (xsymbols)
+
+  evaln	:: "[prog, state, term,         nat, vals \<times> state] \<Rightarrow> bool"
+				("_\<turnstile>_ \<midarrow>_\<succ>\<midarrow>_\<rightarrow> _"   [61,61,80,   61,61] 60)
+  evarn	:: "[prog, state, var  , vvar         , nat, state] \<Rightarrow> bool"
+				("_\<turnstile>_ \<midarrow>_=\<succ>_\<midarrow>_\<rightarrow> _" [61,61,90,61,61,61] 60)
+  eval_n:: "[prog, state, expr , val ,          nat, state] \<Rightarrow> bool"
+				("_\<turnstile>_ \<midarrow>_-\<succ>_\<midarrow>_\<rightarrow> _" [61,61,80,61,61,61] 60)
+  evalsn:: "[prog, state, expr list, val  list, nat, state] \<Rightarrow> bool"
+				("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<midarrow>_\<rightarrow> _" [61,61,61,61,61,61] 60)
+  execn	:: "[prog, state, stmt ,                nat, state] \<Rightarrow> bool"
+				("_\<turnstile>_ \<midarrow>_\<midarrow>_\<rightarrow> _"     [61,61,65,   61,61] 60)
+
+translations
+
+  "G\<turnstile>s \<midarrow>t    \<succ>\<midarrow>n\<rightarrow>  w___s' " == "(s,t,n,w___s') \<in> evaln G"
+  "G\<turnstile>s \<midarrow>t    \<succ>\<midarrow>n\<rightarrow> (w,  s')" <= "(s,t,n,w,  s') \<in> evaln G"
+  "G\<turnstile>s \<midarrow>t    \<succ>\<midarrow>n\<rightarrow> (w,x,s')" <= "(s,t,n,w,x,s') \<in> evaln G"
+  "G\<turnstile>s \<midarrow>c     \<midarrow>n\<rightarrow> (x,s')" <= "G\<turnstile>s \<midarrow>In1r  c\<succ>\<midarrow>n\<rightarrow> (\<diamondsuit>    ,x,s')"
+  "G\<turnstile>s \<midarrow>c     \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In1r  c\<succ>\<midarrow>n\<rightarrow> (\<diamondsuit>    ,  s')"
+  "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')"
+  "G\<turnstile>s \<midarrow>e-\<succ>v  \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In1l e\<succ>\<midarrow>n\<rightarrow> (In1 v ,  s')"
+  "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')"
+  "G\<turnstile>s \<midarrow>e=\<succ>vf \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In2  e\<succ>\<midarrow>n\<rightarrow> (In2 vf,  s')"
+  "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')"
+  "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')"
+
+
+inductive "evaln G" intros
+
+(* propagation of abrupt completion *)
+
+  Abrupt:   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (arbitrary3 t,(Some xc,s))"
+
+
+(* evaluation of variables *)
+
+  LVar:	"G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<midarrow>n\<rightarrow> Norm s"
+
+  FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s2;
+	  (v,s2') = fvar C stat fn a' s2\<rbrakk> \<Longrightarrow>
+	  G\<turnstile>Norm s0 \<midarrow>{C,stat}e..fn=\<succ>v\<midarrow>n\<rightarrow> s2'"
+
+  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; 
+	  (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
+	              G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<midarrow>n\<rightarrow> s2'"
+
+
+
+
+(* evaluation of expressions *)
+
+  NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1;
+	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<midarrow>n\<rightarrow> s2"
+
+  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; 
+	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
+	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<midarrow>n\<rightarrow> s3"
+
+  Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
+	  s2 = abupd (raise_if (\<not>G,snd s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
+			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<midarrow>n\<rightarrow> s2"
+
+  Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
+	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
+			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<midarrow>n\<rightarrow> s1"
+
+  Lit:			   "G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<midarrow>n\<rightarrow> Norm s"
+
+  Super:		   "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<midarrow>n\<rightarrow> Norm s"
+
+  Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<midarrow>n\<rightarrow> s1"
+
+  Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<midarrow>n\<rightarrow> s1;
+          G\<turnstile>     s1 \<midarrow>e-\<succ>v     \<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+				   G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v\<midarrow>n\<rightarrow> assign f v s2"
+
+  Cond:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<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>
+			    G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<midarrow>n\<rightarrow> s2"
+
+  Call:	
+  "\<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;
+    D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>; 
+    G\<turnstile>init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs s2
+            \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<midarrow>n\<rightarrow> s3\<rbrakk>
+   \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>{statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<midarrow>n\<rightarrow> (restore_lvars s2 s3)"
+
+  Methd:"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+				G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<midarrow>Suc n\<rightarrow> s1"
+
+  Body:	"\<lbrakk>G\<turnstile>Norm s0\<midarrow>Init D\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2\<rbrakk>\<Longrightarrow>
+ G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)\<midarrow>n\<rightarrow>abupd (absorb Ret) s2"
+
+(* evaluation of expression lists *)
+
+  Nil:
+				"G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<midarrow>n\<rightarrow> Norm s0"
+
+  Cons:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<midarrow>n\<rightarrow> s1;
+          G\<turnstile>     s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+			     G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<midarrow>n\<rightarrow> s2"
+
+
+(* execution of statements *)
+
+  Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<midarrow>n\<rightarrow> Norm s"
+
+  Expr:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+				  G\<turnstile>Norm s0 \<midarrow>Expr e\<midarrow>n\<rightarrow> s1"
+
+  Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+                             G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<midarrow>n\<rightarrow> abupd (absorb (Break l)) s1"
+
+  Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1;
+	  G\<turnstile>     s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+				 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<midarrow>n\<rightarrow> s2"
+
+  If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
+	  G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<midarrow>n\<rightarrow> s2"
+
+  Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
+	  if normal s1 \<and> the_Bool b 
+             then (G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2 \<and> 
+                   G\<turnstile>(abupd (absorb (Cont l)) s2) \<midarrow>l\<bullet> While(e) c\<midarrow>n\<rightarrow> s3)
+	     else s3 = s1\<rbrakk> \<Longrightarrow>
+			      G\<turnstile>Norm s0 \<midarrow>l\<bullet> While(e) c\<midarrow>n\<rightarrow> s3"
+  
+  Do: "G\<turnstile>Norm s \<midarrow>Do j\<midarrow>n\<rightarrow> (Some (Jump j), s)"
+  
+  Throw:"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+				 G\<turnstile>Norm s0 \<midarrow>Throw e\<midarrow>n\<rightarrow> abupd (throw a') s1"
+
+  Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2;
+	  if G,s2\<turnstile>catch tn then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3 else s3 = s2\<rbrakk>
+          \<Longrightarrow>
+		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(tn vn) c2\<midarrow>n\<rightarrow> s3"
+
+  Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> (x1,s1);
+	  G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+              G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<midarrow>n\<rightarrow> abupd (abrupt_if (x1\<noteq>None) x1) s2"
+  
+  Init:	"\<lbrakk>the (class G C) = c;
+	  if inited C (globs s0) then s3 = Norm s0
+	  else (G\<turnstile>Norm (init_class_obj G C s0)
+	          \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1 \<and>
+	        G\<turnstile>set_lvars empty s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2 \<and> 
+                s3 = restore_lvars s1 s2)\<rbrakk>
+          \<Longrightarrow>
+		 G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s3"
+monos
+  if_def2
+
+lemma evaln_eval: "\<And>ws. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> ws \<Longrightarrow> G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (erule evaln.induct)
+apply (rule eval.intros, (assumption+)?,(force split del: split_if)?)+
+done
+
+
+lemma Suc_le_D_lemma: "\<lbrakk>Suc n <= m'; (\<And>m. n <= m \<Longrightarrow> P (Suc m)) \<rbrakk> \<Longrightarrow> P m'"
+apply (frule Suc_le_D)
+apply fast
+done
+
+lemma evaln_nonstrict [rule_format (no_asm), elim]: 
+  "\<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"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (erule evaln.induct)
+apply (tactic {* ALLGOALS (EVERY'[strip_tac, TRY o etac (thm "Suc_le_D_lemma"),
+  REPEAT o smp_tac 1, 
+  resolve_tac (thms "evaln.intros") THEN_ALL_NEW TRY o atac]) *})
+(* 3 subgoals *)
+apply (auto split del: split_if)
+done
+
+lemmas evaln_nonstrict_Suc = evaln_nonstrict [OF _ le_refl [THEN le_SucI]]
+
+lemma evaln_max2: "\<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>n1\<rightarrow> ws1; G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>n2\<rightarrow> ws2\<rbrakk> \<Longrightarrow> 
+             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"
+apply (fast intro: le_maxI1 le_maxI2)
+done
+
+lemma evaln_max3: 
+"\<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>
+ G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws1 \<and>
+ G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws2 \<and> 
+ G\<turnstile>s3 \<midarrow>t3\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws3"
+apply (drule (1) evaln_max2, erule thin_rl)
+apply (fast intro!: le_maxI1 le_maxI2)
+done
+
+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)"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (erule eval.induct)
+apply (tactic {* ALLGOALS 
+         (asm_full_simp_tac (HOL_basic_ss addsplits [split_if_asm])) *})
+apply (tactic {* ALLGOALS (EVERY'[
+   REPEAT o eresolve_tac [exE, conjE], rtac exI,
+                     TRY o datac (thm "evaln_max3") 2, REPEAT o etac conjE,
+  resolve_tac (thms "evaln.intros") THEN_ALL_NEW 
+  force_tac (HOL_cs, HOL_ss)]) *})
+done
+
+declare split_if     [split del] split_if_asm     [split del]
+        option.split [split del] option.split_asm [split del]
+inductive_cases evaln_cases: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> vs'"
+
+inductive_cases evaln_elim_cases:
+	"G\<turnstile>(Some xc, s) \<midarrow>t                        \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r Skip                      \<succ>\<midarrow>n\<rightarrow> xs'"
+        "G\<turnstile>Norm s \<midarrow>In1r (Do j)                    \<succ>\<midarrow>n\<rightarrow> xs'"
+        "G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> c)                    \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In3  ([])                      \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In3  (e#es)                    \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Lit w)                   \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In2  (LVar vn)                 \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Cast T e)                \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (e InstOf T)              \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Super)                   \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Acc va)                  \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (Expr e)                  \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (c1;; c2)                 \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Methd C sig)             \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Body D c)                \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2)            \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2)        \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> While(e) c)           \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2)           \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (Throw e)                 \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (NewC C)                  \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (New T[e])                \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l (Ass va e)                \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2)  \<succ>\<midarrow>n\<rightarrow> xs'"
+	"G\<turnstile>Norm s \<midarrow>In2  ({C,stat}e..fn)           \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                 \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l ({statT,mode}e\<cdot>mn({pT}p)) \<succ>\<midarrow>n\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1r (Init C)                  \<succ>\<midarrow>n\<rightarrow> xs'"
+declare split_if     [split] split_if_asm     [split] 
+        option.split [split] option.split_asm [split]
+
+lemma evaln_Inj_elim: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (w,s') \<Longrightarrow> case t of In1 ec \<Rightarrow>  
+  (case ec of Inl e \<Rightarrow> (\<exists>v. w = In1 v) | Inr c \<Rightarrow> w = \<diamondsuit>)  
+  | In2 e \<Rightarrow> (\<exists>v. w = In2 v) | In3 e \<Rightarrow> (\<exists>v. w = In3 v)"
+apply (erule evaln_cases , auto)
+apply (induct_tac "t")
+apply   (induct_tac "a")
+apply auto
+done
+
+ML_setup {*
+fun enf nam inj rhs =
+let
+  val name = "evaln_" ^ nam ^ "_eq"
+  val lhs = "G\<turnstile>s \<midarrow>" ^ inj ^ " t\<succ>\<midarrow>n\<rightarrow> (w, s')"
+  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
+	(K [Auto_tac, ALLGOALS (ftac (thm "evaln_Inj_elim")) THEN Auto_tac])
+  fun is_Inj (Const (inj,_) $ _) = true
+    | is_Inj _                   = false
+  fun pred (_ $ (Const ("Pair",_) $ _ $ (Const ("Pair", _) $ _ $ 
+    (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ x $ _ )))) $ _ ) = is_Inj x
+in
+  make_simproc name lhs pred (thm name)
+end;
+
+val evaln_expr_proc = enf "expr" "In1l" "\<exists>v.  w=In1 v  \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<midarrow>n\<rightarrow> s'";
+val evaln_var_proc  = enf "var"  "In2"  "\<exists>vf. w=In2 vf \<and> G\<turnstile>s \<midarrow>t=\<succ>vf\<midarrow>n\<rightarrow> s'";
+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'";
+val evaln_stmt_proc = enf "stmt" "In1r" "     w=\<diamondsuit>      \<and> G\<turnstile>s \<midarrow>t     \<midarrow>n\<rightarrow> s'";
+Addsimprocs [evaln_expr_proc,evaln_var_proc,evaln_exprs_proc,evaln_stmt_proc];
+
+bind_thms ("evaln_AbruptIs", sum3_instantiate (thm "evaln.Abrupt"))
+*}
+declare evaln_AbruptIs [intro!]
+
+lemma evaln_abrupt_lemma: "G\<turnstile>s \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (v,s') \<Longrightarrow> 
+ fst s = Some xc \<longrightarrow> s' = s \<and> v = arbitrary3 e"
+apply (erule evaln_cases , auto)
+done
+
+lemma evaln_abrupt: 
+ "\<And>s'. G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (w,s') = (s' = (Some xc,s) \<and>  
+  w=arbitrary3 e \<and> G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (arbitrary3 e,(Some xc,s)))"
+apply auto
+apply (frule evaln_abrupt_lemma, auto)+
+done
+
+ML {*
+local
+  fun is_Some (Const ("Pair",_) $ (Const ("Option.option.Some",_) $ _)$ _) =true
+    | is_Some _ = false
+  fun pred (_ $ (Const ("Pair",_) $
+     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
+       (Const ("Pair", _) $ _ $ x)))) $ _ ) = is_Some x
+in
+  val evaln_abrupt_proc = 
+ make_simproc "evaln_abrupt" "G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (w,s')" pred (thm "evaln_abrupt")
+end;
+Addsimprocs [evaln_abrupt_proc]
+*}
+
+lemma evaln_LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else arbitrary)\<midarrow>n\<rightarrow> s"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: evaln.Lit)
+
+lemma CondI: 
+ "\<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> 
+  G\<turnstile>s \<midarrow>e ? e1 : e2-\<succ>(if normal s1 then v else arbitrary)\<midarrow>n\<rightarrow> s2"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: evaln.Cond)
+
+lemma evaln_SkipI [intro!]: "G\<turnstile>s \<midarrow>Skip\<midarrow>n\<rightarrow> s"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: evaln.Skip)
+
+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'"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: evaln.Expr)
+
+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"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: evaln.Comp)
+
+lemma evaln_IfI: 
+ "\<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> 
+  G\<turnstile>s \<midarrow>If(e) c1 Else c2\<midarrow>n\<rightarrow> s2"
+apply (case_tac "s", case_tac "a = None")
+by (auto intro!: evaln.If)
+
+lemma evaln_SkipD [dest!]: "G\<turnstile>s \<midarrow>Skip\<midarrow>n\<rightarrow> s' \<Longrightarrow> s' = s" 
+by (erule evaln_cases, auto)
+
+lemma evaln_Skip_eq [simp]: "G\<turnstile>s \<midarrow>Skip\<midarrow>n\<rightarrow> s' = (s = s')"
+apply auto
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Example.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,1248 @@
+(*  Title:      isabelle/Bali/Example.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* Example Bali program *}
+
+theory Example = Eval + WellForm:
+
+
+text {*
+The following example Bali program includes:
+\begin{itemize}
+\item class and interface declarations with inheritance, hiding of fields,
+    overriding of methods (with refined result type), array type,
+\item method call (with dynamic binding), parameter access, return expressions,
+\item expression statements, sequential composition, literal values, 
+    local assignment, local access, field assignment, type cast,
+\item exception generation and propagation, try and catch statement, throw 
+      statement
+\item instance creation and (default) static initialization
+\end{itemize}
+\begin{verbatim}
+package java_lang
+
+public interface HasFoo {
+  public Base foo(Base z);
+}
+
+public class Base implements HasFoo {
+  static boolean arr[] = new boolean[2];
+  public HasFoo vee;
+  public Base foo(Base z) {
+    return z;
+  }
+}
+
+public class Ext extends Base {
+  public int vee;
+  public Ext foo(Base z) {
+    ((Ext)z).vee = 1;
+    return null;
+  }
+}
+
+public class Example {
+  public static void main(String args[]) throws Throwable {
+    Base e = new Ext();
+    try {e.foo(null); }
+    catch(NullPointerException z) { 
+      while(Ext.arr[2]) ;
+    }
+  }
+}
+\end{verbatim}
+*}
+
+declare widen.null [intro]
+
+lemma wf_fdecl_def2: "\<And>fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))"
+apply (unfold wf_fdecl_def)
+apply (simp (no_asm))
+done
+
+declare wf_fdecl_def2 [iff]
+
+
+section "type and expression names"
+
+(** unfortunately cannot simply instantiate tnam **)
+datatype tnam_  = HasFoo_ | Base_ | Ext_
+datatype vnam_  = arr_ | vee_ | z_ | e_
+datatype label_ = lab1_
+
+consts
+
+  tnam_ :: "tnam_  \<Rightarrow> tnam"
+  vnam_ :: "vnam_  \<Rightarrow> vname"
+  label_:: "label_ \<Rightarrow> label"
+axioms  (** tnam_, vnam_ and label are intended to be isomorphic 
+            to tnam, vname and label **)
+
+  inj_tnam_  [simp]: "(tnam_  x = tnam_  y) = (x = y)"
+  inj_vnam_  [simp]: "(vnam_  x = vnam_  y) = (x = y)"
+  inj_label_ [simp]: "(label_ x = label_ y) = (x = y)"
+  
+  
+  surj_tnam_:  "\<exists>m. n = tnam_  m"
+  surj_vnam_:  "\<exists>m. n = vnam_  m"
+  surj_label_:" \<exists>m. n = label_ m"
+
+syntax
+
+  HasFoo :: qtname
+  Base   :: qtname
+  Ext    :: qtname
+  arr :: ename
+  vee :: ename
+  z   :: ename
+  e   :: ename
+  lab1:: label
+translations
+
+  "HasFoo" == "\<lparr>pid=java_lang,tid=TName (tnam_ HasFoo_)\<rparr>"
+  "Base"   == "\<lparr>pid=java_lang,tid=TName (tnam_ Base_)\<rparr>"
+  "Ext"    == "\<lparr>pid=java_lang,tid=TName (tnam_ Ext_)\<rparr>"
+  "arr"    ==        "(vnam_ arr_)"
+  "vee"    ==        "(vnam_ vee_)"
+  "z"      ==        "(vnam_ z_)"
+  "e"      ==        "(vnam_ e_)"
+  "lab1"   ==        "label_ lab1_"
+
+
+lemma neq_Base_Object [simp]: "Base\<noteq>Object"
+by (simp add: Object_def)
+
+lemma neq_Ext_Object [simp]: "Ext\<noteq>Object"
+by (simp add: Object_def)
+
+lemma neq_Base_SXcpt [simp]: "Base\<noteq>SXcpt xn"
+by (simp add: SXcpt_def)
+
+lemma neq_Ext_SXcpt [simp]: "Ext\<noteq>SXcpt xn"
+by (simp add: SXcpt_def)
+
+section "classes and interfaces"
+
+defs
+
+  Object_mdecls_def: "Object_mdecls \<equiv> []"
+  SXcpt_mdecls_def:  "SXcpt_mdecls  \<equiv> []"
+
+consts
+  
+  foo    :: mname
+
+constdefs
+  
+  foo_sig   :: sig
+ "foo_sig   \<equiv> \<lparr>name=foo,parTs=[Class Base]\<rparr>"
+  
+  foo_mhead :: mhead
+ "foo_mhead \<equiv> \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>"
+
+constdefs
+  
+  Base_foo :: mdecl
+ "Base_foo \<equiv> (foo_sig, \<lparr>access=Public,static=False,pars=[z],resT=Class Base,
+                        mbody=\<lparr>lcls=[],stmt=Return (!!z)\<rparr>\<rparr>)"
+  
+  Ext_foo  :: mdecl
+ "Ext_foo  \<equiv> (foo_sig, 
+              \<lparr>access=Public,static=False,pars=[z],resT=Class Ext,
+	       mbody=\<lparr>lcls=[]
+                     ,stmt=Expr({Ext,False}Cast (Class Ext) (!!z)..vee := 
+       	                                                     Lit (Intg 1))\<rparr>
+	      \<rparr>)"
+
+constdefs
+  
+arr_viewed_from :: "qtname \<Rightarrow> var"
+"arr_viewed_from C \<equiv> {Base,True}StatRef (ClassT C)..arr"
+
+BaseCl :: class
+"BaseCl \<equiv> \<lparr>access=Public,
+           cfields=[(arr, \<lparr>access=Public,static=True ,type=PrimT Boolean.[]\<rparr>),
+	            (vee, \<lparr>access=Public,static=False,type=Iface HasFoo    \<rparr>)],
+           methods=[Base_foo],
+           init=Expr(arr_viewed_from Base :=New (PrimT Boolean)[Lit (Intg 2)]),
+           super=Object,
+           superIfs=[HasFoo]\<rparr>"
+  
+ExtCl  :: class
+"ExtCl  \<equiv> \<lparr>access=Public,
+           cfields=[(vee, \<lparr>access=Public,static=False,type= PrimT Integer\<rparr>)], 
+           methods=[Ext_foo],
+           init=Skip,
+           super=Base,
+           superIfs=[]\<rparr>"
+
+constdefs
+  
+  HasFooInt :: iface
+ "HasFooInt \<equiv> \<lparr>access=Public,imethods=[(foo_sig, foo_mhead)],isuperIfs=[]\<rparr>"
+
+  Ifaces ::"idecl list"
+ "Ifaces \<equiv> [(HasFoo,HasFooInt)]"
+
+ "Classes" ::"cdecl list"
+ "Classes \<equiv> [(Base,BaseCl),(Ext,ExtCl)]@standard_classes"
+
+lemmas table_classes_defs = 
+     Classes_def standard_classes_def ObjectC_def SXcptC_def
+
+lemma table_ifaces [simp]: "table_of Ifaces = empty(HasFoo\<mapsto>HasFooInt)"
+apply (unfold Ifaces_def)
+apply (simp (no_asm))
+done
+
+lemma table_classes_Object [simp]: 
+ "table_of Classes Object = Some \<lparr>access=Public,cfields=[]
+                                 ,methods=Object_mdecls
+                                 ,init=Skip,super=arbitrary,superIfs=[]\<rparr>"
+apply (unfold table_classes_defs)
+apply (simp (no_asm) add:Object_def)
+done
+
+lemma table_classes_SXcpt [simp]: 
+  "table_of Classes (SXcpt xn) 
+    = Some \<lparr>access=Public,cfields=[],methods=SXcpt_mdecls,
+            init=Skip,
+            super=if xn = Throwable then Object else SXcpt Throwable,
+            superIfs=[]\<rparr>"
+apply (unfold table_classes_defs)
+apply (induct_tac xn)
+apply (simp add: Object_def SXcpt_def)+
+done
+
+lemma table_classes_HasFoo [simp]: "table_of Classes HasFoo = None"
+apply (unfold table_classes_defs)
+apply (simp (no_asm) add: Object_def SXcpt_def)
+done
+
+lemma table_classes_Base [simp]: "table_of Classes Base = Some BaseCl"
+apply (unfold table_classes_defs )
+apply (simp (no_asm) add: Object_def SXcpt_def)
+done
+
+lemma table_classes_Ext [simp]: "table_of Classes Ext = Some ExtCl"
+apply (unfold table_classes_defs )
+apply (simp (no_asm) add: Object_def SXcpt_def)
+done
+
+
+section "program"
+
+syntax
+  tprg :: prog
+
+translations
+  "tprg" == "\<lparr>ifaces=Ifaces,classes=Classes\<rparr>"
+
+constdefs
+  test    :: "(ty)list \<Rightarrow> stmt"
+ "test pTs \<equiv> e:==NewC Ext;; 
+           \<spacespace> Try Expr({ClassT Base,IntVir}!!e\<cdot>foo({pTs}[Lit Null]))
+           \<spacespace> Catch((SXcpt NullPointer) z)
+           (lab1\<bullet> While(Acc (Acc (arr_viewed_from Ext).[Lit (Intg 2)])) Skip)"
+
+
+section "well-structuredness"
+
+lemma not_Object_subcls_any [elim!]: "(Object, C) \<in> (subcls1 tprg)^+ \<Longrightarrow> R"
+apply (auto dest!: tranclD subcls1D)
+done
+
+lemma not_Throwable_subcls_SXcpt [elim!]: 
+  "(SXcpt Throwable, SXcpt xn) \<in> (subcls1 tprg)^+ \<Longrightarrow> R"
+apply (auto dest!: tranclD subcls1D) 
+apply (simp add: Object_def SXcpt_def)
+done
+
+lemma not_SXcpt_n_subcls_SXcpt_n [elim!]: 
+  "(SXcpt xn, SXcpt xn)  \<in> (subcls1 tprg)^+ \<Longrightarrow> R"
+apply (auto dest!: tranclD subcls1D)
+apply (drule rtranclD)
+apply auto
+done
+
+lemma not_Base_subcls_Ext [elim!]: "(Base, Ext) \<in> (subcls1 tprg)^+  \<Longrightarrow> R"
+apply (auto dest!: tranclD subcls1D simp add: BaseCl_def)
+done
+
+lemma not_TName_n_subcls_TName_n [rule_format (no_asm), elim!]: 
+  "(\<lparr>pid=java_lang,tid=TName tn\<rparr>, \<lparr>pid=java_lang,tid=TName tn\<rparr>) 
+   \<in> (subcls1 tprg)^+ \<longrightarrow> R"
+apply (rule_tac n1 = "tn" in surj_tnam_ [THEN exE])
+apply (erule ssubst)
+apply (rule tnam_.induct)
+apply  safe
+apply (auto dest!: tranclD subcls1D simp add: BaseCl_def ExtCl_def)
+apply (drule rtranclD)
+apply auto
+done
+
+
+lemma ws_idecl_HasFoo: "ws_idecl tprg HasFoo []"
+apply (unfold ws_idecl_def)
+apply (simp (no_asm))
+done
+
+lemma ws_cdecl_Object: "ws_cdecl tprg Object any"
+apply (unfold ws_cdecl_def)
+apply auto
+done
+
+lemma ws_cdecl_Throwable: "ws_cdecl tprg (SXcpt Throwable) Object"
+apply (unfold ws_cdecl_def)
+apply auto
+done
+
+lemma ws_cdecl_SXcpt: "ws_cdecl tprg (SXcpt xn) (SXcpt Throwable)"
+apply (unfold ws_cdecl_def)
+apply auto
+done
+
+lemma ws_cdecl_Base: "ws_cdecl tprg Base Object"
+apply (unfold ws_cdecl_def)
+apply auto
+done
+
+lemma ws_cdecl_Ext: "ws_cdecl tprg Ext Base"
+apply (unfold ws_cdecl_def)
+apply auto
+done
+
+lemmas ws_cdecls = ws_cdecl_SXcpt ws_cdecl_Object ws_cdecl_Throwable
+                   ws_cdecl_Base ws_cdecl_Ext
+
+declare not_Object_subcls_any [rule del]
+          not_Throwable_subcls_SXcpt [rule del] 
+          not_SXcpt_n_subcls_SXcpt_n [rule del] 
+          not_Base_subcls_Ext [rule del] not_TName_n_subcls_TName_n [rule del]
+
+lemma ws_idecl_all: 
+ "G=tprg \<Longrightarrow> (\<forall>(I,i)\<in>set Ifaces. ws_idecl G I (isuperIfs i))"
+apply (simp (no_asm) add: Ifaces_def HasFooInt_def)
+apply (auto intro!: ws_idecl_HasFoo)
+done
+
+lemma ws_cdecl_all: "G=tprg \<Longrightarrow> (\<forall>(C,c)\<in>set Classes. ws_cdecl G C (super c))"
+apply (simp (no_asm) add: Classes_def BaseCl_def ExtCl_def)
+apply (auto intro!: ws_cdecls simp add: standard_classes_def ObjectC_def 
+        SXcptC_def)
+done
+
+lemma ws_tprg: "ws_prog tprg"
+apply (unfold ws_prog_def)
+apply (auto intro!: ws_idecl_all ws_cdecl_all)
+done
+
+
+section "misc program properties (independent of well-structuredness)"
+
+lemma single_iface [simp]: "is_iface tprg I = (I = HasFoo)"
+apply (unfold Ifaces_def)
+apply (simp (no_asm))
+done
+
+lemma empty_subint1 [simp]: "subint1 tprg = {}"
+apply (unfold subint1_def Ifaces_def HasFooInt_def)
+apply auto
+done
+
+lemma unique_ifaces: "unique Ifaces"
+apply (unfold Ifaces_def)
+apply (simp (no_asm))
+done
+
+lemma unique_classes: "unique Classes"
+apply (unfold table_classes_defs )
+apply (simp )
+done
+
+lemma SXcpt_subcls_Throwable [simp]: "tprg\<turnstile>SXcpt xn\<preceq>\<^sub>C SXcpt Throwable"
+apply (rule SXcpt_subcls_Throwable_lemma)
+apply auto
+done
+
+lemma Ext_subclseq_Base [simp]: "tprg\<turnstile>Ext \<preceq>\<^sub>C Base"
+apply (rule subcls_direct1)
+apply  (simp (no_asm) add: ExtCl_def)
+apply  (simp add: Object_def)
+apply (simp (no_asm))
+done
+
+lemma Ext_subcls_Base [simp]: "tprg\<turnstile>Ext \<prec>\<^sub>C Base"
+apply (rule subcls_direct2)
+apply  (simp (no_asm) add: ExtCl_def)
+apply  (simp add: Object_def)
+apply (simp (no_asm))
+done
+
+
+
+section "fields and method lookup"
+
+lemma fields_tprg_Object [simp]: "DeclConcepts.fields tprg Object = []"
+by (rule ws_tprg [THEN fields_emptyI], force+)
+
+lemma fields_tprg_Throwable [simp]: 
+  "DeclConcepts.fields tprg (SXcpt Throwable) = []"
+by (rule ws_tprg [THEN fields_emptyI], force+)
+
+lemma fields_tprg_SXcpt [simp]: "DeclConcepts.fields tprg (SXcpt xn) = []"
+apply (case_tac "xn = Throwable")
+apply  (simp (no_asm_simp))
+by (rule ws_tprg [THEN fields_emptyI], force+)
+
+lemmas fields_rec_ = fields_rec [OF _ ws_tprg]
+
+lemma fields_Base [simp]: 
+"DeclConcepts.fields tprg Base 
+  = [((arr,Base), \<lparr>access=Public,static=True ,type=PrimT Boolean.[]\<rparr>),
+     ((vee,Base), \<lparr>access=Public,static=False,type=Iface HasFoo    \<rparr>)]"
+apply (subst fields_rec_)
+apply   (auto simp add: BaseCl_def)
+done
+
+lemma fields_Ext [simp]: 
+ "DeclConcepts.fields tprg Ext  
+  = [((vee,Ext), \<lparr>access=Public,static=False,type= PrimT Integer\<rparr>)] 
+    @ DeclConcepts.fields tprg Base"
+apply (rule trans)
+apply (rule fields_rec_)
+apply   (auto simp add: ExtCl_def Object_def)
+done
+
+lemmas imethds_rec_ = imethds_rec [OF _ ws_tprg]
+lemmas methd_rec_  = methd_rec  [OF _ ws_tprg]
+
+lemma imethds_HasFoo [simp]: 
+  "imethds tprg HasFoo = o2s \<circ> empty(foo_sig\<mapsto>(HasFoo, foo_mhead))"
+apply (rule trans)
+apply (rule imethds_rec_)
+apply  (auto simp add: HasFooInt_def)
+done
+
+lemma methd_tprg_Object [simp]: "methd tprg Object = empty"
+apply (subst methd_rec_)
+apply (auto simp add: Object_mdecls_def)
+done
+
+lemma methd_Base [simp]: 
+  "methd tprg Base = table_of [(\<lambda>(s,m). (s, Base, m)) Base_foo]"
+apply (rule trans)
+apply (rule methd_rec_)
+apply   (auto simp add: BaseCl_def)
+done
+
+(* ### To Table *)
+lemma filter_tab_all_False: 
+ "\<forall> k y. t k = Some y \<longrightarrow> \<not> p k y \<Longrightarrow>filter_tab p t = empty"
+by (auto simp add: filter_tab_def expand_fun_eq)
+
+
+lemma memberid_Base_foo_simp [simp]:
+ "memberid (mdecl Base_foo) = mid foo_sig"
+by (simp add: Base_foo_def)
+
+lemma memberid_Ext_foo_simp [simp]:
+ "memberid (mdecl Ext_foo) = mid foo_sig"
+by (simp add: Ext_foo_def)
+
+lemma Base_declares_foo:
+  "tprg\<turnstile>mdecl Base_foo declared_in Base"
+by (auto simp add: declared_in_def cdeclaredmethd_def BaseCl_def Base_foo_def)
+
+lemma foo_sig_not_undeclared_in_Base:
+  "\<not> tprg\<turnstile>mid foo_sig undeclared_in Base"
+proof -
+  from Base_declares_foo
+  show ?thesis
+    by (auto dest!: declared_not_undeclared )
+qed
+
+lemma Ext_declares_foo:
+  "tprg\<turnstile>mdecl Ext_foo declared_in Ext"
+by (auto simp add: declared_in_def cdeclaredmethd_def ExtCl_def Ext_foo_def)
+
+lemma foo_sig_not_undeclared_in_Ext:
+  "\<not> tprg\<turnstile>mid foo_sig undeclared_in Ext"
+proof -
+  from Ext_declares_foo
+  show ?thesis
+    by (auto dest!: declared_not_undeclared )
+qed
+
+lemma Base_foo_not_inherited_in_Ext:
+ "\<not> tprg \<turnstile> Ext inherits (Base,mdecl Base_foo)"
+by (auto simp add: inherits_def foo_sig_not_undeclared_in_Ext)
+
+lemma Ext_method_inheritance:
+ "filter_tab (\<lambda>sig m. tprg \<turnstile> Ext inherits method sig m)
+     (empty(fst ((\<lambda>(s, m). (s, Base, m)) Base_foo)\<mapsto>
+      snd ((\<lambda>(s, m). (s, Base, m)) Base_foo)))
+  = empty"
+proof -
+  from Base_foo_not_inherited_in_Ext
+  show ?thesis
+    by (auto intro: filter_tab_all_False simp add: Base_foo_def)
+qed
+
+
+lemma methd_Ext [simp]: "methd tprg Ext =   
+  table_of [(\<lambda>(s,m). (s, Ext, m)) Ext_foo]"
+apply (rule trans)
+apply (rule methd_rec_)
+apply   (auto simp add: ExtCl_def Object_def Ext_method_inheritance)
+done
+
+section "accessibility"
+
+lemma classesDefined: 
+ "\<lbrakk>class tprg C = Some c; C\<noteq>Object\<rbrakk> \<Longrightarrow> \<exists> sc. class tprg (super c) = Some sc"
+apply (auto simp add: Classes_def standard_classes_def 
+                      BaseCl_def ExtCl_def
+                      SXcptC_def ObjectC_def) 
+done
+
+lemma superclassesBase [simp]: "superclasses tprg Base={Object}"
+proof -
+  have ws: "ws_prog tprg" by (rule ws_tprg)
+  then show ?thesis
+    by (auto simp add: superclasses_rec  BaseCl_def)
+qed
+
+lemma superclassesExt [simp]: "superclasses tprg Ext={Base,Object}"
+proof -
+  have ws: "ws_prog tprg" by (rule ws_tprg)
+  then show ?thesis
+    by (auto simp add: superclasses_rec  ExtCl_def BaseCl_def)
+qed
+
+lemma HasFoo_accessible[simp]:"tprg\<turnstile>(Iface HasFoo) accessible_in P" 
+by (simp add: accessible_in_RefT_simp is_public_def HasFooInt_def)
+
+lemma HasFoo_is_acc_iface[simp]: "is_acc_iface tprg P HasFoo"
+by (simp add: is_acc_iface_def)
+
+lemma HasFoo_is_acc_type[simp]: "is_acc_type tprg P (Iface HasFoo)"
+by (simp add: is_acc_type_def)
+
+lemma Base_accessible[simp]:"tprg\<turnstile>(Class Base) accessible_in P" 
+by (simp add: accessible_in_RefT_simp is_public_def BaseCl_def)
+
+lemma Base_is_acc_class[simp]: "is_acc_class tprg P Base"
+by (simp add: is_acc_class_def)
+
+lemma Base_is_acc_type[simp]: "is_acc_type tprg P (Class Base)"
+by (simp add: is_acc_type_def)
+
+lemma Ext_accessible[simp]:"tprg\<turnstile>(Class Ext) accessible_in P" 
+by (simp add: accessible_in_RefT_simp is_public_def ExtCl_def)
+
+lemma Ext_is_acc_class[simp]: "is_acc_class tprg P Ext"
+by (simp add: is_acc_class_def)
+
+lemma Ext_is_acc_type[simp]: "is_acc_type tprg P (Class Ext)"
+by (simp add: is_acc_type_def)
+
+lemma accmethd_tprg_Object [simp]: "accmethd tprg S Object = empty"
+apply (unfold accmethd_def)
+apply (simp)
+done
+
+lemma  snd_special_simp: "snd ((\<lambda>(s, m). (s, a, m)) x) = (a,snd x)"
+by (cases x) (auto)
+
+lemma  fst_special_simp: "fst ((\<lambda>(s, m). (s, a, m)) x) = fst x"
+by (cases x) (auto)
+
+lemma foo_sig_undeclared_in_Object:
+  "tprg\<turnstile>mid foo_sig undeclared_in Object"
+by (auto simp add: undeclared_in_def cdeclaredmethd_def Object_mdecls_def)
+
+(* ### To DeclConcepts *)
+lemma undeclared_not_declared:
+ "G\<turnstile> memberid m undeclared_in C \<Longrightarrow> \<not> G\<turnstile> m declared_in C" 
+by (cases m) (auto simp add: declared_in_def undeclared_in_def)
+
+
+lemma unique_sig_Base_foo:
+ "tprg\<turnstile> mdecl (sig, snd Base_foo) declared_in Base \<Longrightarrow> sig=foo_sig"
+by (auto simp add: declared_in_def cdeclaredmethd_def 
+                   Base_foo_def BaseCl_def)
+
+lemma Base_foo_no_override:
+ "tprg,sig\<turnstile>(Base,(snd Base_foo)) overrides old \<Longrightarrow> P"
+apply (drule overrides_commonD)
+apply (clarsimp)
+apply (frule subclsEval)
+apply    (rule ws_tprg)
+apply    (simp)
+apply    (rule classesDefined) 
+apply    assumption+
+apply (frule unique_sig_Base_foo) 
+apply (auto dest!: declared_not_undeclared intro: foo_sig_undeclared_in_Object
+            dest: unique_sig_Base_foo)
+done
+
+lemma Base_foo_no_stat_override:
+ "tprg,sig\<turnstile>(Base,(snd Base_foo)) overrides\<^sub>S old \<Longrightarrow> P"
+apply (drule stat_overrides_commonD)
+apply (clarsimp)
+apply (frule subclsEval)
+apply    (rule ws_tprg)
+apply    (simp)
+apply    (rule classesDefined) 
+apply    assumption+
+apply (frule unique_sig_Base_foo) 
+apply (auto dest!: declared_not_undeclared intro: foo_sig_undeclared_in_Object
+            dest: unique_sig_Base_foo)
+done
+
+
+lemma Base_foo_no_hide:
+ "tprg,sig\<turnstile>(Base,(snd Base_foo)) hides old \<Longrightarrow> P"
+by (auto dest: hidesD simp add: Base_foo_def member_is_static_simp)
+
+lemma Ext_foo_no_hide:
+ "tprg,sig\<turnstile>(Ext,(snd Ext_foo)) hides old \<Longrightarrow> P"
+by (auto dest: hidesD simp add: Ext_foo_def member_is_static_simp)
+
+lemma unique_sig_Ext_foo:
+ "tprg\<turnstile> mdecl (sig, snd Ext_foo) declared_in Ext \<Longrightarrow> sig=foo_sig"
+by (auto simp add: declared_in_def cdeclaredmethd_def 
+                   Ext_foo_def ExtCl_def)
+
+(* ### To DeclConcepts *)
+lemma unique_declaration: 
+ "\<lbrakk>G\<turnstile>m declared_in C;  G\<turnstile>n declared_in C; memberid m = memberid n \<rbrakk> 
+  \<Longrightarrow> m = n"
+apply (cases m)
+apply (cases n,
+        auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)+
+done
+
+
+lemma Ext_foo_override:
+ "tprg,sig\<turnstile>(Ext,(snd Ext_foo)) overrides old 
+  \<Longrightarrow> old = (Base,(snd Base_foo))"
+apply (drule overrides_commonD)
+apply (clarsimp)
+apply (frule subclsEval)
+apply    (rule ws_tprg)
+apply    (simp)
+apply    (rule classesDefined) 
+apply    assumption+
+apply (frule unique_sig_Ext_foo) 
+apply (case_tac "old")
+apply (insert Base_declares_foo foo_sig_undeclared_in_Object) 
+apply (auto simp add: ExtCl_def Ext_foo_def
+                      BaseCl_def Base_foo_def Object_mdecls_def
+                      split_paired_all
+                      member_is_static_simp
+           dest: declared_not_undeclared unique_declaration)
+done
+
+lemma Ext_foo_stat_override:
+ "tprg,sig\<turnstile>(Ext,(snd Ext_foo)) overrides\<^sub>S old 
+  \<Longrightarrow> old = (Base,(snd Base_foo))"
+apply (drule stat_overrides_commonD)
+apply (clarsimp)
+apply (frule subclsEval)
+apply    (rule ws_tprg)
+apply    (simp)
+apply    (rule classesDefined) 
+apply    assumption+
+apply (frule unique_sig_Ext_foo) 
+apply (case_tac "old")
+apply (insert Base_declares_foo foo_sig_undeclared_in_Object) 
+apply (auto simp add: ExtCl_def Ext_foo_def
+                      BaseCl_def Base_foo_def Object_mdecls_def
+                      split_paired_all
+                      member_is_static_simp
+           dest: declared_not_undeclared unique_declaration)
+done
+
+(*### weiter hoch *)
+lemma Base_foo_member_of_Base: 
+  "tprg\<turnstile>(Base,mdecl Base_foo) member_of Base"
+by (auto intro!: members.Immediate Base_declares_foo)
+
+(*### weiter hoch *)
+lemma Ext_foo_member_of_Ext: 
+  "tprg\<turnstile>(Ext,mdecl Ext_foo) member_of Ext"
+by (auto intro!: members.Immediate Ext_declares_foo)
+
+lemma Base_foo_permits_acc:
+ "tprg \<turnstile> (Base, mdecl Base_foo) in Base permits_acc_to S"
+by ( simp add: permits_acc_def Base_foo_def)
+
+lemma Base_foo_accessible [simp]:
+ "tprg\<turnstile>(Base,mdecl Base_foo) of Base accessible_from S"
+by (auto intro: accessible_fromR.immediate 
+                Base_foo_member_of_Base Base_foo_permits_acc)
+
+lemma accmethd_Base [simp]: 
+  "accmethd tprg S Base = methd tprg Base"
+apply (simp add: accmethd_def)
+apply (rule filter_tab_all_True)
+apply (simp add: snd_special_simp fst_special_simp)
+done
+
+lemma Ext_foo_permits_acc:
+ "tprg \<turnstile> (Ext, mdecl Ext_foo) in Ext permits_acc_to S"
+by ( simp add: permits_acc_def Ext_foo_def)
+
+lemma Ext_foo_accessible [simp]:
+ "tprg\<turnstile>(Ext,mdecl Ext_foo) of Ext accessible_from S"
+by (auto intro: accessible_fromR.immediate 
+                Ext_foo_member_of_Ext Ext_foo_permits_acc)
+
+(*
+lemma Base_foo_accessible_through_inheritance_in_Ext [simp]:
+ "tprg\<turnstile>(Base,snd Base_foo) accessible_through_inheritance_in Ext"
+apply (rule accessible_through_inheritance.Direct)
+apply   simp
+apply   (simp add: accessible_for_inheritance_in_def Base_foo_def)
+done
+*)
+
+lemma Ext_foo_overrides_Base_foo:
+ "tprg\<turnstile>(Ext,Ext_foo) overrides (Base,Base_foo)"
+proof (rule overridesR.Direct, simp_all)
+  show "\<not> is_static Ext_foo" 
+    by (simp add: member_is_static_simp Ext_foo_def)
+  show "\<not> is_static Base_foo"
+    by (simp add: member_is_static_simp Base_foo_def)
+  show "accmodi Ext_foo \<noteq> Private"
+    by (simp add: Ext_foo_def)
+  show "msig (Ext, Ext_foo) = msig (Base, Base_foo)"
+    by (simp add: Ext_foo_def Base_foo_def)
+  show "tprg\<turnstile>mdecl Ext_foo declared_in Ext"
+    by (auto intro: Ext_declares_foo)
+  show "tprg\<turnstile>mdecl Base_foo declared_in Base"
+    by (auto intro: Base_declares_foo)
+  show "tprg \<turnstile>(Base, mdecl Base_foo) inheritable_in java_lang"
+    by (simp add: inheritable_in_def Base_foo_def)
+  show "tprg\<turnstile>resTy Ext_foo\<preceq>resTy Base_foo"
+    by (simp add: Ext_foo_def Base_foo_def mhead_resTy_simp)
+qed
+
+(*
+lemma Base_foo_of_Ext_accessible[simp]:
+ "tprg\<turnstile>(Base, mdecl Base_foo) of Ext accessible_from S"
+apply (auto intro: accessible_fromR.immediate 
+                Base_foo_member_of_Base Base_foo_permits_acc)
+apply (rule accessible_fromR.immediate)
+apply (rule_tac "old"="(Base,Base_foo)" and  sup="Base" 
+       in accessible_fromR.overriding)
+apply (auto intro!: Ext_foo_overrides_Base_foo)
+apply (auto 
+apply (insert Ext_foo_overrides_Base_foo)
+apply (rule accessible_fromR.overriding, simp_all)
+apply (auto intro!: Ext_foo_overrides_Base_foo)
+apply (auto intro!: accessible_fromR.overriding
+             intro:   Ext_foo_overrides_Base_foo)
+by
+                Ext_foo_member_of_Ext Ext_foo_permits_acc)
+apply (auto intro!: accessible 
+apply (auto simp add: method_accessible_from_def accessible_from_def) 
+apply (simp add: Base_foo_def)
+done 
+*)
+
+lemma accmethd_Ext [simp]: 
+  "accmethd tprg S Ext = methd tprg Ext"
+apply (simp add: accmethd_def)
+apply (rule filter_tab_all_True)
+apply (auto simp add: snd_special_simp fst_special_simp)
+done
+
+(* ### Weiter hoch *)
+lemma cls_Ext: "class tprg Ext = Some ExtCl"
+by simp
+lemma dynmethd_Ext_foo:
+ "dynmethd tprg Base Ext \<lparr>name = foo, parTs = [Class Base]\<rparr> 
+  = Some (Ext,snd Ext_foo)"
+proof -
+  have "methd tprg Base \<lparr>name = foo, parTs = [Class Base]\<rparr> 
+          = Some (Base,snd Base_foo)" and
+       "methd tprg Ext \<lparr>name = foo, parTs = [Class Base]\<rparr> 
+          = Some (Ext,snd Ext_foo)" 
+    by (auto simp add: Ext_foo_def Base_foo_def foo_sig_def)
+  with cls_Ext ws_tprg Ext_foo_overrides_Base_foo
+  show ?thesis
+    by (auto simp add: dynmethd_rec simp add: Ext_foo_def Base_foo_def)
+qed
+
+lemma Base_fields_accessible[simp]:
+ "accfield tprg S Base 
+  = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (DeclConcepts.fields tprg Base))"
+apply (auto simp add: accfield_def expand_fun_eq Let_def 
+                      accessible_in_RefT_simp
+                      is_public_def
+                      BaseCl_def
+                      permits_acc_def
+                      declared_in_def 
+                      cdeclaredfield_def
+               intro!: filter_tab_all_True_Some filter_tab_None
+                       accessible_fromR.immediate
+               intro: members.Immediate)
+done
+
+
+lemma arr_member_of_Base:
+  "tprg\<turnstile>(Base, fdecl (arr, 
+                 \<lparr>access = Public, static = True, type = PrimT Boolean.[]\<rparr>))
+          member_of Base"
+by (auto intro: members.Immediate 
+       simp add: declared_in_def cdeclaredfield_def BaseCl_def)
+ 
+lemma arr_member_of_Ext:
+  "tprg\<turnstile>(Base, fdecl (arr, 
+                    \<lparr>access = Public, static = True, type = PrimT Boolean.[]\<rparr>))
+             member_of Ext"
+apply (rule members.Inherited)
+apply   (simp add: inheritable_in_def)
+apply   (simp add: undeclared_in_def cdeclaredfield_def ExtCl_def)
+apply   (auto intro: arr_member_of_Base simp add: subcls1_def ExtCl_def)
+done
+
+lemma Ext_fields_accessible[simp]:
+"accfield tprg S Ext 
+  = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (DeclConcepts.fields tprg Ext))"
+apply (auto simp add: accfield_def expand_fun_eq Let_def 
+                      accessible_in_RefT_simp
+                      is_public_def
+                      BaseCl_def
+                      ExtCl_def
+                      permits_acc_def
+               intro!: filter_tab_all_True_Some filter_tab_None
+                       accessible_fromR.immediate)
+apply (auto intro: members.Immediate arr_member_of_Ext
+            simp add: declared_in_def cdeclaredfield_def ExtCl_def)
+done
+
+lemma array_of_PrimT_acc [simp]:
+ "is_acc_type tprg java_lang (PrimT t.[])"
+apply (simp add: is_acc_type_def accessible_in_RefT_simp)
+done
+
+lemma PrimT_acc [simp]: 
+ "is_acc_type tprg java_lang (PrimT t)"
+apply (simp add: is_acc_type_def accessible_in_RefT_simp)
+done
+
+lemma Object_acc [simp]:
+ "is_acc_class tprg java_lang Object"
+apply (auto simp add: is_acc_class_def accessible_in_RefT_simp is_public_def)
+done
+
+
+section "well-formedness"
+
+
+lemma wf_HasFoo: "wf_idecl tprg (HasFoo, HasFooInt)"
+apply (unfold wf_idecl_def HasFooInt_def)
+apply (auto intro!: wf_mheadI ws_idecl_HasFoo 
+            simp add: foo_sig_def foo_mhead_def mhead_resTy_simp
+                      member_is_static_simp )
+done
+
+declare wt.Skip [rule del] wt.Init [rule del]
+lemmas Base_foo_defs = Base_foo_def foo_sig_def foo_mhead_def
+lemmas Ext_foo_defs  = Ext_foo_def  foo_sig_def
+ML {* bind_thms ("wt_intros",map (rewrite_rule [id_def]) (thms "wt.intros")) *}
+lemmas wtIs = wt_Call wt_Super wt_FVar wt_StatRef wt_intros
+
+lemma wf_Base_foo: "wf_mdecl tprg Base Base_foo"
+apply (unfold Base_foo_defs )
+apply (auto intro!: wf_mdeclI wf_mheadI intro!: wtIs
+            simp add: mhead_resTy_simp)
+done
+
+lemma wf_Ext_foo: "wf_mdecl tprg Ext Ext_foo"
+apply (unfold Ext_foo_defs )
+apply (auto intro!: wf_mdeclI wf_mheadI intro!: wtIs 
+            simp add: mhead_resTy_simp )
+apply  (rule wt.Cast)
+prefer 2
+apply    simp
+apply   (rule_tac [2] narrow.subcls [THEN cast.narrow])
+apply   (auto intro!: wtIs)
+done
+
+declare mhead_resTy_simp [simp add]
+declare member_is_static_simp [simp add]
+
+lemma wf_BaseC: "wf_cdecl tprg (Base,BaseCl)"
+apply (unfold wf_cdecl_def BaseCl_def arr_viewed_from_def)
+apply (auto intro!: wf_Base_foo)
+apply       (auto intro!: ws_cdecl_Base simp add: Base_foo_def foo_mhead_def)
+apply (auto intro!: wtIs)
+apply (auto simp add: Base_foo_defs entails_def Let_def)
+apply   (insert Base_foo_no_stat_override, simp add: Base_foo_def,blast)+
+apply   (insert Base_foo_no_hide         , simp add: Base_foo_def,blast)
+done
+
+
+lemma wf_ExtC: "wf_cdecl tprg (Ext,ExtCl)"
+apply (unfold wf_cdecl_def ExtCl_def)
+apply (auto intro!: wf_Ext_foo ws_cdecl_Ext)
+apply (auto simp add: entails_def snd_special_simp)
+apply (insert Ext_foo_stat_override)
+apply  (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
+apply  (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
+apply  (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
+apply (insert Ext_foo_no_hide)
+apply  (simp_all add: qmdecl_def)
+apply  blast+
+done
+
+lemma wf_idecl_all: "p=tprg \<Longrightarrow> Ball (set Ifaces) (wf_idecl p)"
+apply (simp (no_asm) add: Ifaces_def)
+apply (simp (no_asm_simp))
+apply (rule wf_HasFoo)
+done
+
+lemma wf_cdecl_all_standard_classes: 
+  "Ball (set standard_classes) (wf_cdecl tprg)"
+apply (unfold standard_classes_def Let_def 
+       ObjectC_def SXcptC_def Object_mdecls_def SXcpt_mdecls_def)
+apply (simp (no_asm) add: wf_cdecl_def ws_cdecls)
+apply (auto simp add:is_acc_class_def accessible_in_RefT_simp SXcpt_def)
+apply (auto simp add: Object_def Classes_def standard_classes_def 
+                      SXcptC_def SXcpt_def)
+done
+
+lemma wf_cdecl_all: "p=tprg \<Longrightarrow> Ball (set Classes) (wf_cdecl p)"
+apply (simp (no_asm) add: Classes_def)
+apply (simp (no_asm_simp))
+apply   (rule wf_BaseC [THEN conjI])
+apply  (rule wf_ExtC [THEN conjI])
+apply (rule wf_cdecl_all_standard_classes)
+done
+
+theorem wf_tprg: "wf_prog tprg"
+apply (unfold wf_prog_def Let_def)
+apply (simp (no_asm) add: unique_ifaces unique_classes)
+apply (rule conjI)
+apply  ((simp (no_asm) add: Classes_def standard_classes_def))
+apply (rule conjI)
+apply (simp add: Object_mdecls_def)
+apply safe
+apply  (cut_tac xn_cases_old)   (* FIXME (insert xn_cases) *)
+apply  (simp (no_asm_simp) add: Classes_def standard_classes_def)
+apply  (insert wf_idecl_all)
+apply  (insert wf_cdecl_all)
+apply  auto
+done
+
+
+section "max spec"
+
+lemma appl_methds_Base_foo: 
+"appl_methds tprg S (ClassT Base) \<lparr>name=foo, parTs=[NT]\<rparr> =  
+  {((ClassT Base, \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>)
+    ,[Class Base])}"
+apply (unfold appl_methds_def)
+apply (simp (no_asm))
+apply (subgoal_tac "tprg\<turnstile>NT\<preceq> Class Base")
+apply  (auto simp add: cmheads_def Base_foo_defs)
+done
+
+lemma max_spec_Base_foo: "max_spec tprg S (ClassT Base) \<lparr>name=foo,parTs=[NT]\<rparr> = 
+  {((ClassT Base, \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>)
+   , [Class Base])}"
+apply (unfold max_spec_def)
+apply (simp (no_asm) add: appl_methds_Base_foo)
+apply auto
+done
+
+
+section "well-typedness"
+
+lemma wt_test: "\<lparr>prg=tprg,cls=S,lcl=empty(VName e\<mapsto>Class Base)\<rparr>\<turnstile>test ?pTs\<Colon>\<surd>"
+apply (unfold test_def arr_viewed_from_def)
+(* ?pTs = [Class Base] *)
+apply (rule wtIs (* ;; *))
+apply  (rule wtIs (* Ass *))
+apply  (rule wtIs (* NewC *))
+apply     (rule wtIs (* LVar *))
+apply      (simp)
+apply     (simp)
+apply    (simp)
+apply   (rule wtIs (* NewC *))
+apply   (simp)
+apply  (simp)
+apply (rule wtIs (* Try *))
+prefer 4
+apply    (simp)
+defer
+apply    (rule wtIs (* Expr *))
+apply    (rule wtIs (* Call *))
+apply       (rule wtIs (* Acc *))
+apply       (rule wtIs (* LVar *))
+apply        (simp)
+apply       (simp)
+apply      (rule wtIs (* Cons *))
+apply       (rule wtIs (* Lit *))
+apply       (simp)
+apply      (rule wtIs (* Nil *))
+apply     (simp)
+apply     (rule max_spec_Base_foo)
+apply    (simp)
+apply   (simp)
+apply  (simp)
+apply  (simp)
+apply (rule wtIs (* While *))
+apply  (rule wtIs (* Acc *))
+apply   (rule wtIs (* AVar *))
+apply   (rule wtIs (* Acc *))
+apply   (rule wtIs (* FVar *))
+apply    (rule wtIs (* StatRef *))
+apply    (simp)
+apply   (simp)
+apply   (simp )
+apply   (simp)
+apply  (rule wtIs (* LVar *))
+apply  (simp)
+apply (rule wtIs (* Skip *))
+done
+
+
+section "execution"
+
+lemma alloc_one: "\<And>a obj. \<lbrakk>the (new_Addr h) = a; atleast_free h (Suc n)\<rbrakk> \<Longrightarrow>  
+  new_Addr h = Some a \<and> atleast_free (h(a\<mapsto>obj)) n"
+apply (frule atleast_free_SucD)
+apply (drule atleast_free_Suc [THEN iffD1])
+apply clarsimp
+apply (frule new_Addr_SomeI)
+apply force
+done
+
+declare fvar_def2 [simp] avar_def2 [simp] init_lvars_def2 [simp]
+declare init_obj_def [simp] var_tys_def [simp] fields_table_def [simp]
+declare BaseCl_def [simp] ExtCl_def [simp] Ext_foo_def [simp]
+        Base_foo_defs  [simp]
+
+ML {* bind_thms ("eval_intros", map 
+        (simplify (simpset() delsimps [thm "Skip_eq"]
+                             addsimps [thm "lvar_def"]) o 
+         rewrite_rule [thm "assign_def",Let_def]) (thms "eval.intros")) *}
+lemmas eval_Is = eval_Init eval_StatRef AbruptIs eval_intros
+
+consts
+  a :: loc
+  b :: loc
+  c :: loc
+  
+syntax
+
+  tprg :: prog
+
+  obj_a :: obj
+  obj_b :: obj
+  obj_c :: obj
+  arr_N :: "(vn, val) table"
+  arr_a :: "(vn, val) table"
+  globs1 :: globs
+  globs2 :: globs
+  globs3 :: globs
+  globs8 :: globs
+  locs3 :: locals
+  locs4 :: locals
+  locs8 :: locals
+  s0  :: state
+  s0' :: state
+  s9' :: state
+  s1  :: state
+  s1' :: state
+  s2  :: state
+  s2' :: state
+  s3  :: state
+  s3' :: state
+  s4  :: state
+  s4' :: state
+  s6' :: state
+  s7' :: state
+  s8  :: state
+  s8' :: state
+
+translations
+
+  "tprg"    == "\<lparr>ifaces=Ifaces,classes=Classes\<rparr>"
+  
+  "obj_a"   <= "\<lparr>tag=Arr (PrimT Boolean) two
+                ,values=empty(Inr 0\<mapsto>Bool False)(Inr one\<mapsto>Bool False)\<rparr>"
+  "obj_b"   <= "\<lparr>tag=CInst Ext
+                ,values=(empty(Inl (vee, Base)\<mapsto>Null   ) 
+			      (Inl (vee, Ext )\<mapsto>Intg 0))\<rparr>"
+  "obj_c"   == "\<lparr>tag=CInst (SXcpt NullPointer),values=empty\<rparr>"
+  "arr_N"   == "empty(Inl (arr, Base)\<mapsto>Null)"
+  "arr_a"   == "empty(Inl (arr, Base)\<mapsto>Addr a)"
+  "globs1"  == "empty(Inr Ext   \<mapsto>\<lparr>tag=arbitrary, values=empty\<rparr>)
+		     (Inr Base  \<mapsto>\<lparr>tag=arbitrary, values=arr_N\<rparr>)
+		     (Inr Object\<mapsto>\<lparr>tag=arbitrary, values=empty\<rparr>)"
+  "globs2"  == "empty(Inr Ext   \<mapsto>\<lparr>tag=arbitrary, values=empty\<rparr>)
+		     (Inr Object\<mapsto>\<lparr>tag=arbitrary, values=empty\<rparr>)
+		     (Inl a\<mapsto>obj_a)
+		     (Inr Base  \<mapsto>\<lparr>tag=arbitrary, values=arr_a\<rparr>)"
+  "globs3"  == "globs2(Inl b\<mapsto>obj_b)"
+  "globs8"  == "globs3(Inl c\<mapsto>obj_c)"
+  "locs3"   == "empty(VName e\<mapsto>Addr b)"
+  "locs4"   == "empty(VName z\<mapsto>Null)(Inr()\<mapsto>Addr b)"
+  "locs8"   == "locs3(VName z\<mapsto>Addr c)"
+  "s0"  == "       st empty  empty"
+  "s0'" == " Norm  s0"
+  "s1"  == "       st globs1 empty"
+  "s1'" == " Norm  s1"
+  "s2"  == "       st globs2 empty"
+  "s2'" == " Norm  s2"
+  "s3"  == "       st globs3 locs3 "
+  "s3'" == " Norm  s3"
+  "s4"  == "       st globs3 locs4"
+  "s4'" == " Norm  s4"
+  "s6'" == "(Some (Xcpt (Std NullPointer)), s4)"
+  "s7'" == "(Some (Xcpt (Std NullPointer)), s3)"
+  "s8"  == "       st globs8 locs8"
+  "s8'" == " Norm  s8"
+  "s9'" == "(Some (Xcpt (Std IndOutBound)), s8)"
+
+
+syntax "four"::nat
+       "tree"::nat
+       "two" ::nat
+       "one" ::nat
+translations 
+  "one"  == "Suc 0"
+  "two"  == "Suc one"
+  "tree" == "Suc two"
+  "four" == "Suc tree"
+
+declare Pair_eq [simp del]
+lemma exec_test: 
+"\<lbrakk>the (new_Addr (heap  s1)) = a;  
+  the (new_Addr (heap ?s2)) = b;  
+  the (new_Addr (heap ?s3)) = c\<rbrakk> \<Longrightarrow>  
+  atleast_free  (heap s0) four \<Longrightarrow>  
+  tprg\<turnstile>s0' \<midarrow>test [Class Base]\<rightarrow> ?s9'"
+apply (unfold test_def arr_viewed_from_def)
+(* ?s9' = s9' *)
+apply (simp (no_asm_use))
+apply (drule (1) alloc_one,clarsimp)
+apply (rule eval_Is (* ;; *))
+apply  (erule_tac V = "the (new_Addr ?h) = c" in thin_rl)
+apply  (erule_tac [2] V = "new_Addr ?h = Some a" in thin_rl)
+apply  (erule_tac [2] V = "atleast_free ?h four" in thin_rl)
+apply  (rule eval_Is (* Expr *))
+apply  (rule eval_Is (* Ass *))
+apply   (rule eval_Is (* LVar *))
+apply  (rule eval_Is (* NewC *))
+      (* begin init Ext *)
+apply   (erule_tac V = "the (new_Addr ?h) = b" in thin_rl)
+apply   (erule_tac V = "atleast_free ?h tree" in thin_rl)
+apply   (erule_tac [2] V = "atleast_free ?h four" in thin_rl)
+apply   (erule_tac [2] V = "new_Addr ?h = Some a" in thin_rl)
+apply   (rule eval_Is (* Init Ext *))
+apply   (simp)
+apply   (rule conjI)
+prefer 2 apply (rule conjI HOL.refl)+
+apply   (rule eval_Is (* Init Base *))
+apply   (simp add: arr_viewed_from_def)
+apply   (rule conjI)
+apply    (rule eval_Is (* Init Object *))
+apply    (simp)
+apply    (rule conjI, rule HOL.refl)+
+apply    (rule HOL.refl)
+apply   (simp)
+apply   (rule conjI, rule_tac [2] HOL.refl)
+apply   (rule eval_Is (* Expr *))
+apply   (rule eval_Is (* Ass *))
+apply    (rule eval_Is (* FVar *))
+apply      (rule init_done, simp)
+apply     (rule eval_Is (* StatRef *))
+apply    (simp)
+apply   (rule eval_Is (* NewA *))
+apply     (simp)
+apply    (rule eval_Is (* Lit *))
+apply   (simp)
+apply   (rule halloc.New)
+apply    (simp (no_asm_simp))
+apply   (drule atleast_free_weaken,rotate_tac -1,drule atleast_free_weaken)
+apply   (simp (no_asm_simp))
+apply  (simp add: upd_gobj_def)
+      (* end init Ext *)
+apply  (rule halloc.New)
+apply   (drule alloc_one)
+prefer 2 apply fast
+apply   (simp (no_asm_simp))
+apply  (drule atleast_free_weaken)
+apply  force
+apply (simp)
+apply (drule alloc_one)
+apply  (simp (no_asm_simp))
+apply clarsimp
+apply (erule_tac V = "atleast_free ?h tree" in thin_rl)
+apply (drule_tac x = "a" in new_AddrD2 [THEN spec])
+apply (simp (no_asm_use))
+apply (rule eval_Is (* Try *))
+apply   (rule eval_Is (* Expr *))
+      (* begin method call *)
+apply   (rule eval_Is (* Call *))
+apply      (rule eval_Is (* Acc *))
+apply      (rule eval_Is (* LVar *))
+apply     (rule eval_Is (* Cons *))
+apply      (rule eval_Is (* Lit *))
+apply     (rule eval_Is (* Nil *))
+apply    (simp)
+apply   (simp)
+apply   (rule eval_Is (* Methd *))
+apply   (simp add: body_def Let_def)
+apply   (rule eval_Is (* Body *))
+apply     (rule init_done, simp)
+apply       (simp add: invocation_declclass_def dynlookup_def dynmethd_Ext_foo)
+apply    (simp add: invocation_declclass_def dynlookup_def dynmethd_Ext_foo)
+apply    (rule eval_Is (* Expr *))
+apply    (rule eval_Is (* Ass *))
+apply     (rule eval_Is (* FVar *))
+apply       (rule init_done, simp)
+apply      (rule eval_Is (* Cast *))
+apply       (rule eval_Is (* Acc *))
+apply       (rule eval_Is (* LVar *))
+apply      (simp)
+apply     (simp split del: split_if)
+apply    (rule eval_Is (* XcptE *))
+apply   (simp)
+      (* end method call *)
+apply  (rule sxalloc.intros)
+apply  (rule halloc.New)
+apply   (erule alloc_one [THEN conjunct1])
+apply   (simp (no_asm_simp))
+apply  (simp (no_asm_simp))
+apply (simp add: gupd_def lupd_def obj_ty_def split del: split_if)
+apply (drule alloc_one [THEN conjunct1])
+apply  (simp (no_asm_simp))
+apply (erule_tac V = "atleast_free ?h two" in thin_rl)
+apply (drule_tac x = "a" in new_AddrD2 [THEN spec])
+apply simp
+apply (rule eval_Is (* While *))
+apply  (rule eval_Is (* Acc *))
+apply  (rule eval_Is (* AVar *))
+apply    (rule eval_Is (* Acc *))
+apply    (rule eval_Is (* FVar *))
+apply      (rule init_done, simp)
+apply     (rule eval_Is (* StatRef *))
+apply    (simp)
+apply   (rule eval_Is (* Lit *))
+apply  (simp (no_asm_simp))
+apply (auto simp add: in_bounds_def)
+done
+declare Pair_eq [simp]
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Name.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,123 @@
+(*  Title:      isabelle/Bali/Name.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* Java names *}
+
+theory Name = Basis:
+
+(* cf. 6.5 *) 
+typedecl tnam		(* ordinary type name, i.e. class or interface name *)
+typedecl pname          (* package name *)
+typedecl mname		(* method name *)
+typedecl vname          (* variable or field name *)
+typedecl label          (* label as destination of break or continue *)
+
+arities
+  tnam  :: "type"
+  pname :: "type"
+  vname :: "type"
+  mname :: "type"
+  label :: "type"
+
+
+datatype ename        (* expression name *) 
+        = VNam vname 
+        | Res         (* special name to model the return value of methods *)
+
+datatype lname        (* names for local variables and the This pointer *)
+        = EName ename 
+        | This
+syntax   
+  VName  :: "vname \<Rightarrow> lname"
+  Result :: lname
+
+translations
+  "VName n" == "EName (VNam n)"
+  "Result"  == "EName Res"
+
+datatype xname		(* names of standard exceptions *)
+	= Throwable
+	| NullPointer | OutOfMemory | ClassCast   
+	| NegArrSize  | IndOutBound | ArrStore
+
+lemma xn_cases: 
+  "xn = Throwable   \<or> xn = NullPointer \<or>  
+         xn = OutOfMemory \<or> xn = ClassCast \<or> 
+         xn = NegArrSize  \<or> xn = IndOutBound \<or> xn = ArrStore"
+apply (induct xn)
+apply auto
+done
+
+lemma xn_cases_old:   (* FIXME cf. Example.thy *)
+  "ALL xn. xn = Throwable   \<or> xn = NullPointer \<or>  
+         xn = OutOfMemory \<or> xn = ClassCast \<or> 
+         xn = NegArrSize  \<or> xn = IndOutBound \<or> xn = ArrStore"
+apply (rule allI)
+apply (induct_tac xn)
+apply auto
+done
+
+datatype tname	(* type names for standard classes and other type names *)
+	= Object_
+	| SXcpt_   xname
+	| TName   tnam
+
+record   qtname = (* qualified tname cf. 6.5.3, 6.5.4*)
+          pid :: pname 
+          tid :: tname
+
+axclass has_pname < "type"
+consts pname::"'a::has_pname \<Rightarrow> pname"
+
+instance pname::has_pname
+by (intro_classes)
+
+defs (overloaded)
+pname_pname_def: "pname (p::pname) \<equiv> p"
+
+axclass has_tname < "type"
+consts tname::"'a::has_tname \<Rightarrow> tname"
+
+instance tname::has_tname
+by (intro_classes)
+
+defs (overloaded)
+tname_tname_def: "tname (t::tname) \<equiv> t"
+
+axclass has_qtname < "type"
+consts qtname:: "'a::has_qtname \<Rightarrow> qtname"
+
+(* Declare qtname as instance of has_qtname *)
+instance pid_field_type::(has_pname,"type") has_qtname
+by intro_classes 
+
+defs (overloaded)
+qtname_qtname_def: "qtname (q::qtname) \<equiv> q"
+
+translations
+  "mname"  <= "Name.mname"
+  "xname"  <= "Name.xname"
+  "tname"  <= "Name.tname"
+  "ename"  <= "Name.ename"
+  "qtname" <= (type) "\<lparr>pid::pname,tid::tname\<rparr>"
+  (type) "'a qtname_scheme" <= (type) "\<lparr>pid::pname,tid::tname,\<dots>::'a\<rparr>"
+
+
+consts java_lang::pname (* package java.lang *)
+
+consts 
+  Object :: qtname
+  SXcpt  :: "xname \<Rightarrow> qtname"
+defs
+  Object_def: "Object  \<equiv> \<lparr>pid = java_lang, tid = Object_\<rparr>"
+  SXcpt_def:  "SXcpt   \<equiv> \<lambda>x.  \<lparr>pid = java_lang, tid = SXcpt_ x\<rparr>"
+
+lemma Object_neq_SXcpt [simp]: "Object \<noteq> SXcpt xn"
+by (simp add: Object_def SXcpt_def)
+
+lemma SXcpt_inject [simp]: "(SXcpt xn = SXcpt xm) = (xn = xm)"
+by (simp add: SXcpt_def)
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/ROOT.ML	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,13 @@
+
+use_thy "WellForm";
+
+(*The dynamic part of Bali, including type-safety*)
+use_thy "Evaln";
+use_thy "Example"; 
+use_thy "TypeSafe";
+(*###use_thy "Trans";*)
+
+(*The Hoare logic for Bali*)
+use_thy "AxExample";
+use_thy "AxSound";
+use_thy "AxCompl";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/State.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,739 @@
+(*  Title:      isabelle/Bali/State.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* State for evaluation of Java expressions and statements *}
+
+theory State = DeclConcepts:
+
+text {*
+design issues:
+\begin{itemize}
+\item all kinds of objects (class instances, arrays, and class objects)
+  are handeled via a general object abstraction
+\item the heap and the map for class objects are combined into a single table
+  @{text "(recall (loc, obj) table \<times> (qtname, obj) table  ~=  (loc + qtname, obj) table)"}
+\end{itemize}
+*}
+
+section "objects"
+
+datatype  obj_tag =     (* tag for generic object   *)
+	  CInst qtname   (* class instance           *)
+	| Arr  ty int   (* array with component type and length *)
+     (* | CStat            the tag is irrelevant for a class object,
+			   i.e. the static fields of a class *)
+
+types	vn   = "fspec + int"                    (* variable name      *)
+record	obj  = 
+          tag :: "obj_tag"                      (* generalized object *)
+          values :: "(vn, val) table"      
+
+translations 
+  "fspec" <= (type) "vname \<times> qtname" 
+  "vn"    <= (type) "fspec + int"
+  "obj"   <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option\<rparr>"
+  "obj"   <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option,\<dots>::'a\<rparr>"
+
+constdefs
+  
+  the_Arr :: "obj option \<Rightarrow> ty \<times> int \<times> (vn, val) table"
+ "the_Arr obj \<equiv> \<epsilon>(T,k,t). obj = Some \<lparr>tag=Arr T k,values=t\<rparr>"
+
+lemma the_Arr_Arr [simp]: "the_Arr (Some \<lparr>tag=Arr T k,values=cs\<rparr>) = (T,k,cs)"
+apply (auto simp: the_Arr_def)
+done
+
+lemma the_Arr_Arr1 [simp,intro,dest]:
+ "\<lbrakk>tag obj = Arr T k\<rbrakk> \<Longrightarrow> the_Arr (Some obj) = (T,k,values obj)"
+apply (auto simp add: the_Arr_def)
+done
+
+constdefs
+
+  upd_obj       :: "vn \<Rightarrow> val \<Rightarrow> obj \<Rightarrow> obj" 
+ "upd_obj n v \<equiv> \<lambda> obj . obj \<lparr>values:=(values obj)(n\<mapsto>v)\<rparr>"
+
+lemma upd_obj_def2 [simp]: 
+  "upd_obj n v obj = obj \<lparr>values:=(values obj)(n\<mapsto>v)\<rparr>" 
+apply (auto simp: upd_obj_def)
+done
+
+constdefs
+  obj_ty        :: "obj \<Rightarrow> ty"
+ "obj_ty obj    \<equiv> case tag obj of 
+                    CInst C \<Rightarrow> Class C 
+                  | Arr T k \<Rightarrow> T.[]"
+
+lemma obj_ty_eq [intro!]: "obj_ty \<lparr>tag=oi,values=x\<rparr> = obj_ty \<lparr>tag=oi,values=y\<rparr>" 
+by (simp add: obj_ty_def)
+
+
+lemma obj_ty_eq1 [intro!,dest]: 
+  "tag obj = tag obj' \<Longrightarrow> obj_ty obj = obj_ty obj'" 
+by (simp add: obj_ty_def)
+
+lemma obj_ty_cong [simp]: 
+  "obj_ty (obj \<lparr>values:=vs\<rparr>) = obj_ty obj" 
+by auto
+(*
+lemma obj_ty_cong [simp]: 
+  "obj_ty (obj \<lparr>values:=vs(n\<mapsto>v)\<rparr>) = obj_ty (obj \<lparr>values:=vs\<rparr>)" 
+by auto
+*)
+lemma obj_ty_CInst [simp]: 
+ "obj_ty \<lparr>tag=CInst C,values=vs\<rparr> = Class C" 
+by (simp add: obj_ty_def)
+
+lemma obj_ty_CInst1 [simp,intro!,dest]: 
+ "\<lbrakk>tag obj = CInst C\<rbrakk> \<Longrightarrow> obj_ty obj = Class C" 
+by (simp add: obj_ty_def)
+
+lemma obj_ty_Arr [simp]: 
+ "obj_ty \<lparr>tag=Arr T i,values=vs\<rparr> = T.[]"
+by (simp add: obj_ty_def)
+
+lemma obj_ty_Arr1 [simp,intro!,dest]: 
+ "\<lbrakk>tag obj = Arr T i\<rbrakk> \<Longrightarrow> obj_ty obj = T.[]"
+by (simp add: obj_ty_def)
+
+lemma obj_ty_widenD: 
+ "G\<turnstile>obj_ty obj\<preceq>RefT t \<Longrightarrow> (\<exists>C. tag obj = CInst C) \<or> (\<exists>T k. tag obj = Arr T k)"
+apply (unfold obj_ty_def)
+apply (auto split add: obj_tag.split_asm)
+done
+
+constdefs
+
+  obj_class :: "obj \<Rightarrow> qtname"
+ "obj_class obj \<equiv> case tag obj of 
+                    CInst C \<Rightarrow> C 
+                  | Arr T k \<Rightarrow> Object"
+
+
+lemma obj_class_CInst [simp]: "obj_class \<lparr>tag=CInst C,values=vs\<rparr> = C" 
+by (auto simp: obj_class_def)
+
+lemma obj_class_CInst1 [simp,intro!,dest]: 
+  "tag obj = CInst C \<Longrightarrow> obj_class obj = C" 
+by (auto simp: obj_class_def)
+
+lemma obj_class_Arr [simp]: "obj_class \<lparr>tag=Arr T k,values=vs\<rparr> = Object" 
+by (auto simp: obj_class_def)
+
+lemma obj_class_Arr1 [simp,intro!,dest]: 
+ "tag obj = Arr T k \<Longrightarrow> obj_class obj = Object" 
+by (auto simp: obj_class_def)
+
+lemma obj_ty_obj_class: "G\<turnstile>obj_ty obj\<preceq> Class statC = G\<turnstile>obj_class obj \<preceq>\<^sub>C statC"
+apply (case_tac "tag obj")
+apply (auto simp add: obj_ty_def obj_class_def)
+apply (case_tac "statC = Object")
+apply (auto dest: widen_Array_Class)
+done
+
+section "object references"
+
+types oref = "loc + qtname"          (* generalized object reference *)
+syntax
+  Heap  :: "loc   \<Rightarrow> oref"
+  Stat  :: "qtname \<Rightarrow> oref"
+
+translations
+  "Heap" => "Inl"
+  "Stat" => "Inr"
+  "oref" <= (type) "loc + qtname"
+
+constdefs
+  fields_table::
+    "prog \<Rightarrow> qtname \<Rightarrow> (fspec \<Rightarrow> field \<Rightarrow> bool)  \<Rightarrow> (fspec, ty) table"
+ "fields_table G C P 
+    \<equiv> option_map type \<circ> table_of (filter (split P) (DeclConcepts.fields G C))"
+
+lemma fields_table_SomeI: 
+"\<lbrakk>table_of (DeclConcepts.fields G C) n = Some f; P n f\<rbrakk> 
+ \<Longrightarrow> fields_table G C P n = Some (type f)"
+apply (unfold fields_table_def)
+apply clarsimp
+apply (rule exI)
+apply (rule conjI)
+apply (erule map_of_filter_in)
+apply assumption
+apply simp
+done
+
+(* unused *)
+lemma fields_table_SomeD': "fields_table G C P fn = Some T \<Longrightarrow>  
+  \<exists>f. (fn,f)\<in>set(DeclConcepts.fields G C) \<and> type f = T"
+apply (unfold fields_table_def)
+apply clarsimp
+apply (drule map_of_SomeD)
+apply auto
+done
+
+lemma fields_table_SomeD: 
+"\<lbrakk>fields_table G C P fn = Some T; unique (DeclConcepts.fields G C)\<rbrakk> \<Longrightarrow>  
+  \<exists>f. table_of (DeclConcepts.fields G C) fn = Some f \<and> type f = T"
+apply (unfold fields_table_def)
+apply clarsimp
+apply (rule exI)
+apply (rule conjI)
+apply (erule table_of_filter_unique_SomeD)
+apply assumption
+apply simp
+done
+
+constdefs
+  in_bounds :: "int \<Rightarrow> int \<Rightarrow> bool"            ("(_/ in'_bounds _)" [50, 51] 50)
+ "i in_bounds k \<equiv> 0 \<le> i \<and> i < k"
+
+  arr_comps :: "'a \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a option"
+ "arr_comps T k \<equiv> \<lambda>i. if i in_bounds k then Some T else None"
+  
+  var_tys       :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> (vn, ty) table"
+"var_tys G oi r 
+  \<equiv> case r of 
+      Heap a \<Rightarrow> (case oi of 
+                   CInst C \<Rightarrow> fields_table G C (\<lambda>n f. \<not>static f) (+) empty
+                 | Arr T k \<Rightarrow> empty (+) arr_comps T k)
+    | Stat C \<Rightarrow> fields_table G C (\<lambda>fn f. declclassf fn = C \<and> static f) 
+                (+) empty"
+
+lemma var_tys_Some_eq: 
+ "var_tys G oi r n = Some T 
+  = (case r of 
+       Inl a \<Rightarrow> (case oi of  
+                   CInst C \<Rightarrow> (\<exists>nt. n = Inl nt \<and> fields_table G C (\<lambda>n f. 
+                               \<not>static f) nt = Some T)  
+                 | Arr t k \<Rightarrow> (\<exists> i. n = Inr i  \<and> i in_bounds k \<and> t = T))  
+     | Inr C \<Rightarrow> (\<exists>nt. n = Inl nt \<and> 
+                 fields_table G C (\<lambda>fn f. declclassf fn = C \<and> static f) nt 
+                  = Some T))"
+apply (unfold var_tys_def arr_comps_def)
+apply (force split add: sum.split_asm sum.split obj_tag.split)
+done
+
+
+section "stores"
+
+types	globs                  (* global variables: heap and static variables *)
+	= "(oref , obj) table"
+	heap
+	= "(loc  , obj) table"
+	locals
+	= "(lname, val) table" (* local variables *)
+
+translations
+ "globs"  <= (type) "(oref , obj) table"
+ "heap"   <= (type) "(loc  , obj) table"
+ "locals" <= (type) "(lname, val) table"
+
+datatype st = (* pure state, i.e. contents of all variables *)
+	 st globs locals
+
+subsection "access"
+
+constdefs
+
+  globs  :: "st \<Rightarrow> globs"
+ "globs  \<equiv> st_case (\<lambda>g l. g)"
+  
+  locals :: "st \<Rightarrow> locals"
+ "locals \<equiv> st_case (\<lambda>g l. l)"
+
+  heap   :: "st \<Rightarrow> heap"
+ "heap s \<equiv> globs s \<circ> Heap"
+
+
+lemma globs_def2 [simp]: " globs (st g l) = g"
+by (simp add: globs_def)
+
+lemma locals_def2 [simp]: "locals (st g l) = l"
+by (simp add: locals_def)
+
+lemma heap_def2 [simp]:  "heap s a=globs s (Heap a)"
+by (simp add: heap_def)
+
+
+syntax
+  val_this     :: "st \<Rightarrow> val"
+  lookup_obj   :: "st \<Rightarrow> val \<Rightarrow> obj"
+
+translations
+ "val_this s"       == "the (locals s This)" 
+ "lookup_obj s a'"  == "the (heap s (the_Addr a'))"
+
+subsection "memory allocation"
+
+constdefs
+  new_Addr     :: "heap \<Rightarrow> loc option"
+ "new_Addr h   \<equiv> if (\<forall>a. h a \<noteq> None) then None else Some (\<epsilon>a. h a = None)"
+
+lemma new_AddrD: "new_Addr h = Some a \<Longrightarrow> h a = None"
+apply (unfold new_Addr_def)
+apply auto
+apply (case_tac "h (SOME a\<Colon>loc. h a = None)")
+apply simp
+apply (fast intro: someI2)
+done
+
+lemma new_AddrD2: "new_Addr h = Some a \<Longrightarrow> \<forall>b. h b \<noteq> None \<longrightarrow> b \<noteq> a"
+apply (drule new_AddrD)
+apply auto
+done
+
+lemma new_Addr_SomeI: "h a = None \<Longrightarrow> \<exists>b. new_Addr h = Some b \<and> h b = None"
+apply (unfold new_Addr_def)
+apply (frule not_Some_eq [THEN iffD2])
+apply auto
+apply  (drule not_Some_eq [THEN iffD2])
+apply  auto
+apply (fast intro!: someI2)
+done
+
+
+subsection "initialization"
+
+syntax
+
+  init_vals     :: "('a, ty) table \<Rightarrow> ('a, val) table"
+
+translations
+ "init_vals vs"    == "option_map default_val \<circ> vs"
+
+lemma init_arr_comps_base [simp]: "init_vals (arr_comps T 0) = empty"
+apply (unfold arr_comps_def in_bounds_def)
+apply (rule ext)
+apply auto
+done
+
+lemma init_arr_comps_step [simp]: 
+"0 < j \<Longrightarrow> init_vals (arr_comps T  j    ) =  
+           init_vals (arr_comps T (j - 1))(j - 1\<mapsto>default_val T)"
+apply (unfold arr_comps_def in_bounds_def)
+apply (rule ext)
+apply auto
+done
+
+subsection "update"
+
+constdefs
+  gupd       :: "oref  \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st"        ("gupd'(_\<mapsto>_')"[10,10]1000)
+ "gupd r obj  \<equiv> st_case (\<lambda>g l. st (g(r\<mapsto>obj)) l)"
+
+  lupd       :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st"        ("lupd'(_\<mapsto>_')"[10,10]1000)
+ "lupd vn v   \<equiv> st_case (\<lambda>g l. st g (l(vn\<mapsto>v)))"
+
+  upd_gobj   :: "oref \<Rightarrow> vn \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" 
+ "upd_gobj r n v \<equiv> st_case (\<lambda>g l. st (chg_map (upd_obj n v) r g) l)"
+
+  set_locals  :: "locals \<Rightarrow> st \<Rightarrow> st"
+ "set_locals l \<equiv> st_case (\<lambda>g l'. st g l)"
+  
+  init_obj    :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> st \<Rightarrow> st"
+ "init_obj G oi r \<equiv> gupd(r\<mapsto>\<lparr>tag=oi, values=init_vals (var_tys G oi r)\<rparr>)"
+
+syntax
+  init_class_obj :: "prog \<Rightarrow> qtname \<Rightarrow> st \<Rightarrow> st"
+
+translations
+ "init_class_obj G C" == "init_obj G arbitrary (Inr C)"
+
+lemma gupd_def2 [simp]: "gupd(r\<mapsto>obj) (st g l) = st (g(r\<mapsto>obj)) l"
+apply (unfold gupd_def)
+apply (simp (no_asm))
+done
+
+lemma lupd_def2 [simp]: "lupd(vn\<mapsto>v) (st g l) = st g (l(vn\<mapsto>v))"
+apply (unfold lupd_def)
+apply (simp (no_asm))
+done
+
+lemma globs_gupd [simp]: "globs  (gupd(r\<mapsto>obj) s) = globs s(r\<mapsto>obj)"
+apply (induct "s")
+by (simp add: gupd_def)
+
+lemma globs_lupd [simp]: "globs  (lupd(vn\<mapsto>v ) s) = globs  s"
+apply (induct "s")
+by (simp add: lupd_def)
+
+lemma locals_gupd [simp]: "locals (gupd(r\<mapsto>obj) s) = locals s"
+apply (induct "s")
+by (simp add: gupd_def)
+
+lemma locals_lupd [simp]: "locals (lupd(vn\<mapsto>v ) s) = locals s(vn\<mapsto>v )"
+apply (induct "s")
+by (simp add: lupd_def)
+
+lemma globs_upd_gobj_new [rule_format (no_asm), simp]: 
+  "globs s r = None \<longrightarrow> globs (upd_gobj r n v s) = globs s"
+apply (unfold upd_gobj_def)
+apply (induct "s")
+apply auto
+done
+
+lemma globs_upd_gobj_upd [rule_format (no_asm), simp]: 
+"globs s r=Some obj\<longrightarrow> globs (upd_gobj r n v s) = globs s(r\<mapsto>upd_obj n v obj)"
+apply (unfold upd_gobj_def)
+apply (induct "s")
+apply auto
+done
+
+lemma locals_upd_gobj [simp]: "locals (upd_gobj r n v s) = locals s"
+apply (induct "s")
+by (simp add: upd_gobj_def) 
+
+
+lemma globs_init_obj [simp]: "globs (init_obj G oi r s) t =  
+  (if t=r then Some \<lparr>tag=oi,values=init_vals (var_tys G oi r)\<rparr> else globs s t)"
+apply (unfold init_obj_def)
+apply (simp (no_asm))
+done
+
+lemma locals_init_obj [simp]: "locals (init_obj G oi r s) = locals s"
+by (simp add: init_obj_def)
+  
+lemma surjective_st [simp]: "st (globs s) (locals s) = s"
+apply (induct "s")
+by auto
+
+lemma surjective_st_init_obj: 
+ "st (globs (init_obj G oi r s)) (locals s) = init_obj G oi r s"
+apply (subst locals_init_obj [THEN sym])
+apply (rule surjective_st)
+done
+
+lemma heap_heap_upd [simp]: 
+  "heap (st (g(Inl a\<mapsto>obj)) l) = heap (st g l)(a\<mapsto>obj)"
+apply (rule ext)
+apply (simp (no_asm))
+done
+lemma heap_stat_upd [simp]: "heap (st (g(Inr C\<mapsto>obj)) l) = heap (st g l)"
+apply (rule ext)
+apply (simp (no_asm))
+done
+lemma heap_local_upd [simp]: "heap (st g (l(vn\<mapsto>v))) = heap (st g l)"
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+lemma heap_gupd_Heap [simp]: "heap (gupd(Heap a\<mapsto>obj) s) = heap s(a\<mapsto>obj)"
+apply (rule ext)
+apply (simp (no_asm))
+done
+lemma heap_gupd_Stat [simp]: "heap (gupd(Stat C\<mapsto>obj) s) = heap s"
+apply (rule ext)
+apply (simp (no_asm))
+done
+lemma heap_lupd [simp]: "heap (lupd(vn\<mapsto>v) s) = heap s"
+apply (rule ext)
+apply (simp (no_asm))
+done
+
+(*
+lemma heap_upd_gobj_Heap: "!!a. heap (upd_gobj (Heap a) n v s) = ?X"
+apply (rule ext)
+apply (simp (no_asm))
+apply (case_tac "globs s (Heap a)")
+apply  auto
+*)
+
+lemma heap_upd_gobj_Stat [simp]: "heap (upd_gobj (Stat C) n v s) = heap s"
+apply (rule ext)
+apply (simp (no_asm))
+apply (case_tac "globs s (Stat C)")
+apply  auto
+done
+
+lemma set_locals_def2 [simp]: "set_locals l (st g l') = st g l"
+apply (unfold set_locals_def)
+apply (simp (no_asm))
+done
+
+lemma set_locals_id [simp]: "set_locals (locals s) s = s"
+apply (unfold set_locals_def)
+apply (induct_tac "s")
+apply (simp (no_asm))
+done
+
+lemma set_set_locals [simp]: "set_locals l (set_locals l' s) = set_locals l s"
+apply (unfold set_locals_def)
+apply (induct_tac "s")
+apply (simp (no_asm))
+done
+
+lemma locals_set_locals [simp]: "locals (set_locals l s) = l"
+apply (unfold set_locals_def)
+apply (induct_tac "s")
+apply (simp (no_asm))
+done
+
+lemma globs_set_locals [simp]: "globs (set_locals l s) = globs s"
+apply (unfold set_locals_def)
+apply (induct_tac "s")
+apply (simp (no_asm))
+done
+
+lemma heap_set_locals [simp]: "heap (set_locals l s) = heap s"
+apply (unfold heap_def)
+apply (induct_tac "s")
+apply (simp (no_asm))
+done
+
+
+section "abrupt completion"
+
+
+datatype xcpt        (* exception *)
+	= Loc loc    (* location of allocated execption object *)
+	| Std xname  (* intermediate standard exception, see Eval.thy *)
+
+datatype abrupt      (* abrupt completion *) 
+        = Xcpt xcpt  (* exception *)
+        | Jump jump  (* break, continue, return *)
+consts
+
+  the_Xcpt :: "abrupt \<Rightarrow> xcpt"
+  the_Jump :: "abrupt => jump"
+  the_Loc  :: "xcpt \<Rightarrow> loc"
+  the_Std  :: "xcpt \<Rightarrow> xname"
+
+primrec "the_Xcpt (Xcpt x) = x"
+primrec "the_Jump (Jump j) = j"
+primrec "the_Loc (Loc a) = a"
+primrec "the_Std (Std x) = x"
+
+types
+  abopt  = "abrupt option"
+	
+
+constdefs
+  abrupt_if    :: "bool \<Rightarrow> abopt \<Rightarrow> abopt \<Rightarrow> abopt"
+ "abrupt_if c x' x \<equiv> if c \<and> (x = None) then x' else x"
+
+lemma abrupt_if_True_None [simp]: "abrupt_if True x None = x"
+by (simp add: abrupt_if_def)
+
+lemma abrupt_if_True_not_None [simp]: "x \<noteq> None \<Longrightarrow> abrupt_if True x y \<noteq> None"
+by (simp add: abrupt_if_def)
+
+lemma abrupt_if_False [simp]: "abrupt_if False x y = y"
+by (simp add: abrupt_if_def)
+
+lemma abrupt_if_Some [simp]: "abrupt_if c x (Some y) = Some y"
+by (simp add: abrupt_if_def)
+
+lemma abrupt_if_not_None [simp]: "y \<noteq> None \<Longrightarrow> abrupt_if c x y = y"
+apply (simp add: abrupt_if_def)
+by auto
+
+
+lemma split_abrupt_if: 
+"P (abrupt_if c x' x) = 
+      ((c \<and> x = None \<longrightarrow> P x') \<and> (\<not> (c \<and> x = None) \<longrightarrow> P x))"
+apply (unfold abrupt_if_def)
+apply (split split_if)
+apply auto
+done
+
+syntax
+
+  raise_if :: "bool \<Rightarrow> xname \<Rightarrow> abopt \<Rightarrow> abopt"
+  np       :: "val  \<spacespace>        \<Rightarrow> abopt \<Rightarrow> abopt"
+  check_neg:: "val  \<spacespace>        \<Rightarrow> abopt \<Rightarrow> abopt"
+  
+translations
+
+ "raise_if c xn" == "abrupt_if c (Some (Xcpt (Std xn)))"
+ "np v"          == "raise_if (v = Null)      NullPointer"
+ "check_neg i'"  == "raise_if (the_Intg i'<0) NegArrSize"
+
+lemma raise_if_None [simp]: "(raise_if c x y = None) = (\<not>c \<and> y = None)"
+apply (simp add: abrupt_if_def)
+by auto
+declare raise_if_None [THEN iffD1, dest!]
+
+lemma if_raise_if_None [simp]: 
+  "((if b then y else raise_if c x y) = None) = ((c \<longrightarrow> b) \<and> y = None)"
+apply (simp add: abrupt_if_def)
+apply auto
+done
+
+lemma raise_if_SomeD [dest!]:
+  "raise_if c x y = Some z \<Longrightarrow> c \<and> z=(Xcpt (Std x)) \<and> y=None \<or> (y=Some z)"
+apply (case_tac y)
+apply (case_tac c)
+apply (simp add: abrupt_if_def)
+apply (simp add: abrupt_if_def)
+apply auto
+done
+
+constdefs
+   absorb :: "jump \<Rightarrow> abopt \<Rightarrow> abopt"
+  "absorb j a \<equiv> if a=Some (Jump j) then None else a"
+
+lemma absorb_SomeD [dest!]: "absorb j a = Some x \<Longrightarrow> a = Some x"
+by (auto simp add: absorb_def)
+
+lemma absorb_same [simp]: "absorb j (Some (Jump j)) = None"
+by (auto simp add: absorb_def)
+
+lemma absorb_other [simp]: "a \<noteq> Some (Jump j) \<Longrightarrow> absorb j a = a"
+by (auto simp add: absorb_def)
+
+
+section "full program state"
+
+types
+  state = "abopt \<times> st"          (* state including exception information *)
+
+syntax 
+  Norm   :: "st \<Rightarrow> state"
+  abrupt :: "state \<Rightarrow> abopt"
+  store  :: "state \<Rightarrow> st"
+
+translations
+   
+  "Norm s"     == "(None,s)" 
+  "abrupt"     => "fst"
+  "store"      => "snd"
+  "abopt"       <= (type) "State.abrupt option"
+  "abopt"       <= (type) "abrupt option"
+  "state"      <= (type) "abopt \<times> State.st"
+  "state"      <= (type) "abopt \<times> st"
+
+
+
+lemma single_stateE: "\<forall>Z. Z = (s::state) \<Longrightarrow> False"
+apply (erule_tac x = "(Some k,y)" in all_dupE)
+apply (erule_tac x = "(None,y)" in allE)
+apply clarify
+done
+
+lemma state_not_single: "All (op = (x::state)) \<Longrightarrow> R"
+apply (drule_tac x = "(if abrupt x = None then Some ?x else None,?y)" in spec)
+apply clarsimp
+done
+
+constdefs
+
+  normal     :: "state \<Rightarrow> bool"
+ "normal \<equiv> \<lambda>s. abrupt s = None"
+
+lemma normal_def2 [simp]: "normal s = (abrupt s = None)"
+apply (unfold normal_def)
+apply (simp (no_asm))
+done
+
+constdefs
+  heap_free :: "nat \<Rightarrow> state \<Rightarrow> bool"
+ "heap_free n \<equiv> \<lambda>s. atleast_free (heap (store s)) n"
+
+lemma heap_free_def2 [simp]: "heap_free n s = atleast_free (heap (store s)) n"
+apply (unfold heap_free_def)
+apply simp
+done
+
+subsection "update"
+
+constdefs
+ 
+  abupd     :: "(abopt \<Rightarrow> abopt) \<Rightarrow> state \<Rightarrow> state"
+ "abupd f \<equiv> prod_fun f id"
+
+  supd     :: "(st \<Rightarrow> st) \<Rightarrow> state \<Rightarrow> state" 
+ "supd \<equiv> prod_fun id"
+  
+lemma abupd_def2 [simp]: "abupd f (x,s) = (f x,s)"
+by (simp add: abupd_def)
+
+lemma abupd_abrupt_if_False [simp]: "\<And> s. abupd (abrupt_if False xo) s = s"
+by simp
+
+lemma supd_def2 [simp]: "supd f (x,s) = (x,f s)"
+by (simp add: supd_def)
+
+lemma supd_lupd [simp]: 
+ "\<And> s. supd (lupd vn v ) s = (abrupt s,lupd vn v (store s))"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (simp (no_asm))
+done
+
+
+lemma supd_gupd [simp]: 
+ "\<And> s. supd (gupd r obj) s = (abrupt s,gupd r obj (store s))"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (simp (no_asm))
+done
+
+lemma supd_init_obj [simp]: 
+ "supd (init_obj G oi r) s = (abrupt s,init_obj G oi r (store s))"
+apply (unfold init_obj_def)
+apply (simp (no_asm))
+done
+
+syntax
+
+  set_lvars     :: "locals \<Rightarrow> state \<Rightarrow> state"
+  restore_lvars :: "state  \<Rightarrow> state \<Rightarrow> state"
+  
+translations
+
+ "set_lvars l" == "supd (set_locals l)"
+ "restore_lvars s' s" == "set_lvars (locals (store s')) s"
+
+lemma set_set_lvars [simp]: "\<And> s. set_lvars l (set_lvars l' s) = set_lvars l s"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (simp (no_asm))
+done
+
+lemma set_lvars_id [simp]: "\<And> s. set_lvars (locals (store s)) s = s"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (simp (no_asm))
+done
+
+section "initialisation test"
+
+constdefs
+
+  inited   :: "qtname \<Rightarrow> globs \<Rightarrow> bool"
+ "inited C g \<equiv> g (Stat C) \<noteq> None"
+
+  initd    :: "qtname \<Rightarrow> state \<Rightarrow> bool"
+ "initd C \<equiv> inited C \<circ> globs \<circ> store"
+
+lemma not_inited_empty [simp]: "\<not>inited C empty"
+apply (unfold inited_def)
+apply (simp (no_asm))
+done
+
+lemma inited_gupdate [simp]: "inited C (g(r\<mapsto>obj)) = (inited C g \<or> r = Stat C)"
+apply (unfold inited_def)
+apply (auto split add: st.split)
+done
+
+lemma inited_init_class_obj [intro!]: "inited C (globs (init_class_obj G C s))"
+apply (unfold inited_def)
+apply (simp (no_asm))
+done
+
+lemma not_initedD: "\<not> inited C g \<Longrightarrow> g (Stat C) = None"
+apply (unfold inited_def)
+apply (erule notnotD)
+done
+
+lemma initedD: "inited C g \<Longrightarrow> \<exists> obj. g (Stat C) = Some obj"
+apply (unfold inited_def)
+apply auto
+done
+
+lemma initd_def2 [simp]: "initd C s = inited C (globs (store s))"
+apply (unfold initd_def)
+apply (simp (no_asm))
+done
+
+
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Table.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,445 @@
+(*  Title:      isabelle/Bali/Table.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* Abstract tables and their implementation as lists *}
+
+theory Table = Basis:
+
+text {*
+design issues:
+\begin{itemize}
+\item definition of table: infinite map vs. list vs. finite set
+      list chosen, because:
+  \begin{itemize} 
+  \item[+]  a priori finite
+  \item[+]  lookup is more operational than for finite set
+  \item[-]  not very abstract, but function table converts it to abstract 
+            mapping
+  \end{itemize}
+\item coding of lookup result: Some/None vs. value/arbitrary
+   Some/None chosen, because:
+  \begin{itemize}
+  \item[++] makes definedness check possible (applies also to finite set),
+     which is important for the type standard, hiding/overriding, etc.
+     (though it may perhaps be possible at least for the operational semantics
+      to treat programs as infinite, i.e. where classes, fields, methods etc.
+      of any name are considered to be defined)
+  \item[-]  sometimes awkward case distinctions, alleviated by operator 'the'
+  \end{itemize}
+\end{itemize}
+*}
+
+
+types ('a, 'b) table    (* table with key type 'a and contents type 'b *)
+      = "'a \<leadsto> 'b"
+      ('a, 'b) tables   (* non-unique table with key 'a and contents 'b *)
+      = "'a \<Rightarrow> 'b set"
+
+
+section "map of / table of"
+
+syntax
+  table_of      :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) table"    (* concrete table *)
+
+translations
+  "table_of" == "map_of"
+  
+  (type)"'a \<leadsto> 'b"       <= (type)"'a \<Rightarrow> 'b Option.option"
+  (type)"('a, 'b) table" <= (type)"'a \<leadsto> 'b"
+
+(* ### To map *)
+lemma override_find_left[simp]:
+"n k = None \<Longrightarrow> (m ++ n) k = m k"
+by (simp add: override_def)
+
+section {* Conditional Override *}
+constdefs
+cond_override:: 
+  "('b \<Rightarrow>'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b) table"
+
+(* when merging tables old and new, only override an entry of table old when  
+   the condition cond holds *)
+"cond_override cond old new \<equiv>
+ \<lambda> k.
+  (case new k of
+     None         \<Rightarrow> old k                       
+   | Some new_val \<Rightarrow> (case old k of
+                        None         \<Rightarrow> Some new_val          
+                      | Some old_val \<Rightarrow> (if cond new_val old_val
+                                         then Some new_val     
+                                         else Some old_val)))"
+
+lemma cond_override_empty1[simp]: "cond_override c empty t = t"
+by (simp add: cond_override_def expand_fun_eq)
+
+lemma cond_override_empty2[simp]: "cond_override c t empty = t"
+by (simp add: cond_override_def expand_fun_eq)
+
+lemma cond_override_None[simp]:
+ "old k = None \<Longrightarrow> (cond_override c old new) k = new k"
+by (simp add: cond_override_def)
+
+lemma cond_override_override:
+ "\<lbrakk>old k = Some ov;new k = Some nv; C nv ov\<rbrakk> 
+  \<Longrightarrow> (cond_override C old new) k = Some nv"
+by (auto simp add: cond_override_def)
+
+lemma cond_override_noOverride:
+ "\<lbrakk>old k = Some ov;new k = Some nv; \<not> (C nv ov)\<rbrakk> 
+  \<Longrightarrow> (cond_override C old new) k = Some ov"
+by (auto simp add: cond_override_def)
+
+lemma dom_cond_override: "dom (cond_override C s t) \<subseteq> dom s \<union> dom t"
+by (auto simp add: cond_override_def dom_def)
+
+lemma finite_dom_cond_override:
+ "\<lbrakk> finite (dom s); finite (dom t) \<rbrakk> \<Longrightarrow> finite (dom (cond_override C s t))"
+apply (rule_tac B="dom s \<union> dom t" in finite_subset)
+apply (rule dom_cond_override)
+by (rule finite_UnI)
+
+section {* Filter on Tables *}
+
+constdefs
+filter_tab:: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) table \<Rightarrow> ('a, 'b) table"
+"filter_tab c t \<equiv> \<lambda> k. (case t k of 
+                           None   \<Rightarrow> None
+                         | Some x \<Rightarrow> if c k x then Some x else None)"
+
+lemma filter_tab_empty[simp]: "filter_tab c empty = empty"
+by (simp add: filter_tab_def empty_def)
+
+lemma filter_tab_True[simp]: "filter_tab (\<lambda>x y. True) t = t"
+by (simp add: expand_fun_eq filter_tab_def)
+
+lemma filter_tab_False[simp]: "filter_tab (\<lambda>x y. False) t = empty"
+by (simp add: expand_fun_eq filter_tab_def empty_def)
+
+lemma filter_tab_ran_subset: "ran (filter_tab c t) \<subseteq> ran t"
+by (auto simp add: filter_tab_def ran_def)
+
+lemma filter_tab_range_subset: "range (filter_tab c t) \<subseteq> range t \<union> {None}"
+apply (auto simp add: filter_tab_def)
+apply (drule sym, blast)
+done
+
+lemma finite_range_filter_tab:
+  "finite (range t) \<Longrightarrow> finite (range (filter_tab c t))"
+apply (rule_tac B="range t \<union> {None} " in finite_subset)
+apply (rule filter_tab_range_subset)
+apply (auto intro: finite_UnI)
+done
+
+lemma filter_tab_SomeD[dest!]: 
+"filter_tab c t k = Some x \<Longrightarrow> (t k = Some x) \<and> c k x"
+by (auto simp add: filter_tab_def)
+
+lemma filter_tab_SomeI: "\<lbrakk>t k = Some x;C k x\<rbrakk> \<Longrightarrow>filter_tab C t k = Some x"
+by (simp add: filter_tab_def)
+
+lemma filter_tab_all_True: 
+ "\<forall> k y. t k = Some y \<longrightarrow> p k y \<Longrightarrow>filter_tab p t = t"
+apply (auto simp add: filter_tab_def expand_fun_eq)
+done
+
+lemma filter_tab_all_True_Some:
+ "\<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"
+by (auto simp add: filter_tab_def expand_fun_eq)
+
+lemma filter_tab_None: "t k = None \<Longrightarrow> filter_tab p t k = None"
+apply (simp add: filter_tab_def expand_fun_eq)
+done
+
+lemma filter_tab_dom_subset: "dom (filter_tab C t) \<subseteq> dom t"
+by (auto simp add: filter_tab_def dom_def)
+
+lemma filter_tab_eq: "\<lbrakk>a=b\<rbrakk> \<Longrightarrow> filter_tab C a = filter_tab C b"
+by (auto simp add: expand_fun_eq filter_tab_def)
+
+lemma finite_dom_filter_tab:
+"finite (dom t) \<Longrightarrow> finite (dom (filter_tab C t))"
+apply (rule_tac B="dom t" in finite_subset)
+by (rule filter_tab_dom_subset)
+
+
+lemma filter_tab_weaken:
+"\<lbrakk>\<forall> a \<in> t k: \<exists> b \<in> s k: P a b; 
+  \<And> k x y. \<lbrakk>t k = Some x;s k = Some y\<rbrakk> \<Longrightarrow> cond k x \<longrightarrow> cond k y
+ \<rbrakk> \<Longrightarrow> \<forall> a \<in> filter_tab cond t k: \<exists> b \<in> filter_tab cond s k: P a b"
+apply (auto simp add: filter_tab_def)
+done
+
+lemma cond_override_filter: 
+  "\<lbrakk>\<And> k old new. \<lbrakk>s k = Some new; t k = Some old\<rbrakk> 
+    \<Longrightarrow> (\<not> overC new old \<longrightarrow>  \<not> filterC k new) \<and> 
+        (overC new old \<longrightarrow> filterC k old \<longrightarrow> filterC k new)
+   \<rbrakk> \<Longrightarrow>
+    cond_override overC (filter_tab filterC t) (filter_tab filterC s) 
+    = filter_tab filterC (cond_override overC t s)"
+by (auto simp add: expand_fun_eq cond_override_def filter_tab_def )
+
+section {* Misc. *}
+
+lemma Ball_set_table: "(\<forall> (x,y)\<in> set l. P x y) \<Longrightarrow> \<forall> x. \<forall> y\<in> map_of l x: P x y"
+apply (erule make_imp)
+apply (induct l)
+apply simp
+apply (simp (no_asm))
+apply auto
+done
+
+lemma Ball_set_tableD: 
+  "\<lbrakk>(\<forall> (x,y)\<in> set l. P x y); x \<in> o2s (table_of l xa)\<rbrakk> \<Longrightarrow> P xa x"
+apply (frule Ball_set_table)
+by auto
+
+declare map_of_SomeD [elim]
+
+lemma table_of_Some_in_set:
+"table_of l k = Some x \<Longrightarrow> (k,x) \<in> set l"
+by auto
+
+lemma set_get_eq: 
+  "unique l \<Longrightarrow> (k, the (table_of l k)) \<in> set l = (table_of l k \<noteq> None)"
+apply safe
+apply (fast dest!: weak_map_of_SomeI)
+apply auto
+done
+
+lemma inj_Pair_const2: "inj (\<lambda>k. (k, C))"
+apply (rule injI)
+apply auto
+done
+
+lemma table_of_mapconst_SomeI:
+  "\<lbrakk>table_of t k = Some y'; snd y=y'; fst y=c\<rbrakk> \<Longrightarrow>
+   table_of (map (\<lambda>(k,x). (k,c,x)) t) k = Some y"
+apply (induct t)
+apply auto
+done
+
+lemma table_of_mapconst_NoneI:
+  "\<lbrakk>table_of t k = None\<rbrakk> \<Longrightarrow>
+   table_of (map (\<lambda>(k,x). (k,c,x)) t) k = None"
+apply (induct t)
+apply auto
+done
+
+lemmas table_of_map2_SomeI = inj_Pair_const2 [THEN map_of_mapk_SomeI, standard]
+
+lemma table_of_map_SomeI [rule_format (no_asm)]: "table_of t k = Some x \<longrightarrow>
+   table_of (map (\<lambda>(k,x). (k, f x)) t) k = Some (f x)"
+apply (induct_tac "t")
+apply auto
+done
+
+lemma table_of_remap_SomeD [rule_format (no_asm)]: 
+  "table_of (map (\<lambda>((k,k'),x). (k,(k',x))) t) k = Some (k',x) \<longrightarrow>
+  table_of t (k, k') = Some x"
+apply (induct_tac "t")
+apply  auto
+done
+
+lemma table_of_mapf_Some [rule_format (no_asm)]: "\<forall>x y. f x = f y \<longrightarrow> x = y \<Longrightarrow> 
+  table_of (map (\<lambda>(k,x). (k,f x)) t) k = Some (f x) \<longrightarrow> table_of t k = Some x"
+apply (induct_tac "t")
+apply  auto
+done
+
+lemma table_of_mapf_SomeD [rule_format (no_asm), dest!]: 
+"table_of (map (\<lambda>(k,x). (k, f x)) t) k = Some z \<longrightarrow> (\<exists>y\<in>table_of t k: z=f y)"
+apply (induct_tac "t")
+apply  auto
+done
+
+lemma table_of_mapf_NoneD [rule_format (no_asm), dest!]: 
+"table_of (map (\<lambda>(k,x). (k, f x)) t) k = None \<longrightarrow> (table_of t k = None)"
+apply (induct_tac "t")
+apply auto
+done
+
+lemma table_of_mapkey_SomeD [rule_format (no_asm), dest!]: 
+  "table_of (map (\<lambda>(k,x). ((k,C),x)) t) (k,D) = Some x \<longrightarrow> C = D \<and> table_of t k = Some x"
+apply (induct_tac "t")
+apply  auto
+done
+lemma table_of_mapkey_SomeD2 [rule_format (no_asm), dest!]: 
+  "table_of (map (\<lambda>(k,x). ((k,C),x)) t) ek = Some x 
+   \<longrightarrow> C = snd ek \<and> table_of t (fst ek) = Some x"
+apply (induct_tac "t")
+apply  auto
+done
+
+lemma table_append_Some_iff: "table_of (xs@ys) k = Some z = 
+ (table_of xs k = Some z \<or> (table_of xs k = None \<and> table_of ys k = Some z))"
+apply (simp only: map_of_override [THEN sym])
+apply (rule override_Some_iff)
+done
+
+lemma table_of_filter_unique_SomeD [rule_format (no_asm)]:
+  "table_of (filter P xs) k = Some z \<Longrightarrow> unique xs \<longrightarrow> table_of xs k = Some z"
+apply (induct xs)
+apply (auto del: map_of_SomeD intro!: map_of_SomeD)
+done
+
+
+consts
+  Un_tables      :: "('a, 'b) tables set \<Rightarrow> ('a, 'b) tables"
+  overrides_t    :: "('a, 'b) tables     \<Rightarrow> ('a, 'b) tables \<Rightarrow>
+                     ('a, 'b) tables"             (infixl "\<oplus>\<oplus>" 100)
+  hidings_entails:: "('a, 'b) tables \<Rightarrow> ('a, 'c) tables \<Rightarrow> 
+                     ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"   ("_ hidings _ entails _" 20)
+  (* variant for unique table: *)
+  hiding_entails :: "('a, 'b) table  \<Rightarrow> ('a, 'c) table  \<Rightarrow> 
+                     ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"   ("_ hiding _ entails _"  20)
+  (* variant for a unique table and conditional overriding: *)
+  cond_hiding_entails :: "('a, 'b) table  \<Rightarrow> ('a, 'c) table  
+                          \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"  
+                          ("_ hiding _ under _ entails _"  20)
+
+defs
+  Un_tables_def:       "Un_tables ts\<spacespace>\<spacespace>\<equiv> \<lambda>k. \<Union>t\<in>ts. t k"
+  overrides_t_def:     "s \<oplus>\<oplus> t        \<equiv> \<lambda>k. if t k = {} then s k else t k"
+  hidings_entails_def: "t hidings s entails R \<equiv> \<forall>k. \<forall>x\<in>t k. \<forall>y\<in>s k. R x y"
+  hiding_entails_def:  "t hiding  s entails R \<equiv> \<forall>k. \<forall>x\<in>t k: \<forall>y\<in>s k: R x y"
+  cond_hiding_entails_def: "t hiding  s under C entails R
+                            \<equiv> \<forall>k. \<forall>x\<in>t k: \<forall>y\<in>s k: C x y \<longrightarrow> R x y"
+
+section "Untables"
+
+lemma Un_tablesI [intro]:  "\<And>x. \<lbrakk>t \<in> ts; x \<in> t k\<rbrakk> \<Longrightarrow> x \<in> Un_tables ts k"
+apply (simp add: Un_tables_def)
+apply auto
+done
+
+lemma Un_tablesD [dest!]: "\<And>x. x \<in> Un_tables ts k \<Longrightarrow> \<exists>t. t \<in> ts \<and> x \<in> t k"
+apply (simp add: Un_tables_def)
+apply auto
+done
+
+lemma Un_tables_empty [simp]: "Un_tables {} = (\<lambda>k. {})"
+apply (unfold Un_tables_def)
+apply (simp (no_asm))
+done
+
+
+section "overrides"
+
+lemma empty_overrides_t [simp]: "(\<lambda>k. {}) \<oplus>\<oplus> m = m"
+apply (unfold overrides_t_def)
+apply (simp (no_asm))
+done
+lemma overrides_empty_t [simp]: "m \<oplus>\<oplus> (\<lambda>k. {}) = m"
+apply (unfold overrides_t_def)
+apply (simp (no_asm))
+done
+
+lemma overrides_t_Some_iff: 
+ "(x \<in> (s \<oplus>\<oplus> t) k) = (x \<in> t k \<or> t k = {} \<and> x \<in> s k)"
+by (simp add: overrides_t_def)
+
+lemmas overrides_t_SomeD = overrides_t_Some_iff [THEN iffD1, dest!]
+
+lemma overrides_t_right_empty [simp]: "n k = {} \<Longrightarrow> (m \<oplus>\<oplus> n) k = m k"  
+by (simp add: overrides_t_def)
+
+lemma overrides_t_find_right [simp]: "n k \<noteq> {} \<Longrightarrow> (m \<oplus>\<oplus> n) k = n k"  
+by (simp add: overrides_t_def)
+
+section "hiding entails"
+
+lemma hiding_entailsD: 
+  "\<lbrakk>t hiding s entails R; t k = Some x; s k = Some y\<rbrakk> \<Longrightarrow> R x y"
+by (simp add: hiding_entails_def)
+
+lemma empty_hiding_entails: "empty hiding s entails R"
+by (simp add: hiding_entails_def)
+
+lemma hiding_empty_entails: "t hiding empty entails R"
+by (simp add: hiding_entails_def)
+declare empty_hiding_entails [simp] hiding_empty_entails [simp]
+
+section "cond hiding entails"
+
+lemma cond_hiding_entailsD: 
+"\<lbrakk>t hiding s under C entails R; t k = Some x; s k = Some y; C x y\<rbrakk> \<Longrightarrow> R x y"
+by (simp add: cond_hiding_entails_def)
+
+lemma empty_cond_hiding_entails[simp]: "empty hiding s under C entails R"
+by (simp add: cond_hiding_entails_def)
+
+lemma cond_hiding_empty_entails[simp]: "t hiding empty under C entails R"
+by (simp add: cond_hiding_entails_def)
+
+lemma hidings_entailsD: "\<lbrakk>t hidings s entails R; x \<in> t k; y \<in> s k\<rbrakk> \<Longrightarrow> R x y"
+by (simp add: hidings_entails_def)
+
+lemma hidings_empty_entails: "t hidings (\<lambda>k. {}) entails R"
+apply (unfold hidings_entails_def)
+apply (simp (no_asm))
+done
+
+lemma empty_hidings_entails: 
+  "(\<lambda>k. {}) hidings s entails R"apply (unfold hidings_entails_def)
+by (simp (no_asm))
+declare empty_hidings_entails [intro!] hidings_empty_entails [intro!]
+
+
+
+(*###TO Map?*)
+consts
+  atleast_free :: "('a ~=> 'b) => nat => bool"
+primrec
+ "atleast_free m 0       = True"
+ atleast_free_Suc: 
+ "atleast_free m (Suc n) = (? a. m a = None & (!b. atleast_free (m(a|->b)) n))"
+
+lemma atleast_free_weaken [rule_format (no_asm)]: 
+  "!m. atleast_free m (Suc n) \<longrightarrow> atleast_free m n"
+apply (induct_tac "n")
+apply (simp (no_asm))
+apply clarify
+apply (simp (no_asm))
+apply (drule atleast_free_Suc [THEN iffD1])
+apply fast
+done
+
+lemma atleast_free_SucI: 
+"[| h a = None; !obj. atleast_free (h(a|->obj)) n |] ==> atleast_free h (Suc n)"
+by force
+
+declare fun_upd_apply [simp del]
+lemma atleast_free_SucD_lemma [rule_format (no_asm)]: 
+" !m a. m a = None --> (!c. atleast_free (m(a|->c)) n) -->  
+  (!b d. a ~= b --> atleast_free (m(b|->d)) n)"
+apply (induct_tac "n")
+apply  auto
+apply (rule_tac x = "a" in exI)
+apply  (rule conjI)
+apply  (force simp add: fun_upd_apply)
+apply (erule_tac V = "m a = None" in thin_rl)
+apply clarify
+apply (subst fun_upd_twist)
+apply  (erule not_sym)
+apply (rename_tac "ba")
+apply (drule_tac x = "ba" in spec)
+apply clarify
+apply (tactic "smp_tac 2 1")
+apply (erule (1) notE impE)
+apply (case_tac "aa = b")
+apply fast+
+done
+declare fun_upd_apply [simp]
+
+lemma atleast_free_SucD [rule_format (no_asm)]: "atleast_free h (Suc n) ==> atleast_free (h(a|->b)) n"
+apply auto
+apply (case_tac "aa = a")
+apply auto
+apply (erule atleast_free_SucD_lemma)
+apply auto
+done
+
+declare atleast_free_Suc [simp del]
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Term.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,231 @@
+(*  Title:      isabelle/Bali/Term.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+
+header {* Java expressions and statements *}
+
+theory Term = Value:
+
+text {*
+design issues:
+\begin{itemize}
+\item invocation frames for local variables could be reduced to special static
+  objects (one per method). This would reduce redundancy, but yield a rather
+  non-standard execution model more difficult to understand.
+\item method bodies separated from calls to handle assumptions in axiomat. 
+  semantics
+  NB: Body is intended to be in the environment of the called method.
+\item class initialization is regarded as (auxiliary) statement 
+      (required for AxSem)
+\item result expression of method return is handled by a special result variable
+  result variable is treated uniformly with local variables
+  \begin{itemize}
+  \item[+] welltypedness and existence of the result/return expression is 
+           ensured without extra efford
+  \end{itemize}
+\end{itemize}
+
+simplifications:
+\begin{itemize}
+\item expression statement allowed for any expression
+\item no unary, binary, etc, operators
+\item This  is modeled as a special non-assignable local variable
+\item Super is modeled as a general expression with the same value as This
+\item access to field x in current class via This.x
+\item NewA creates only one-dimensional arrays;
+  initialization of further subarrays may be simulated with nested NewAs
+\item The 'Lit' constructor is allowed to contain a reference value.
+  But this is assumed to be prohibited in the input language, which is enforced
+  by the type-checking rules.
+\item a call of a static method via a type name may be simulated by a dummy 
+      variable
+\item no nested blocks with inner local variables
+\item no synchronized statements
+\item no secondary forms of if, while (e.g. no for) (may be easily simulated)
+\item no switch (may be simulated with if)
+\item the @{text try_catch_finally} statement is divided into the 
+      @{text try_catch} statement 
+      and a finally statement, which may be considered as try..finally with 
+      empty catch
+\item the @{text try_catch} statement has exactly one catch clause; 
+      multiple ones can be
+  simulated with instanceof
+\item the compiler is supposed to add the annotations {@{text _}} during 
+      type-checking. This
+  transformation is left out as its result is checked by the type rules anyway
+\end{itemize}
+*}
+
+datatype inv_mode                  (* invocation mode for method calls *)
+	= Static                   (* static *)
+	| SuperM                   (* super  *)
+	| IntVir                   (* interface or virtual *)
+
+record  sig =            (* signature of a method, cf. 8.4.2  *)
+	  name ::"mname"      (* acutally belongs to Decl.thy *)
+          parTs::"ty list"        
+
+translations
+  "sig" <= (type) "\<lparr>name::mname,parTs::ty list\<rparr>"
+  "sig" <= (type) "\<lparr>name::mname,parTs::ty list,\<dots>::'a\<rparr>"
+
+datatype jump
+        = Break label (* break *)
+        | Cont label  (* continue *)
+        | Ret         (* return from method *)
+
+datatype var
+	= LVar                  lname(* local variable (incl. parameters) *)
+        | FVar qtname bool expr vname(*class field*)("{_,_}_.._"[10,10,85,99]90)
+	| AVar        expr expr      (* array component *) ("_.[_]"[90,10   ]90)
+
+and expr
+	= NewC qtname              (* class instance creation *)
+	| NewA ty expr             (* array creation *) ("New _[_]"[99,10   ]85)
+	| Cast ty expr             (* type cast  *)
+	| Inst expr ref_ty         (* instanceof *)     ("_ InstOf _"[85,99] 85)
+	| Lit  val                 (* literal value, references not allowed *)
+	| Super                    (* special Super keyword *)
+	| Acc  var                 (* variable access *)
+	| Ass  var expr            (* variable assign *) ("_:=_"   [90,85   ]85)
+	| Cond expr expr expr      (* conditional *)  ("_ ? _ : _" [85,85,80]80)
+        | Call ref_ty inv_mode expr mname "(ty list)" (* method call *)
+                  "(expr list)" ("{_,_}_\<cdot>_'( {_}_')"[10,10,85,99,10,10]85)
+        | Methd qtname sig          (*   (folded) method (see below) *)
+        | Body qtname stmt          (* (unfolded) method body *)
+and  stmt
+	= Skip                     (* empty      statement *)
+	| Expr  expr               (* expression statement *)
+        | Lab   label stmt         ("_\<bullet> _"(* labeled statement*)[      99,66]66)
+                                   (* handles break *)
+	| Comp  stmt stmt          ("_;; _"                     [      66,65]65)
+	| If_   expr stmt stmt     ("If'(_') _ Else _"          [   80,79,79]70)
+	| Loop  label expr stmt    ("_\<bullet> While'(_') _"           [   99,80,79]70)
+        | Do jump                  (* break, continue, return *)
+	| Throw expr
+        | TryC  stmt
+	        qtname vname stmt   ("Try _ Catch'(_ _') _"     [79,99,80,79]70)
+	| Fin   stmt stmt          ("_ Finally _"               [      79,79]70)
+	| Init  qtname              (* class initialization *)
+
+
+text {*
+The expressions Methd and Body are artificial program constructs, in the
+sense that they are not used to define a concrete Bali program. In the 
+evaluation semantic definition they are "generated on the fly" to decompose 
+the task to define the behaviour of the Call expression. They are crucial 
+for the axiomatic semantics to give a syntactic hook to insert 
+some assertions (cf. AxSem.thy, Eval.thy).
+Also the Init statement (to initialize a class on its first use) is inserted 
+in various places by the evaluation semantics.   
+*}
+ 
+types "term" = "(expr+stmt, var, expr list) sum3"
+translations
+  "sig"   <= (type) "mname \<times> ty list"
+  "var"   <= (type) "Term.var"
+  "expr"  <= (type) "Term.expr"
+  "stmt"  <= (type) "Term.stmt"
+  "term"  <= (type) "(expr+stmt, var, expr list) sum3"
+
+syntax
+  
+  this    :: expr
+  LAcc    :: "vname \<Rightarrow>         expr" ("!!")
+  LAss    :: "vname \<Rightarrow> expr \<Rightarrow> stmt" ("_:==_" [90,85] 85)
+  Return  :: "expr \<Rightarrow> stmt"
+  StatRef :: "ref_ty \<Rightarrow> expr"
+
+translations
+  
+ "this"       == "Acc (LVar This)"
+ "!!v"        == "Acc (LVar (EName (VNam v)))"
+ "v:==e"      == "Expr (Ass (LVar (EName (VNam  v))) e)"
+ "Return e"   == "Expr (Ass (LVar (EName Res)) e);; Do Ret" 
+                                                   (* Res := e;; Do Ret *)
+ "StatRef rt" == "Cast (RefT rt) (Lit Null)"
+  
+constdefs
+
+  is_stmt :: "term \<Rightarrow> bool"
+ "is_stmt t \<equiv> \<exists>c. t=In1r c"
+
+ML {*
+bind_thms ("is_stmt_rews", sum3_instantiate (thm "is_stmt_def"));
+*}
+
+declare is_stmt_rews [simp]
+
+
+(* ############# Just testing syntax *)
+(** unfortunately cannot simply instantiate tnam **)
+(*
+datatype tnam_  = HasFoo_ | Base_ | Ext_
+datatype vnam_  = arr_ | vee_ | z_ | e_
+datatype label_ = lab1_
+
+consts
+
+  tnam_ :: "tnam_  \<Rightarrow> tnam"
+  vnam_ :: "vnam_  \<Rightarrow> vname"
+  label_:: "label_ \<Rightarrow> label"
+axioms  
+
+  inj_tnam_  [simp]: "(tnam_  x = tnam_  y) = (x = y)"
+  inj_vnam_  [simp]: "(vnam_  x = vnam_  y) = (x = y)"
+  inj_label_ [simp]: "(label_ x = label_ y) = (x = y)"
+  
+  
+  surj_tnam_:  "\<exists>m. n = tnam_  m"
+  surj_vnam_:  "\<exists>m. n = vnam_  m"
+  surj_label_:" \<exists>m. n = label_ m"
+
+syntax
+
+  HasFoo :: qtname
+  Base   :: qtname
+  Ext    :: qtname
+  arr :: ename
+  vee :: ename
+  z   :: ename
+  e   :: ename
+  lab1:: label
+
+consts
+  
+  foo    :: mname
+translations
+
+  "HasFoo" == "\<lparr>pid=java_lang,tid=TName (tnam_ HasFoo_)\<rparr>"
+  "Base"   == "\<lparr>pid=java_lang,tid=TName (tnam_ Base_)\<rparr>"
+  "Ext"    == "\<lparr>pid=java_lang,tid=TName (tnam_ Ext_)\<rparr>"
+  "arr"    ==        "(vnam_ arr_)"
+  "vee"    ==        "(vnam_ vee_)"
+  "z"      ==        "(vnam_ z_)"
+  "e"      ==        "(vnam_ e_)"
+  "lab1"   ==        "label_ lab1_"
+
+constdefs test::stmt
+"test \<equiv>
+(lab1@ While(Acc  
+      (Acc ({Base,True}StatRef (ClassT Object).arr).[Lit (Intg #2)])) Skip)"
+
+consts
+ pTs :: "ty list"
+   
+constdefs 
+
+test1::stmt
+"test1 \<equiv> 
+  Expr({ClassT Base,IntVir}!!e\<cdot>foo({pTs}[Lit Null]))"
+
+
+
+constdefs test::stmt
+"test \<equiv>
+(lab1\<cdot> While(Acc 
+      (Acc ({Base,True}StatRef (ClassT Object)..arr).[Lit (Intg #2)])) Skip)"
+*)
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Trans.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,144 @@
+(*  Title:      isabelle/Bali/Trans.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+
+Operational transition (small-step) semantics of the 
+execution of Java expressions and statements
+
+PRELIMINARY!!!!!!!!
+
+*)
+
+Trans = Eval +
+
+consts
+  texpr_tstmt	:: "prog \<Rightarrow> (((expr \<times> state) \<times> (expr \<times> state)) +
+		            ((stmt \<times> state) \<times> (stmt \<times> state))) set"
+
+syntax (symbols)
+  texpr :: "[prog, expr \<times> state, expr \<times> state] \<Rightarrow> bool "("_\<turnstile>_ \<rightarrow>1 _"[61,82,82] 81)
+  tstmt :: "[prog, stmt \<times> state, stmt \<times> state] \<Rightarrow> bool "("_\<turnstile>_ \<mapsto>1 _"[61,82,82] 81)
+  Ref   :: "loc \<Rightarrow> expr"
+
+translations
+
+  "G\<turnstile>e_s \<rightarrow>1 ex_s'" == "Inl (e_s, ex_s') \<in> texpr_tstmt G"
+  "G\<turnstile>s_s \<mapsto>1 s'_s'" == "Inr (s_s, s'_s') \<in> texpr_tstmt G"
+  "Ref a" == "Lit (Addr a)"
+
+inductive "texpr_tstmt G" intrs 
+
+(* evaluation of expression *)
+  (* cf. 15.5 *)
+  XcptE	"\<lbrakk>\<forall>v. e \<noteq> Lit v\<rbrakk> \\<Longrightarrow>
+				  G\<turnstile>(e,Some xc,s) \<rightarrow>1 (Lit arbitrary,Some xc,s)"
+
+  (* cf. 15.8.1 *)
+  NewC	"\<lbrakk>new_Addr (heap s) = Some (a,x);
+	  s' = assign (hupd[a\<mapsto>init_Obj G C]s) (x,s)\<rbrakk> \\<Longrightarrow>
+				G\<turnstile>(NewC C,None,s) \<rightarrow>1 (Ref a,s')"
+
+  (* cf. 15.9.1 *)
+  NewA1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
+			      G\<turnstile>(New T[e],None,s) \<rightarrow>1 (New T[e'],s')"
+  NewA	"\<lbrakk>i = the_Intg i'; new_Addr (heap s) = Some (a, x);
+	  s' = assign (hupd[a\<mapsto>init_Arr T i]s)(raise_if (i<#0) NegArrSize x,s)\<rbrakk> \\<Longrightarrow>
+			G\<turnstile>(New T[Lit i'],None,s) \<rightarrow>1 (Ref a,s')"
+  (* cf. 15.15 *)
+  Cast1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
+			      G\<turnstile>(Cast T e,None,s) \<rightarrow>1 (Cast T e',s')"
+  Cast	"\<lbrakk>x'= raise_if (\<questiondown>G,s\<turnstile>v fits T) ClassCast None\<rbrakk> \\<Longrightarrow>
+		        G\<turnstile>(Cast T (Lit v),None,s) \<rightarrow>1 (Lit v,x',s)"
+
+  (* cf. 15.7.1 *)
+(*Lit				"G\<turnstile>(Lit v,None,s) \<rightarrow>1 (Lit v,None,s)"*)
+
+  (* cf. 15.13.1, 15.2 *)
+  LAcc	"\<lbrakk>v = the (locals s vn)\<rbrakk> \\<Longrightarrow>
+			       G\<turnstile>(LAcc vn,None,s) \<rightarrow>1 (Lit v,None,s)"
+
+  (* cf. 15.25.1 *)
+  LAss1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
+				 G\<turnstile>(f vn:=e,None,s) \<rightarrow>1 (vn:=e',s')"
+  LAss			    "G\<turnstile>(f vn:=Lit v,None,s) \<rightarrow>1 (Lit v,None,lupd[vn\<mapsto>v]s)"
+
+  (* cf. 15.10.1, 15.2 *)
+  FAcc1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
+			       G\<turnstile>({T}e..fn,None,s) \<rightarrow>1 ({T}e'..fn,s')"
+  FAcc	"\<lbrakk>v = the (snd (the_Obj (heap s (the_Addr a'))) (fn,T))\<rbrakk> \\<Longrightarrow>
+			  G\<turnstile>({T}Lit a'..fn,None,s) \<rightarrow>1 (Lit v,np a' None,s)"
+
+  (* cf. 15.25.1 *)
+  FAss1	"\<lbrakk>G\<turnstile>(e1,None,s) \<rightarrow>1 (e1',s')\<rbrakk> \\<Longrightarrow>
+			  G\<turnstile>(f ({T}e1..fn):=e2,None,s) \<rightarrow>1 (f({T}e1'..fn):=e2,s')"
+  FAss2	"\<lbrakk>G\<turnstile>(e2,np a' None,s) \<rightarrow>1 (e2',s')\<rbrakk> \\<Longrightarrow>
+		      G\<turnstile>(f({T}Lit a'..fn):=e2,None,s) \<rightarrow>1 (f({T}Lit a'..fn):=e2',s')"
+  FAss	"\<lbrakk>a = the_Addr a'; (c,fs) = the_Obj (heap s a);
+	  s'= assign (hupd[a\<mapsto>Obj c (fs[(fn,T)\<mapsto>v])]s) (None,s)\<rbrakk> \\<Longrightarrow>
+		   G\<turnstile>(f({T}Lit a'..fn):=Lit v,None,s) \<rightarrow>1 (Lit v,s')"
+
+
+
+
+
+  (* cf. 15.12.1 *)
+  AAcc1	"\<lbrakk>G\<turnstile>(e1,None,s) \<rightarrow>1 (e1',s')\<rbrakk> \\<Longrightarrow>
+				G\<turnstile>(e1[e2],None,s) \<rightarrow>1 (e1'[e2],s')"
+  AAcc2	"\<lbrakk>G\<turnstile>(e2,None,s) \<rightarrow>1 (e2',s')\<rbrakk> \\<Longrightarrow>
+			    G\<turnstile>(Lit a'[e2],None,s) \<rightarrow>1 (Lit a'[e2'],s')"
+  AAcc	"\<lbrakk>vo = snd (the_Arr (heap s (the_Addr a'))) (the_Intg i');
+	  x' = raise_if (vo = None) IndOutBound (np a' None)\<rbrakk> \\<Longrightarrow>
+			G\<turnstile>(Lit a'[Lit i'],None,s) \<rightarrow>1 (Lit (the vo),x',s)"
+
+
+  (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *)
+  Call1	"\<lbrakk>G\<turnstile>(e,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
+			  G\<turnstile>(e..mn({pT}p),None,s) \<rightarrow>1 (e'..mn({pT}p),s')"
+  Call2	"\<lbrakk>G\<turnstile>(p,None,s) \<rightarrow>1 (p',s')\<rbrakk> \\<Longrightarrow>
+		     G\<turnstile>(Lit a'..mn({pT}p),None,s) \<rightarrow>1 (Lit a'..mn({pT}p'),s')"
+  Call	"\<lbrakk>a = the_Addr a'; (md,(pn,rT),lvars,blk,res) = 
+ 			   the (cmethd G (fst (the_Obj (h a))) (mn,pT))\<rbrakk> \\<Longrightarrow>
+	    G\<turnstile>(Lit a'..mn({pT}Lit pv),None,(h,l)) \<rightarrow>1 
+      (Body blk res l,np a' x,(h,init_vals lvars[This\<mapsto>a'][Super\<mapsto>a'][pn\<mapsto>pv]))"
+  Body1	"\<lbrakk>G\<turnstile>(s0,None,s) \<mapsto>1 (s0',s')\<rbrakk> \\<Longrightarrow>
+		   G\<turnstile>(Body s0    e      l,None,s) \<rightarrow>1 (Body s0'  e  l,s')"
+  Body2	"\<lbrakk>G\<turnstile>(e ,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
+		   G\<turnstile>(Body Skip  e      l,None,s) \<rightarrow>1 (Body Skip e' l,s')"
+  Body		  "G\<turnstile>(Body Skip (Lit v) l,None,s) \<rightarrow>1 (Lit v,None,(heap s,l))"
+
+(* execution of statements *)
+
+  (* cf. 14.1 *)
+  XcptS	"\<lbrakk>s0 \<noteq> Skip\<rbrakk> \\<Longrightarrow>
+				 G\<turnstile>(s0,Some xc,s) \<mapsto>1 (Skip,Some xc,s)"
+
+  (* cf. 14.5 *)
+(*Skip	 			 "G\<turnstile>(Skip,None,s) \<mapsto>1 (Skip,None,s)"*)
+
+  (* cf. 14.2 *)
+  Comp1	"\<lbrakk>G\<turnstile>(s1,None,s) \<mapsto>1 (s1',s')\<rbrakk> \\<Longrightarrow>
+			       G\<turnstile>(s1;; s2,None,s) \<mapsto>1 (s1';; s2,s')"
+  Comp			    "G\<turnstile>(Skip;; s2,None,s) \<mapsto>1 (s2,None,s)"
+
+
+
+
+
+
+  (* cf. 14.7 *)
+  Expr1	"\<lbrakk>G\<turnstile>(e ,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
+				G\<turnstile>(Expr e,None,s) \<mapsto>1 (Expr e',s')"
+  Expr			 "G\<turnstile>(Expr (Lit v),None,s) \<mapsto>1 (Skip,None,s)"
+
+  (* cf. 14.8.2 *)
+  If1	"\<lbrakk>G\<turnstile>(e ,None,s) \<rightarrow>1 (e',s')\<rbrakk> \\<Longrightarrow>
+		      G\<turnstile>(If(e) s1 Else s2,None,s) \<mapsto>1 (If(e') s1 Else s2,s')"
+  If		 "G\<turnstile>(If(Lit v) s1 Else s2,None,s) \<mapsto>1 
+		    (if the_Bool v then s1 else s2,None,s)"
+
+  (* cf. 14.10, 14.10.1 *)
+  Loop			  "G\<turnstile>(While(e) s0,None,s) \<mapsto>1 
+			     (If(e) (s0;; While(e) s0) Else Skip,None,s)"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Type.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,59 @@
+(*  Title:      isabelle/Bali/Type.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+
+header {* Java types *}
+
+theory Type = Name:
+
+text {*
+simplifications:
+\begin{itemize}
+\item only the most important primitive types
+\item the null type is regarded as reference type
+\end{itemize}
+*}
+
+datatype prim_ty       	(* primitive type, cf. 4.2 *)
+	= Void    	(* 'result type' of void methods *)
+	| Boolean
+	| Integer
+
+
+datatype ref_ty		(* reference type, cf. 4.3 *)
+	= NullT		(* null type, cf. 4.1 *)
+	| IfaceT qtname	(* interface type *)
+	| ClassT qtname	(* class type *)
+	| ArrayT ty	(* array type *)
+
+and ty	    		(* any type, cf. 4.1 *)
+	= PrimT prim_ty	(* primitive type *)
+	| RefT  ref_ty	(* reference type *)
+
+translations
+  "prim_ty" <= (type) "Type.prim_ty"
+  "ref_ty"  <= (type) "Type.ref_ty"
+  "ty"      <= (type) "Type.ty"
+
+syntax
+	 NT	:: "       \<spacespace> ty"
+	 Iface	:: "qtname  \<Rightarrow> ty"
+	 Class	:: "qtname  \<Rightarrow> ty"
+	 Array	:: "ty     \<Rightarrow> ty"	("_.[]" [90] 90)
+
+translations
+	"NT"      == "RefT   NullT"
+	"Iface I" == "RefT (IfaceT I)"
+	"Class C" == "RefT (ClassT C)"
+	"T.[]"    == "RefT (ArrayT T)"
+
+constdefs
+  the_Class :: "ty \<Rightarrow> qtname"
+ "the_Class T \<equiv> \<epsilon>C. T = Class C" (** primrec not possible here **)
+ 
+lemma the_Class_eq [simp]: "the_Class (Class C)= C"
+by (auto simp add: the_Class_def)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/TypeRel.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,667 @@
+(*  Title:      isabelle/Bali/TypeRel.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* The relations between Java types *}
+
+theory TypeRel = Decl:
+
+text {*
+simplifications:
+\begin{itemize}
+\item subinterface, subclass and widening relation includes identity
+\end{itemize}
+improvements over Java Specification 1.0:
+\begin{itemize}
+\item narrowing reference conversion also in cases where the return types of a 
+      pair of methods common to both types are in widening (rather identity) 
+      relation
+\item one could add similar constraints also for other cases
+\end{itemize}
+design issues:
+\begin{itemize}
+\item the type relations do not require @{text is_type} for their arguments
+\item the subint1 and subcls1 relations imply @{text is_iface}/@{text is_class}
+      for their first arguments, which is required for their finiteness
+\end{itemize}
+*}
+
+consts
+
+(*subint1, in Decl.thy*)                     (* direct subinterface       *)
+(*subint , by translation*)                  (* subinterface (+ identity) *)
+(*subcls1, in Decl.thy*)                     (* direct subclass           *)
+(*subcls , by translation*)                  (*        subclass           *)
+(*subclseq, by translation*)                 (* subclass + identity       *)
+  implmt1   :: "prog \<Rightarrow> (qtname \<times> qtname) set" (* direct implementation *)
+  implmt    :: "prog \<Rightarrow> (qtname \<times> qtname) set" (*        implementation *)
+  widen     :: "prog \<Rightarrow> (ty    \<times> ty   ) set" (*        widening       *)
+  narrow    :: "prog \<Rightarrow> (ty    \<times> ty   ) set" (*        narrowing      *)
+  cast     :: "prog \<Rightarrow> (ty    \<times> ty   ) set"  (*        casting        *)
+
+syntax
+
+ "@subint1" :: "prog => [qtname, qtname] => bool" ("_|-_<:I1_" [71,71,71] 70)
+ "@subint"  :: "prog => [qtname, qtname] => bool" ("_|-_<=:I _"[71,71,71] 70)
+ (* Defined in Decl.thy:
+ "@subcls1" :: "prog => [qtname, qtname] => bool" ("_|-_<:C1_" [71,71,71] 70)
+ "@subclseq":: "prog => [qtname, qtname] => bool" ("_|-_<=:C _"[71,71,71] 70)
+ "@subcls"  :: "prog => [qtname, qtname] => bool" ("_|-_<:C _"[71,71,71] 70)
+ *)
+ "@implmt1" :: "prog => [qtname, qtname] => bool" ("_|-_~>1_"  [71,71,71] 70)
+ "@implmt"  :: "prog => [qtname, qtname] => bool" ("_|-_~>_"   [71,71,71] 70)
+ "@widen"   :: "prog => [ty   , ty   ] => bool" ("_|-_<=:_"  [71,71,71] 70)
+ "@narrow"  :: "prog => [ty   , ty   ] => bool" ("_|-_:>_"   [71,71,71] 70)
+ "@cast"    :: "prog => [ty   , ty   ] => bool" ("_|-_<=:? _"[71,71,71] 70)
+
+syntax (symbols)
+
+  "@subint1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>I1_"  [71,71,71] 70)
+  "@subint"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>I _"  [71,71,71] 70)
+  (* Defined in Decl.thy:
+\  "@subcls1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>\<^sub>C\<^sub>1_"  [71,71,71] 70)
+  "@subclseq":: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>\<^sub>C _"  [71,71,71] 70)
+  "@subcls"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>\<^sub>C _"  [71,71,71] 70)
+  *)
+  "@implmt1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<leadsto>1_"  [71,71,71] 70)
+  "@implmt"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<leadsto>_"   [71,71,71] 70)
+  "@widen"   :: "prog \<Rightarrow> [ty   , ty   ] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>_"    [71,71,71] 70)
+  "@narrow"  :: "prog \<Rightarrow> [ty   , ty   ] \<Rightarrow> bool" ("_\<turnstile>_\<succ>_"    [71,71,71] 70)
+  "@cast"    :: "prog \<Rightarrow> [ty   , ty   ] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>? _"  [71,71,71] 70)
+
+translations
+
+	"G\<turnstile>I \<prec>I1 J" == "(I,J) \<in> subint1 G"
+	"G\<turnstile>I \<preceq>I  J" == "(I,J) \<in>(subint1 G)^*" (* cf. 9.1.3 *)
+  	(* Defined in Decl.thy:
+        "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" == "(C,D) \<in> subcls1 G"
+	"G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" 
+	*)
+        "G\<turnstile>C \<leadsto>1 I" == "(C,I) \<in> implmt1 G"
+	"G\<turnstile>C \<leadsto>  I" == "(C,I) \<in> implmt  G"
+	"G\<turnstile>S \<preceq>   T" == "(S,T) \<in> widen   G"
+	"G\<turnstile>S \<succ>   T" == "(S,T) \<in> narrow  G"
+	"G\<turnstile>S \<preceq>?  T" == "(S,T) \<in> cast    G"
+
+
+section "subclass and subinterface relations"
+
+  (* direct subinterface in Decl.thy, cf. 9.1.3 *)
+  (* direct subclass     in Decl.thy, cf. 8.1.3 *)
+
+lemmas subcls_direct = subcls1I [THEN r_into_rtrancl, standard]
+
+lemma subcls_direct1:
+ "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<preceq>\<^sub>C D"
+apply (auto dest: subcls_direct)
+done
+
+lemma subcls1I1:
+ "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>\<^sub>C\<^sub>1 D"
+apply (auto dest: subcls1I)
+done
+
+lemma subcls_direct2:
+ "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>\<^sub>C D"
+apply (auto dest: subcls1I1)
+done
+
+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"
+by (blast intro: rtrancl_trans)
+
+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"
+by (blast intro: trancl_trans)
+
+lemma SXcpt_subcls_Throwable_lemma: 
+"\<lbrakk>class G (SXcpt xn) = Some xc;
+  super xc = (if xn = Throwable then Object else  SXcpt Throwable)\<rbrakk> 
+\<Longrightarrow> G\<turnstile>SXcpt xn\<preceq>\<^sub>C SXcpt Throwable"
+apply (case_tac "xn = Throwable")
+apply  simp_all
+apply (drule subcls_direct)
+apply (auto dest: sym)
+done
+
+lemma subcls_ObjectI: "\<lbrakk>is_class G C; ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>C\<preceq>\<^sub>C Object"
+apply (erule ws_subcls1_induct)
+apply clarsimp
+apply (case_tac "C = Object")
+apply  (fast intro: r_into_rtrancl [THEN rtrancl_trans])+
+done
+
+lemma subclseq_ObjectD [dest!]: "G\<turnstile>Object\<preceq>\<^sub>C C \<Longrightarrow> C = Object"
+apply (erule rtrancl_induct)
+apply  (auto dest: subcls1D)
+done
+
+lemma subcls_ObjectD [dest!]: "G\<turnstile>Object\<prec>\<^sub>C C \<Longrightarrow> False"
+apply (erule trancl_induct)
+apply  (auto dest: subcls1D)
+done
+
+lemma subcls_ObjectI1 [intro!]: 
+ "\<lbrakk>C \<noteq> Object;is_class G C;ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>C \<prec>\<^sub>C Object" 
+apply (drule (1) subcls_ObjectI)
+apply (auto intro: rtrancl_into_trancl3)
+done
+
+lemma subcls_is_class: "(C,D) \<in> (subcls1 G)^+ \<Longrightarrow> is_class G C"
+apply (erule trancl_trans_induct)
+apply (auto dest!: subcls1D)
+done
+
+lemma subcls_is_class2 [rule_format (no_asm)]: 
+ "G\<turnstile>C\<preceq>\<^sub>C D \<Longrightarrow> is_class G D \<longrightarrow> is_class G C"
+apply (erule rtrancl_induct)
+apply (drule_tac [2] subcls1D)
+apply  auto
+done
+
+lemma single_inheritance: 
+"\<lbrakk>G\<turnstile>A \<prec>\<^sub>C\<^sub>1 B; G\<turnstile>A \<prec>\<^sub>C\<^sub>1 C\<rbrakk> \<Longrightarrow> B = C"
+by (auto simp add: subcls1_def)
+  
+lemma subcls_compareable:
+"\<lbrakk>G\<turnstile>A \<preceq>\<^sub>C X; G\<turnstile>A \<preceq>\<^sub>C Y 
+ \<rbrakk> \<Longrightarrow> G\<turnstile>X \<preceq>\<^sub>C Y \<or> G\<turnstile>Y \<preceq>\<^sub>C X"
+by (rule triangle_lemma)  (auto intro: single_inheritance) 
+
+lemma subcls1_irrefl: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D; ws_prog G \<rbrakk>
+ \<Longrightarrow> C \<noteq> D"
+proof 
+  assume ws: "ws_prog G" and
+    subcls1: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" and
+     eq_C_D: "C=D"
+  from subcls1 obtain c 
+    where
+      neq_C_Object: "C\<noteq>Object" and
+              clsC: "class G C = Some c" and
+           super_c: "super c = D"
+    by (auto simp add: subcls1_def)
+  with super_c subcls1 eq_C_D 
+  have subcls_super_c_C: "G\<turnstile>super c \<prec>\<^sub>C C"
+    by auto
+  from ws clsC neq_C_Object 
+  have "\<not> G\<turnstile>super c \<prec>\<^sub>C C"
+    by (auto dest: ws_prog_cdeclD)
+  from this subcls_super_c_C 
+  show "False"
+    by (rule notE)
+qed
+
+lemma no_subcls_Object: "G\<turnstile>C \<prec>\<^sub>C D \<Longrightarrow> C \<noteq> Object"
+by (erule converse_trancl_induct) (auto dest: subcls1D)
+
+lemma subcls_acyclic: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D; ws_prog G\<rbrakk> \<Longrightarrow> \<not> G\<turnstile>D \<prec>\<^sub>C C"
+proof -
+  assume         ws: "ws_prog G"
+  assume subcls_C_D: "G\<turnstile>C \<prec>\<^sub>C D"
+  then show ?thesis
+  proof (induct rule: converse_trancl_induct)
+    fix C
+    assume subcls1_C_D: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D"
+    then obtain c  where
+        "C\<noteq>Object" and
+        "class G C = Some c" and
+        "super c = D"
+      by (auto simp add: subcls1_def)
+    with ws 
+    show "\<not> G\<turnstile>D \<prec>\<^sub>C C"
+      by (auto dest: ws_prog_cdeclD)
+  next
+    fix C Z
+    assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 Z" and
+            subcls_Z_D: "G\<turnstile>Z \<prec>\<^sub>C D" and
+           nsubcls_D_Z: "\<not> G\<turnstile>D \<prec>\<^sub>C Z"
+    show "\<not> G\<turnstile>D \<prec>\<^sub>C C"
+    proof
+      assume subcls_D_C: "G\<turnstile>D \<prec>\<^sub>C C"
+      show "False"
+      proof -
+	from subcls_D_C subcls1_C_Z
+	have "G\<turnstile>D \<prec>\<^sub>C Z"
+	  by (auto dest: r_into_trancl trancl_trans)
+	with nsubcls_D_Z
+	show ?thesis
+	  by (rule notE)
+      qed
+    qed
+  qed  
+qed
+
+lemma subclseq_cases [consumes 1, case_names Eq Subcls]:
+ "\<lbrakk>G\<turnstile>C \<preceq>\<^sub>C D; C = D \<Longrightarrow> P; G\<turnstile>C \<prec>\<^sub>C D \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+by (blast intro: rtrancl_cases)
+
+lemma subclseq_acyclic: 
+ "\<lbrakk>G\<turnstile>C \<preceq>\<^sub>C D; G\<turnstile>D \<preceq>\<^sub>C C; ws_prog G\<rbrakk> \<Longrightarrow> C=D"
+by (auto elim: subclseq_cases dest: subcls_acyclic)
+
+lemma subcls_irrefl: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D; ws_prog G\<rbrakk>
+ \<Longrightarrow> C \<noteq> D"
+proof -
+  assume     ws: "ws_prog G"
+  assume subcls: "G\<turnstile>C \<prec>\<^sub>C D"
+  then show ?thesis
+  proof (induct rule: converse_trancl_induct)
+    fix C
+    assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D"
+    with ws 
+    show "C\<noteq>D" 
+      by (blast dest: subcls1_irrefl)
+  next
+    fix C Z
+    assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 Z" and
+            subcls_Z_D: "G\<turnstile>Z \<prec>\<^sub>C D" and
+               neq_Z_D: "Z \<noteq> D"
+    show "C\<noteq>D"
+    proof
+      assume eq_C_D: "C=D"
+      show "False"
+      proof -
+	from subcls1_C_Z eq_C_D 
+	have "G\<turnstile>D \<prec>\<^sub>C Z"
+	  by (auto)
+	also
+	from subcls_Z_D ws   
+	have "\<not> G\<turnstile>D \<prec>\<^sub>C Z"
+	  by (rule subcls_acyclic)
+	ultimately 
+	show ?thesis 
+	  by - (rule notE)
+      qed
+    qed
+  qed
+qed
+
+lemma invert_subclseq:
+"\<lbrakk>G\<turnstile>C \<preceq>\<^sub>C D; ws_prog G\<rbrakk>
+ \<Longrightarrow> \<not> G\<turnstile>D \<prec>\<^sub>C C"
+proof -
+  assume       ws: "ws_prog G" and
+     subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D"
+  show ?thesis
+  proof (cases "D=C")
+    case True
+    with ws 
+    show ?thesis 
+      by (auto dest: subcls_irrefl)
+  next
+    case False
+    with subclseq_C_D 
+    have "G\<turnstile>C \<prec>\<^sub>C D"
+      by (blast intro: rtrancl_into_trancl3) 
+    with ws 
+    show ?thesis 
+      by (blast dest: subcls_acyclic)
+  qed
+qed
+
+lemma invert_subcls:
+"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D; ws_prog G\<rbrakk>
+ \<Longrightarrow> \<not> G\<turnstile>D \<preceq>\<^sub>C C"
+proof -
+  assume        ws: "ws_prog G" and
+        subcls_C_D: "G\<turnstile>C \<prec>\<^sub>C D"
+  then 
+  have nsubcls_D_C: "\<not> G\<turnstile>D \<prec>\<^sub>C C"
+    by (blast dest: subcls_acyclic)
+  show ?thesis
+  proof
+    assume "G\<turnstile>D \<preceq>\<^sub>C C"
+    then show "False"
+    proof (cases rule: subclseq_cases)
+      case Eq
+      with ws subcls_C_D
+      show ?thesis 
+	by (auto dest: subcls_irrefl)
+    next
+      case Subcls
+      with nsubcls_D_C
+      show ?thesis
+	by blast
+    qed
+  qed
+qed
+
+lemma subcls_superD:
+ "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D; class G C = Some c\<rbrakk> \<Longrightarrow> G\<turnstile>(super c) \<preceq>\<^sub>C D"
+proof -
+  assume       clsC: "class G C = Some c"
+  assume subcls_C_C: "G\<turnstile>C \<prec>\<^sub>C D"
+  then obtain S where 
+                  "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" and
+    subclseq_S_D: "G\<turnstile>S \<preceq>\<^sub>C D"
+    by (blast dest: tranclD)
+  with clsC 
+  have "S=super c" 
+    by (auto dest: subcls1D)
+  with subclseq_S_D show ?thesis by simp
+qed
+ 
+
+lemma subclseq_superD:
+ "\<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"
+proof -
+  assume neq_C_D: "C\<noteq>D"
+  assume    clsC: "class G C = Some c"
+  assume subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" 
+  then show ?thesis
+  proof (cases rule: subclseq_cases)
+    case Eq with neq_C_D show ?thesis by contradiction
+  next
+    case Subcls
+    with clsC show ?thesis by (blast dest: subcls_superD)
+  qed
+qed
+
+section "implementation relation"
+
+defs
+  (* direct implementation, cf. 8.1.3 *)
+implmt1_def:"implmt1 G\<equiv>{(C,I). C\<noteq>Object \<and> (\<exists>c\<in>class G C: I\<in>set (superIfs c))}"
+
+lemma implmt1D: "G\<turnstile>C\<leadsto>1I \<Longrightarrow> C\<noteq>Object \<and> (\<exists>c\<in>class G C: I\<in>set (superIfs c))"
+apply (unfold implmt1_def)
+apply auto
+done
+
+
+inductive "implmt G" intros                    (* cf. 8.1.4 *)
+
+  direct:         "G\<turnstile>C\<leadsto>1J     \<spacespace>\<spacespace>     \<Longrightarrow> G\<turnstile>C\<leadsto>J"
+  subint:        "\<lbrakk>G\<turnstile>C\<leadsto>1I; G\<turnstile>I\<preceq>I J\<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
+  subcls1:       "\<lbrakk>G\<turnstile>C\<prec>\<^sub>C\<^sub>1D; G\<turnstile>D\<leadsto>J \<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
+
+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)" 
+apply (erule implmt.induct)
+apply fast+
+done
+
+lemma implmt_ObjectE [elim!]: "G\<turnstile>Object\<leadsto>I \<Longrightarrow> R"
+by (auto dest!: implmtD implmt1D subcls1D)
+
+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"
+apply (erule rtrancl_induct)
+apply  (auto intro: implmt.subcls1)
+done
+
+lemma implmt_subint2: "\<lbrakk> G\<turnstile>A\<leadsto>J; G\<turnstile>J\<preceq>I K\<rbrakk> \<Longrightarrow> G\<turnstile>A\<leadsto>K"
+apply (erule make_imp, erule implmt.induct)
+apply (auto dest: implmt.subint rtrancl_trans implmt.subcls1)
+done
+
+lemma implmt_is_class: "G\<turnstile>C\<leadsto>I \<Longrightarrow> is_class G C"
+apply (erule implmt.induct)
+apply (blast dest: implmt1D subcls1D)+
+done
+
+
+section "widening relation"
+
+inductive "widen G" intros(*widening, viz. method invocation conversion, cf. 5.3
+			    i.e. kind of syntactic subtyping *)
+  refl:    "G\<turnstile>T\<preceq>T"(*identity conversion, cf. 5.1.1 *)
+  subint:  "G\<turnstile>I\<preceq>I J  \<Longrightarrow> G\<turnstile>Iface I\<preceq> Iface J"(*wid.ref.conv.,cf. 5.1.4 *)
+  int_obj: "G\<turnstile>Iface I\<preceq> Class Object"
+  subcls:  "G\<turnstile>C\<preceq>\<^sub>C D  \<Longrightarrow> G\<turnstile>Class C\<preceq> Class D"
+  implmt:  "G\<turnstile>C\<leadsto>I   \<Longrightarrow> G\<turnstile>Class C\<preceq> Iface I"
+  null:    "G\<turnstile>NT\<preceq> RefT R"
+  arr_obj: "G\<turnstile>T.[]\<preceq> Class Object"
+  array:   "G\<turnstile>RefT S\<preceq>RefT T \<Longrightarrow> G\<turnstile>RefT S.[]\<preceq> RefT T.[]"
+
+declare widen.refl [intro!]
+declare widen.intros [simp]
+
+(* too strong in general:
+lemma widen_PrimT: "G\<turnstile>PrimT x\<preceq>T \<Longrightarrow> T = PrimT x"
+*)
+lemma widen_PrimT: "G\<turnstile>PrimT x\<preceq>T \<Longrightarrow> (\<exists>y. T = PrimT y)"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+(* too strong in general:
+lemma widen_PrimT2: "G\<turnstile>S\<preceq>PrimT x \<Longrightarrow> S = PrimT x"
+*)
+lemma widen_PrimT2: "G\<turnstile>S\<preceq>PrimT x \<Longrightarrow> \<exists>y. S = PrimT y"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_RefT: "G\<turnstile>RefT R\<preceq>T \<Longrightarrow> \<exists>t. T=RefT t"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_RefT2: "G\<turnstile>S\<preceq>RefT R \<Longrightarrow> \<exists>t. S=RefT t"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Iface: "G\<turnstile>Iface I\<preceq>T \<Longrightarrow> T=Class Object \<or> (\<exists>J. T=Iface J)"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Iface2: "G\<turnstile>S\<preceq> Iface J \<Longrightarrow> S = NT \<or> (\<exists>I. S = Iface I) \<or> (\<exists>D. S = Class D)"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Iface_Iface: "G\<turnstile>Iface I\<preceq> Iface J \<Longrightarrow> G\<turnstile>I\<preceq>I J"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Iface_Iface_eq [simp]: "G\<turnstile>Iface I\<preceq> Iface J = G\<turnstile>I\<preceq>I J"
+apply (rule iffI)
+apply  (erule widen_Iface_Iface)
+apply (erule widen.subint)
+done
+
+lemma widen_Class: "G\<turnstile>Class C\<preceq>T \<Longrightarrow>  (\<exists>D. T=Class D) \<or> (\<exists>I. T=Iface I)"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Class2: "G\<turnstile>S\<preceq> Class C \<Longrightarrow> C = Object \<or> S = NT \<or> (\<exists>D. S = Class D)"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Class_Class: "G\<turnstile>Class C\<preceq> Class cm \<Longrightarrow> G\<turnstile>C\<preceq>\<^sub>C cm"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Class_Class_eq [simp]: "G\<turnstile>Class C\<preceq> Class cm = G\<turnstile>C\<preceq>\<^sub>C cm"
+apply (rule iffI)
+apply  (erule widen_Class_Class)
+apply (erule widen.subcls)
+done
+
+lemma widen_Class_Iface: "G\<turnstile>Class C\<preceq> Iface I \<Longrightarrow> G\<turnstile>C\<leadsto>I"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Class_Iface_eq [simp]: "G\<turnstile>Class C\<preceq> Iface I = G\<turnstile>C\<leadsto>I"
+apply (rule iffI)
+apply  (erule widen_Class_Iface)
+apply (erule widen.implmt)
+done
+
+lemma widen_Array: "G\<turnstile>S.[]\<preceq>T \<Longrightarrow> T=Class Object \<or> (\<exists>T'. T=T'.[] \<and> G\<turnstile>S\<preceq>T')"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Array2: "G\<turnstile>S\<preceq>T.[] \<Longrightarrow> S = NT \<or> (\<exists>S'. S=S'.[] \<and> G\<turnstile>S'\<preceq>T)"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+
+lemma widen_ArrayPrimT: "G\<turnstile>PrimT t.[]\<preceq>T \<Longrightarrow> T=Class Object \<or> T=PrimT t.[]"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_ArrayRefT: 
+  "G\<turnstile>RefT t.[]\<preceq>T \<Longrightarrow> T=Class Object \<or> (\<exists>s. T=RefT s.[] \<and> G\<turnstile>RefT t\<preceq>RefT s)"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_ArrayRefT_ArrayRefT_eq [simp]: 
+  "G\<turnstile>RefT T.[]\<preceq>RefT T'.[] = G\<turnstile>RefT T\<preceq>RefT T'"
+apply (rule iffI)
+apply (drule widen_ArrayRefT)
+apply simp
+apply (erule widen.array)
+done
+
+lemma widen_Array_Array: "G\<turnstile>T.[]\<preceq>T'.[] \<Longrightarrow> G\<turnstile>T\<preceq>T'"
+apply (drule widen_Array)
+apply auto
+done
+
+lemma widen_Array_Class: "G\<turnstile>S.[] \<preceq> Class C \<Longrightarrow> C=Object"
+by (auto dest: widen_Array)
+
+(*
+qed_typerel "widen_NT2" "G\<turnstile>S\<preceq>NT \<Longrightarrow> S = NT"
+ [prove_widen_lemma "G\<turnstile>S\<preceq>T \<Longrightarrow> T = NT \<longrightarrow> S = NT"]
+*)
+lemma widen_NT2: "G\<turnstile>S\<preceq>NT \<Longrightarrow> S = NT"
+apply (ind_cases "G\<turnstile>S\<preceq>T")
+by auto
+
+lemma widen_Object:"\<lbrakk>isrtype G T;ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>RefT T \<preceq> Class Object"
+apply (case_tac T)
+apply (auto)
+apply (subgoal_tac "G\<turnstile>pid_field_type\<preceq>\<^sub>C Object")
+apply (auto intro: subcls_ObjectI)
+done
+
+lemma widen_trans_lemma [rule_format (no_asm)]: 
+  "\<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"
+apply (erule widen.induct)
+apply        safe
+prefer      5 apply (drule widen_RefT) apply clarsimp
+apply      (frule_tac [1] widen_Iface)
+apply      (frule_tac [2] widen_Class)
+apply      (frule_tac [3] widen_Class)
+apply      (frule_tac [4] widen_Iface)
+apply      (frule_tac [5] widen_Class)
+apply      (frule_tac [6] widen_Array)
+apply      safe
+apply            (rule widen.int_obj)
+prefer          6 apply (drule implmt_is_class) apply simp
+apply (tactic "ALLGOALS (etac thin_rl)")
+prefer         6 apply simp
+apply          (rule_tac [9] widen.arr_obj)
+apply         (rotate_tac [9] -1)
+apply         (frule_tac [9] widen_RefT)
+apply         (auto elim!: rtrancl_trans subcls_implmt implmt_subint2)
+done
+
+lemma ws_widen_trans: "\<lbrakk>G\<turnstile>S\<preceq>U; G\<turnstile>U\<preceq>T; ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>S\<preceq>T"
+by (auto intro: widen_trans_lemma subcls_ObjectI)
+
+lemma widen_antisym_lemma [rule_format (no_asm)]: "\<lbrakk>G\<turnstile>S\<preceq>T;  
+ \<forall>I J. G\<turnstile>I\<preceq>I J \<and> G\<turnstile>J\<preceq>I I \<longrightarrow> I = J;  
+ \<forall>C D. G\<turnstile>C\<preceq>\<^sub>C D \<and> G\<turnstile>D\<preceq>\<^sub>C C \<longrightarrow> C = D;  
+ \<forall>I  . G\<turnstile>Object\<leadsto>I        \<longrightarrow> False\<rbrakk> \<Longrightarrow> G\<turnstile>T\<preceq>S \<longrightarrow> S = T"
+apply (erule widen.induct)
+apply (auto dest: widen_Iface widen_NT2 widen_Class)
+done
+
+lemmas subint_antisym = 
+       subint1_acyclic [THEN acyclic_impl_antisym_rtrancl, standard]
+lemmas subcls_antisym = 
+       subcls1_acyclic [THEN acyclic_impl_antisym_rtrancl, standard]
+
+lemma widen_antisym: "\<lbrakk>G\<turnstile>S\<preceq>T; G\<turnstile>T\<preceq>S; ws_prog G\<rbrakk> \<Longrightarrow> S=T"
+by (fast elim: widen_antisym_lemma subint_antisym [THEN antisymD] 
+                                   subcls_antisym [THEN antisymD])
+
+lemma widen_ObjectD [dest!]: "G\<turnstile>Class Object\<preceq>T \<Longrightarrow> T=Class Object"
+apply (frule widen_Class)
+apply (fast dest: widen_Class_Class widen_Class_Iface)
+done
+
+constdefs
+  widens :: "prog \<Rightarrow> [ty list, ty list] \<Rightarrow> bool" ("_\<turnstile>_[\<preceq>]_" [71,71,71] 70)
+ "G\<turnstile>Ts[\<preceq>]Ts' \<equiv> list_all2 (\<lambda>T T'. G\<turnstile>T\<preceq>T') Ts Ts'"
+
+lemma widens_Nil [simp]: "G\<turnstile>[][\<preceq>][]"
+apply (unfold widens_def)
+apply auto
+done
+
+lemma widens_Cons [simp]: "G\<turnstile>(S#Ss)[\<preceq>](T#Ts) = (G\<turnstile>S\<preceq>T \<and> G\<turnstile>Ss[\<preceq>]Ts)"
+apply (unfold widens_def)
+apply auto
+done
+
+
+section "narrowing relation"
+
+(* all properties of narrowing and casting conversions we actually need *)
+(* these can easily be proven from the definitions below *)
+(*
+rules
+  cast_RefT2   "G\<turnstile>S\<preceq>? RefT R   \<Longrightarrow> \<exists>t. S=RefT t" 
+  cast_PrimT2  "G\<turnstile>S\<preceq>? PrimT pt \<Longrightarrow> \<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
+*)
+
+(* more detailed than necessary for type-safety, see above rules. *)
+inductive "narrow G" intros (* narrowing reference conversion, cf. 5.1.5 *)
+
+  subcls:  "G\<turnstile>C\<preceq>\<^sub>C D \<Longrightarrow> G\<turnstile>     Class D\<succ>Class C"
+  implmt:  "\<not>G\<turnstile>C\<leadsto>I \<Longrightarrow> G\<turnstile>     Class C\<succ>Iface I"
+  obj_arr: "G\<turnstile>Class Object\<succ>T.[]"
+  int_cls: "G\<turnstile>     Iface I\<succ>Class C"
+  subint:  "imethds G I hidings imethds G J entails 
+	   (\<lambda>(md, mh   ) (md',mh').\<spacespace>G\<turnstile>mrt mh\<preceq>mrt mh') \<Longrightarrow>
+	   \<not>G\<turnstile>I\<preceq>I J         \<spacespace>\<spacespace>\<spacespace>\<Longrightarrow> G\<turnstile>     Iface I\<succ>Iface J"
+  array:   "G\<turnstile>RefT S\<succ>RefT T   \<spacespace>\<Longrightarrow> G\<turnstile>   RefT S.[]\<succ>RefT T.[]"
+
+(*unused*)
+lemma narrow_RefT: "G\<turnstile>RefT R\<succ>T \<Longrightarrow> \<exists>t. T=RefT t"
+apply (ind_cases "G\<turnstile>S\<succ>T")
+by auto
+
+lemma narrow_RefT2: "G\<turnstile>S\<succ>RefT R \<Longrightarrow> \<exists>t. S=RefT t"
+apply (ind_cases "G\<turnstile>S\<succ>T")
+by auto
+
+(*unused*)
+lemma narrow_PrimT: "G\<turnstile>PrimT pt\<succ>T \<Longrightarrow> \<exists>t. T=PrimT t"
+apply (ind_cases "G\<turnstile>S\<succ>T")
+by auto
+
+lemma narrow_PrimT2: "G\<turnstile>S\<succ>PrimT pt \<Longrightarrow>  
+				  \<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
+apply (ind_cases "G\<turnstile>S\<succ>T")
+by auto
+
+
+section "casting relation"
+
+inductive "cast G" intros (* casting conversion, cf. 5.5 *)
+
+  widen:   "G\<turnstile>S\<preceq>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
+  narrow:  "G\<turnstile>S\<succ>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
+
+(*
+lemma ??unknown??: "\<lbrakk>G\<turnstile>S\<preceq>T; G\<turnstile>S\<succ>T\<rbrakk> \<Longrightarrow> R"
+ deferred *)
+
+(*unused*)
+lemma cast_RefT: "G\<turnstile>RefT R\<preceq>? T \<Longrightarrow> \<exists>t. T=RefT t"
+apply (ind_cases "G\<turnstile>S\<preceq>? T")
+by (auto dest: widen_RefT narrow_RefT)
+
+lemma cast_RefT2: "G\<turnstile>S\<preceq>? RefT R \<Longrightarrow> \<exists>t. S=RefT t"
+apply (ind_cases "G\<turnstile>S\<preceq>? T")
+by (auto dest: widen_RefT2 narrow_RefT2)
+
+(*unused*)
+lemma cast_PrimT: "G\<turnstile>PrimT pt\<preceq>? T \<Longrightarrow> \<exists>t. T=PrimT t"
+apply (ind_cases "G\<turnstile>S\<preceq>? T")
+by (auto dest: widen_PrimT narrow_PrimT)
+
+lemma cast_PrimT2: "G\<turnstile>S\<preceq>? PrimT pt \<Longrightarrow> \<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
+apply (ind_cases "G\<turnstile>S\<preceq>? T")
+by (auto dest: widen_PrimT2 narrow_PrimT2)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/TypeSafe.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,1211 @@
+(*  Title:      isabelle/Bali/TypeSafe.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* The type soundness proof for Java *}
+
+
+theory TypeSafe = Eval + WellForm + Conform:
+
+section "result conformance"
+
+constdefs
+  assign_conforms :: "st \<Rightarrow> (val \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> ty \<Rightarrow> env_ \<Rightarrow> bool"
+          ("_\<le>|_\<preceq>_\<Colon>\<preceq>_"                                        [71,71,71,71] 70)
+ "s\<le>|f\<preceq>T\<Colon>\<preceq>E \<equiv>
+  \<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"
+
+  rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool"
+          ("_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_"                               [71,71,71,71,71,71] 70)
+  "G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T 
+    \<equiv> case T of
+        Inl T  \<Rightarrow> if (\<exists>vf. t=In2 vf)
+                  then G,s\<turnstile>fst (the_In2 v)\<Colon>\<preceq>T \<and> s\<le>|snd (the_In2 v)\<preceq>T\<Colon>\<preceq>(G,L)
+                  else G,s\<turnstile>the_In1 v\<Colon>\<preceq>T
+      | Inr Ts \<Rightarrow> list_all2 (conf G s) (the_In3 v) Ts"
+
+lemma rconf_In1 [simp]: 
+ "G,L,s\<turnstile>In1 ec\<succ>In1 v \<Colon>\<preceq>Inl T  =  G,s\<turnstile>v\<Colon>\<preceq>T"
+apply (unfold rconf_def)
+apply (simp (no_asm))
+done
+
+lemma rconf_In2 [simp]: 
+ "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))"
+apply (unfold rconf_def)
+apply (simp (no_asm))
+done
+
+lemma rconf_In3 [simp]: 
+ "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"
+apply (unfold rconf_def)
+apply (simp (no_asm))
+done
+
+section "fits and conf"
+
+(* unused *)
+lemma conf_fits: "G,s\<turnstile>v\<Colon>\<preceq>T \<Longrightarrow> G,s\<turnstile>v fits T"
+apply (unfold fits_def)
+apply clarify
+apply (erule swap, simp (no_asm_use))
+apply (drule conf_RefTD)
+apply auto
+done
+
+lemma fits_conf: 
+  "\<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'"
+apply (auto dest!: fitsD cast_PrimT2 cast_RefT2)
+apply (force dest: conf_RefTD intro: conf_AddrI)
+done
+
+lemma fits_Array: 
+ "\<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'"
+apply (auto dest!: fitsD widen_ArrayPrimT widen_ArrayRefT)
+apply (force dest: conf_RefTD intro: conf_AddrI)
+done
+
+
+
+section "gext"
+
+lemma halloc_gext: "\<And>s1 s2. G\<turnstile>s1 \<midarrow>halloc oi\<succ>a\<rightarrow> s2 \<Longrightarrow> snd s1\<le>|snd s2"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (erule halloc.induct)
+apply  (auto dest!: new_AddrD)
+done
+
+lemma sxalloc_gext: "\<And>s1 s2. G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2 \<Longrightarrow> snd s1\<le>|snd s2"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (erule sxalloc.induct)
+apply   (auto dest!: halloc_gext)
+done
+
+lemma eval_gext_lemma [rule_format (no_asm)]: 
+ "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<Longrightarrow> snd s\<le>|snd s' \<and> (case w of  
+    In1 v \<Rightarrow> True  
+  | In2 vf \<Rightarrow> normal s \<longrightarrow> (\<forall>v x s. s\<le>|snd (assign (snd vf) v (x,s)))  
+  | In3 vs \<Rightarrow> True)"
+apply (erule eval_induct)
+prefer 24 
+  apply (case_tac "inited C (globs s0)", clarsimp, erule thin_rl) (* Init *)
+apply (auto del: conjI  dest!: not_initedD gext_new sxalloc_gext halloc_gext
+ simp  add: lvar_def fvar_def2 avar_def2 init_lvars_def2
+ split del: split_if_asm split add: sum3.split)
+(* 6 subgoals *)
+apply force+
+done
+
+lemma evar_gext_f: 
+  "G\<turnstile>Norm s1 \<midarrow>e=\<succ>vf \<rightarrow> s2 \<Longrightarrow> s\<le>|snd (assign (snd vf) v (x,s))"
+apply (drule eval_gext_lemma [THEN conjunct2])
+apply auto
+done
+
+lemmas eval_gext = eval_gext_lemma [THEN conjunct1]
+
+lemma eval_gext': "G\<turnstile>(x1,s1) \<midarrow>t\<succ>\<rightarrow> (w,x2,s2) \<Longrightarrow> s1\<le>|s2"
+apply (drule eval_gext)
+apply auto
+done
+
+lemma init_yields_initd: "G\<turnstile>Norm s1 \<midarrow>Init C\<rightarrow> s2 \<Longrightarrow> initd C s2"
+apply (erule eval_cases , auto split del: split_if_asm)
+apply (case_tac "inited C (globs s1)")
+apply  (clarsimp split del: split_if_asm)+
+apply (drule eval_gext')+
+apply (drule init_class_obj_inited)
+apply (erule inited_gext)
+apply (simp (no_asm_use))
+done
+
+
+section "Lemmas"
+
+lemma obj_ty_obj_class1: 
+ "\<lbrakk>wf_prog G; is_type G (obj_ty obj)\<rbrakk> \<Longrightarrow> is_class G (obj_class obj)"
+apply (case_tac "tag obj")
+apply (auto simp add: obj_ty_def obj_class_def)
+done
+
+lemma oconf_init_obj: 
+ "\<lbrakk>wf_prog G;  
+ (case r of Heap a \<Rightarrow> is_type G (obj_ty obj) | Stat C \<Rightarrow> is_class G C)
+\<rbrakk> \<Longrightarrow> G,s\<turnstile>obj \<lparr>values:=init_vals (var_tys G (tag obj) r)\<rparr>\<Colon>\<preceq>\<surd>r"
+apply (auto intro!: oconf_init_obj_lemma unique_fields)
+done
+
+(*
+lemma obj_split: "P obj = (\<forall> oi vs. obj = \<lparr>tag=oi,values=vs\<rparr> \<longrightarrow> ?P \<lparr>tag=oi,values=vs\<rparr>)"
+apply auto
+apply (case_tac "obj")
+apply auto
+*)
+
+lemma conforms_newG: "\<lbrakk>globs s oref = None; (x, s)\<Colon>\<preceq>(G,L);   
+  wf_prog G; case oref of Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=vs\<rparr>)  
+                        | Stat C \<Rightarrow> is_class G C\<rbrakk> \<Longrightarrow>  
+  (x, init_obj G oi oref s)\<Colon>\<preceq>(G, L)"
+apply (unfold init_obj_def)
+apply (auto elim!: conforms_gupd dest!: oconf_init_obj 
+            simp add: obj.update_defs)
+done
+
+lemma conforms_init_class_obj: 
+ "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L); wf_prog G; class G C=Some y; \<not> inited C (globs s)\<rbrakk> \<Longrightarrow> 
+  (x,init_class_obj G C s)\<Colon>\<preceq>(G, L)"
+apply (rule not_initedD [THEN conforms_newG])
+apply    (auto)
+done
+
+
+lemma fst_init_lvars[simp]: 
+ "fst (init_lvars G C sig (invmode m e) a' pvs (x,s)) = 
+  (if static m then x else (np a') x)"
+apply (simp (no_asm) add: init_lvars_def2)
+done
+
+
+lemma halloc_conforms: "\<And>s1. \<lbrakk>G\<turnstile>s1 \<midarrow>halloc oi\<succ>a\<rightarrow> s2; wf_prog G; s1\<Colon>\<preceq>(G, L); 
+  is_type G (obj_ty \<lparr>tag=oi,values=fs\<rparr>)\<rbrakk> \<Longrightarrow> s2\<Colon>\<preceq>(G, L)"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (case_tac "aa")
+apply  (auto elim!: halloc_elim_cases dest!: new_AddrD 
+       intro!: conforms_newG [THEN conforms_xconf] conf_AddrI)
+done
+
+lemma halloc_type_sound: "\<And>s1. \<lbrakk>G\<turnstile>s1 \<midarrow>halloc oi\<succ>a\<rightarrow> (x,s); wf_prog G; s1\<Colon>\<preceq>(G, L);
+  T = obj_ty \<lparr>tag=oi,values=fs\<rparr>; is_type G T\<rbrakk> \<Longrightarrow>  
+  (x,s)\<Colon>\<preceq>(G, L) \<and> (x = None \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq>T)"
+apply (auto elim!: halloc_conforms)
+apply (case_tac "aa")
+apply (subst obj_ty_eq)
+apply  (auto elim!: halloc_elim_cases dest!: new_AddrD intro!: conf_AddrI)
+done
+
+lemma sxalloc_type_sound: 
+ "\<And>s1 s2. \<lbrakk>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; wf_prog G\<rbrakk> \<Longrightarrow> case fst s1 of  
+  None \<Rightarrow> s2 = s1 | Some x \<Rightarrow>  
+  (\<exists>a. fst s2 = Some(Xcpt (Loc a)) \<and> (\<forall>L. s1\<Colon>\<preceq>(G,L) \<longrightarrow> s2\<Colon>\<preceq>(G,L)))"
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply (erule sxalloc.induct)
+apply   auto
+apply (rule halloc_conforms [THEN conforms_xconf])
+apply     (auto elim!: halloc_elim_cases dest!: new_AddrD intro!: conf_AddrI)
+done
+
+lemma wt_init_comp_ty: 
+"is_acc_type G (pid C) T \<Longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>"
+apply (unfold init_comp_ty_def)
+apply (clarsimp simp add: accessible_in_RefT_simp 
+                          is_acc_type_def is_acc_class_def)
+done
+
+
+declare fun_upd_same [simp]
+
+declare fun_upd_apply [simp del]
+
+
+constdefs
+  DynT_prop::"[prog,inv_mode,qtname,ref_ty] \<Rightarrow> bool" 
+                                              ("_\<turnstile>_\<rightarrow>_\<preceq>_"[71,71,71,71]70)
+ "G\<turnstile>mode\<rightarrow>D\<preceq>t \<equiv> mode = IntVir \<longrightarrow> is_class G D \<and> 
+                     (if (\<exists>T. t=ArrayT T) then D=Object else G\<turnstile>Class D\<preceq>RefT t)"
+
+lemma DynT_propI: 
+ "\<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> 
+  \<Longrightarrow>  G\<turnstile>mode\<rightarrow>invocation_class mode s a' statT\<preceq>statT"
+proof (unfold DynT_prop_def)
+  assume state_conform: "(x,s)\<Colon>\<preceq>(G, L)"
+     and      statT_a': "G,s\<turnstile>a'\<Colon>\<preceq>RefT statT"
+     and            wf: "wf_prog G"
+     and          mode: "mode = IntVir \<longrightarrow> a' \<noteq> Null"
+  let ?invCls = "(invocation_class mode s a' statT)"
+  let ?IntVir = "mode = IntVir"
+  let ?Concl  = "\<lambda>invCls. is_class G invCls \<and>
+                          (if \<exists>T. statT = ArrayT T
+                                  then invCls = Object
+                                  else G\<turnstile>Class invCls\<preceq>RefT statT)"
+  show "?IntVir \<longrightarrow> ?Concl ?invCls"
+  proof  
+    assume modeIntVir: ?IntVir 
+    with mode have not_Null: "a' \<noteq> Null" ..
+    from statT_a' not_Null state_conform 
+    obtain a obj 
+      where obj_props:  "a' = Addr a" "globs s (Inl a) = Some obj"   
+                        "G\<turnstile>obj_ty obj\<preceq>RefT statT" "is_type G (obj_ty obj)"
+      by (blast dest: conforms_RefTD)
+    show "?Concl ?invCls"
+    proof (cases "tag obj")
+      case CInst
+      with modeIntVir obj_props
+      show ?thesis 
+	by (auto dest!: widen_Array2 split add: split_if)
+    next
+      case Arr
+      from Arr obtain T where "obj_ty obj = T.[]" by (blast dest: obj_ty_Arr1)
+      moreover from Arr have "obj_class obj = Object" 
+	by (blast dest: obj_class_Arr1)
+      moreover note modeIntVir obj_props wf 
+      ultimately show ?thesis by (auto dest!: widen_Array )
+    qed
+  qed
+qed
+
+lemma invocation_methd:
+"\<lbrakk>wf_prog G; statT \<noteq> NullT; 
+ (\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC);
+ (\<forall>     I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> mode \<noteq> SuperM);
+ (\<forall>     T. statT = ArrayT T \<longrightarrow> mode \<noteq> SuperM);
+ G\<turnstile>mode\<rightarrow>invocation_class mode s a' statT\<preceq>statT;  
+ dynlookup G statT (invocation_class mode s a' statT) sig = Some m \<rbrakk> 
+\<Longrightarrow> methd G (invocation_declclass G mode s a' statT sig) sig = Some m"
+proof -
+  assume         wf: "wf_prog G"
+     and  not_NullT: "statT \<noteq> NullT"
+     and statC_prop: "(\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC)"
+     and statI_prop: "(\<forall> I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> mode \<noteq> SuperM)"
+     and statA_prop: "(\<forall>     T. statT = ArrayT T \<longrightarrow> mode \<noteq> SuperM)"
+     and  invC_prop: "G\<turnstile>mode\<rightarrow>invocation_class mode s a' statT\<preceq>statT"
+     and  dynlookup: "dynlookup G statT (invocation_class mode s a' statT) sig 
+                      = Some m"
+  show ?thesis
+  proof (cases statT)
+    case NullT
+    with not_NullT show ?thesis by simp
+  next
+    case IfaceT
+    with statI_prop obtain I 
+      where    statI: "statT = IfaceT I" and 
+            is_iface: "is_iface G I"     and
+          not_SuperM: "mode \<noteq> SuperM" by blast            
+    
+    show ?thesis 
+    proof (cases mode)
+      case Static
+      with wf dynlookup statI is_iface 
+      show ?thesis
+	by (auto simp add: invocation_declclass_def dynlookup_def 
+                           dynimethd_def dynmethd_C_C 
+	            intro: dynmethd_declclass
+	            dest!: wf_imethdsD
+                     dest: table_of_map_SomeI
+                    split: split_if_asm)
+    next	
+      case SuperM
+      with not_SuperM show ?thesis ..
+    next
+      case IntVir
+      with wf dynlookup IfaceT invC_prop show ?thesis 
+	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
+                           DynT_prop_def
+	            intro: methd_declclass dynmethd_declclass
+	            split: split_if_asm)
+    qed
+  next
+    case ClassT
+    show ?thesis
+    proof (cases mode)
+      case Static
+      with wf ClassT dynlookup statC_prop 
+      show ?thesis by (auto simp add: invocation_declclass_def dynlookup_def
+                               intro: dynmethd_declclass)
+    next
+      case SuperM
+      with wf ClassT dynlookup statC_prop 
+      show ?thesis by (auto simp add: invocation_declclass_def dynlookup_def
+                               intro: dynmethd_declclass)
+    next
+      case IntVir
+      with wf ClassT dynlookup statC_prop invC_prop 
+      show ?thesis
+	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
+                           DynT_prop_def
+	            intro: dynmethd_declclass)
+    qed
+  next
+    case ArrayT
+    show ?thesis
+    proof (cases mode)
+      case Static
+      with wf ArrayT dynlookup show ?thesis
+	by (auto simp add: invocation_declclass_def dynlookup_def 
+                           dynimethd_def dynmethd_C_C
+                    intro: dynmethd_declclass
+                     dest: table_of_map_SomeI)
+    next
+      case SuperM
+      with ArrayT statA_prop show ?thesis by blast
+    next
+      case IntVir
+      with wf ArrayT dynlookup invC_prop show ?thesis
+	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
+                           DynT_prop_def dynmethd_C_C
+                    intro: dynmethd_declclass
+                     dest: table_of_map_SomeI)
+    qed
+  qed
+qed
+   
+declare split_paired_All [simp del] split_paired_Ex [simp del] 
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac";
+claset_ref () := claset () delSWrapper "split_all_tac"
+*}
+lemma DynT_mheadsD: 
+"\<lbrakk>G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT; 
+  wf_prog G; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT; 
+  sm \<in> mheads G C statT sig; 
+  invC = invocation_class (invmode (mhd sm) e) s a' statT;
+  declC =invocation_declclass G (invmode (mhd sm) e) s a' statT sig
+ \<rbrakk> \<Longrightarrow> 
+  \<exists> dm. 
+  methd G declC sig = Some dm  \<and> G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm) \<and> 
+  wf_mdecl G declC (sig, mthd dm) \<and>
+  declC = declclass dm \<and>
+  is_static dm = is_static sm \<and>  
+  is_class G invC \<and> is_class G declC  \<and> G\<turnstile>invC\<preceq>\<^sub>C declC \<and>  
+  (if invmode (mhd sm) e = IntVir
+      then (\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)
+      else (  (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
+            \<or> (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object)) \<and> 
+           (declrefT sm) = ClassT (declclass dm))"
+proof -
+  assume invC_prop: "G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT" 
+     and        wf: "wf_prog G" 
+     and      wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
+     and        sm: "sm \<in> mheads G C statT sig" 
+     and      invC: "invC = invocation_class (invmode (mhd sm) e) s a' statT"
+     and     declC: "declC = 
+                    invocation_declclass G (invmode (mhd sm) e) s a' statT sig"
+  from wt_e wf have type_statT: "is_type G (RefT statT)"
+    by (auto dest: ty_expr_is_type)
+  from sm have not_Null: "statT \<noteq> NullT" by auto
+  from type_statT 
+  have wf_C: "(\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC)"
+    by (auto)
+  from type_statT wt_e 
+  have wf_I: "(\<forall>I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> 
+                                        invmode (mhd sm) e \<noteq> SuperM)"
+    by (auto dest: invocationTypeExpr_noClassD)
+  from wt_e
+  have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode (mhd sm) e \<noteq> SuperM)"
+    by (auto dest: invocationTypeExpr_noClassD)
+  show ?thesis
+  proof (cases "invmode (mhd sm) e = IntVir")
+    case True
+    with invC_prop not_Null
+    have invC_prop': " is_class G invC \<and> 
+                      (if (\<exists>T. statT=ArrayT T) then invC=Object 
+                                              else G\<turnstile>Class invC\<preceq>RefT statT)"
+      by (auto simp add: DynT_prop_def) 
+    from True 
+    have "\<not> is_static sm"
+      by (simp add: invmode_IntVir_eq)
+    with invC_prop' not_Null
+    have "G,statT \<turnstile> invC valid_lookup_cls_for (is_static sm)"
+      by (cases statT) auto
+    with sm wf type_statT obtain dm where
+           dm: "dynlookup G statT invC sig = Some dm" and
+      resT_dm: "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm)"      and
+       static: "is_static dm = is_static sm"
+      by - (drule dynamic_mheadsD,auto)
+    with declC invC not_Null 
+    have declC': "declC = (declclass dm)" 
+      by (auto simp add: invocation_declclass_def)
+    with wf invC declC not_Null wf_C wf_I wf_A invC_prop dm 
+    have dm': "methd G declC sig = Some dm"
+      by - (drule invocation_methd,auto)
+    from wf dm invC_prop' declC' type_statT 
+    have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
+      by (auto dest: dynlookup_declC)
+    from wf dm' declC_prop declC' 
+    have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
+      by (auto dest: methd_wf_mdecl)
+    from invC_prop' 
+    have statC_prop: "(\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)"
+      by auto
+    from True dm' resT_dm wf_dm invC_prop' declC_prop statC_prop declC' static
+    show ?thesis by auto
+  next
+    case False
+    with type_statT wf invC not_Null wf_I wf_A
+    have invC_prop': "is_class G invC \<and>  
+                     ((\<exists> statC. statT=ClassT statC \<and> invC=statC) \<or>
+                      (\<forall> statC. statT\<noteq>ClassT statC \<and> invC=Object)) "
+        by (case_tac "statT") (auto simp add: invocation_class_def 
+                                       split: inv_mode.splits)
+    with not_Null wf
+    have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig"
+      by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_C_C
+                                            dynimethd_def)
+    from sm wf wt_e not_Null False invC_prop' obtain "dm" where
+                    dm: "methd G invC sig = Some dm"          and
+	eq_declC_sm_dm:"declrefT sm = ClassT (declclass dm)"  and
+	     eq_mheads:"mhd sm=mhead (mthd dm) "
+      by - (drule static_mheadsD, auto dest: accmethd_SomeD)
+    then have static: "is_static dm = is_static sm" by - (case_tac "sm",auto)
+    with declC invC dynlookup_static dm
+    have declC': "declC = (declclass dm)"  
+      by (auto simp add: invocation_declclass_def)
+    from invC_prop' wf declC' dm 
+    have dm': "methd G declC sig = Some dm"
+      by (auto intro: methd_declclass)
+    from wf dm invC_prop' declC' type_statT 
+    have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
+      by (auto dest: methd_declC )
+    then have declC_prop1: "invC=Object \<longrightarrow> declC=Object"  by auto
+    from wf dm' declC_prop declC' 
+    have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
+      by (auto dest: methd_wf_mdecl)
+    from invC_prop' declC_prop declC_prop1
+    have statC_prop: "(   (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
+                       \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object))" 
+      by auto
+    from False dm' wf_dm invC_prop' declC_prop statC_prop declC' 
+         eq_declC_sm_dm eq_mheads static
+    show ?thesis by auto
+  qed
+qed
+
+(*
+lemma DynT_mheadsD: 
+"\<lbrakk>G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT; 
+  wf_prog G; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT; 
+  sm \<in> mheads G C statT sig; 
+  invC = invocation_class (invmode (mhd sm) e) s a' statT;
+  declC =invocation_declclass G (invmode (mhd sm) e) s a' statT sig
+ \<rbrakk> \<Longrightarrow> 
+  \<exists> dm. 
+  methd G declC sig = Some dm  \<and> G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm) \<and> 
+  wf_mdecl G declC (sig, mthd dm) \<and>  
+  is_class G invC \<and> is_class G declC  \<and> G\<turnstile>invC\<preceq>\<^sub>C declC \<and>  
+  (if invmode (mhd sm) e = IntVir
+      then (\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)
+      else (\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>statC \<preceq>\<^sub>C declC) \<and> 
+           (declrefT sm) = ClassT (declclass dm))"
+proof -
+  assume invC_prop: "G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT" 
+     and        wf: "wf_prog G" 
+     and      wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
+     and        sm: "sm \<in> mheads G C statT sig" 
+     and      invC: "invC = invocation_class (invmode (mhd sm) e) s a' statT"
+     and     declC: "declC = 
+                    invocation_declclass G (invmode (mhd sm) e) s a' statT sig"
+  from wt_e wf have type_statT: "is_type G (RefT statT)"
+    by (auto dest: ty_expr_is_type)
+  from sm have not_Null: "statT \<noteq> NullT" by auto
+  from type_statT 
+  have wf_C: "(\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC)"
+    by (auto)
+  from type_statT wt_e 
+  have wf_I: "(\<forall>I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> 
+                                        invmode (mhd sm) e \<noteq> SuperM)"
+    by (auto dest: invocationTypeExpr_noClassD)
+  from wt_e
+  have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode (mhd sm) e \<noteq> SuperM)"
+    by (auto dest: invocationTypeExpr_noClassD)
+  show ?thesis
+  proof (cases "invmode (mhd sm) e = IntVir")
+    case True
+    with invC_prop not_Null
+    have invC_prop': "is_class G invC \<and>  
+                      (if (\<exists>T. statT=ArrayT T) then invC=Object 
+                                              else G\<turnstile>Class invC\<preceq>RefT statT)"
+      by (auto simp add: DynT_prop_def) 
+    with sm wf type_statT not_Null obtain dm where
+           dm: "dynlookup G statT invC sig = Some dm" and
+      resT_dm: "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm)"
+      by - (clarify,drule dynamic_mheadsD,auto)
+    with declC invC not_Null 
+    have declC': "declC = (declclass dm)" 
+      by (auto simp add: invocation_declclass_def)
+    with wf invC declC not_Null wf_C wf_I wf_A invC_prop dm 
+    have dm': "methd G declC sig = Some dm"
+      by - (drule invocation_methd,auto)
+    from wf dm invC_prop' declC' type_statT 
+    have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
+      by (auto dest: dynlookup_declC)
+    from wf dm' declC_prop declC' 
+    have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
+      by (auto dest: methd_wf_mdecl)
+    from invC_prop' 
+    have statC_prop: "(\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)"
+      by auto
+    from True dm' resT_dm wf_dm invC_prop' declC_prop statC_prop
+    show ?thesis by auto
+  next
+    case False
+    
+    with type_statT wf invC not_Null wf_I wf_A
+    have invC_prop': "is_class G invC \<and>  
+                     ((\<exists> statC. statT=ClassT statC \<and> invC=statC) \<or>
+                      (\<forall> statC. statT\<noteq>ClassT statC \<and> invC=Object)) "
+        
+        by (case_tac "statT") (auto simp add: invocation_class_def 
+                                       split: inv_mode.splits)
+    with not_Null 
+    have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig"
+      by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_def 
+                                            dynimethd_def)
+    from sm wf wt_e not_Null False invC_prop' obtain "dm" where
+                    dm: "methd G invC sig = Some dm"          and
+	eq_declC_sm_dm:"declrefT sm = ClassT (declclass dm)"  and
+	     eq_mheads:"mhd sm=mhead (mthd dm) "
+      by - (drule static_mheadsD, auto dest: accmethd_SomeD)
+    with declC invC dynlookup_static dm
+    have declC': "declC = (declclass dm)"  
+      by (auto simp add: invocation_declclass_def)
+    from invC_prop' wf declC' dm 
+    have dm': "methd G declC sig = Some dm"
+      by (auto intro: methd_declclass)
+    from wf dm invC_prop' declC' type_statT 
+    have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
+      by (auto dest: methd_declC )   
+    from wf dm' declC_prop declC' 
+    have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
+      by (auto dest: methd_wf_mdecl)
+    from invC_prop' declC_prop
+    have statC_prop: "(\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>statC \<preceq>\<^sub>C declC)" 
+      by auto
+    from False dm' wf_dm invC_prop' declC_prop statC_prop 
+         eq_declC_sm_dm eq_mheads
+    show ?thesis by auto
+  qed
+qed	
+*)
+
+declare split_paired_All [simp del] split_paired_Ex [simp del] 
+declare split_if     [split del] split_if_asm     [split del] 
+        option.split [split del] option.split_asm [split del]
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac";
+claset_ref () := claset () delSWrapper "split_all_tac"
+*}
+
+lemma DynT_conf: "\<lbrakk>G\<turnstile>invocation_class mode s a' statT \<preceq>\<^sub>C declC; wf_prog G;
+ isrtype G (statT);
+ G,s\<turnstile>a'\<Colon>\<preceq>RefT statT; mode = IntVir \<longrightarrow> a' \<noteq> Null;  
+  mode \<noteq> IntVir \<longrightarrow>    (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
+                    \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object)\<rbrakk> 
+ \<Longrightarrow>G,s\<turnstile>a'\<Colon>\<preceq> Class declC"
+apply (case_tac "mode = IntVir")
+apply (drule conf_RefTD)
+apply (force intro!: conf_AddrI 
+            simp add: obj_class_def split add: obj_tag.split_asm)
+apply  clarsimp
+apply  safe
+apply    (erule (1) widen.subcls [THEN conf_widen])
+apply    (erule wf_ws_prog)
+
+apply    (frule widen_Object) apply (erule wf_ws_prog)
+apply    (erule (1) conf_widen) apply (erule wf_ws_prog)
+done
+
+
+lemma Ass_lemma: 
+ "\<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'; 
+   s1\<le>|s2 \<longrightarrow> assign f v (Norm s2)\<Colon>\<preceq>(G, L)
+  \<rbrakk> \<Longrightarrow> assign f v (Norm s2)\<Colon>\<preceq>(G, L) \<and> 
+        (\<lambda>(x',s'). x' = None \<longrightarrow> G,s'\<turnstile>v\<Colon>\<preceq>T') (assign f v (Norm s2))"
+apply (drule_tac x = "None" and s = "s2" and v = "v" in evar_gext_f)
+apply (drule eval_gext', clarsimp)
+apply (erule conf_gext)
+apply simp
+done
+
+lemma Throw_lemma: "\<lbrakk>G\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable; wf_prog G; (x1,s1)\<Colon>\<preceq>(G, L);  
+    x1 = None \<longrightarrow> G,s1\<turnstile>a'\<Colon>\<preceq> Class tn\<rbrakk> \<Longrightarrow> (throw a' x1, s1)\<Colon>\<preceq>(G, L)"
+apply (auto split add: split_abrupt_if simp add: throw_def2)
+apply (erule conforms_xconf)
+apply (frule conf_RefTD)
+apply (auto elim: widen.subcls [THEN conf_widen])
+done
+
+lemma Try_lemma: "\<lbrakk>G\<turnstile>obj_ty (the (globs s1' (Heap a)))\<preceq> Class tn; 
+ (Some (Xcpt (Loc a)), s1')\<Colon>\<preceq>(G, L); wf_prog G\<rbrakk> 
+ \<Longrightarrow> Norm (lupd(vn\<mapsto>Addr a) s1')\<Colon>\<preceq>(G, L(vn\<mapsto>Class tn))"
+apply (rule conforms_allocL)
+apply  (erule conforms_NormI)
+apply (drule conforms_XcptLocD [THEN conf_RefTD],rule HOL.refl)
+apply (auto intro!: conf_AddrI)
+done
+
+lemma Fin_lemma: 
+"\<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> 
+\<Longrightarrow>  (abrupt_if True (Some a) x2, s2)\<Colon>\<preceq>(G, L)"
+apply (force elim: eval_gext' conforms_xgext split add: split_abrupt_if)
+done
+
+lemma FVar_lemma1: "\<lbrakk>table_of (DeclConcepts.fields G Ca) (fn, C) = Some f ; 
+  x2 = None \<longrightarrow> G,s2\<turnstile>a\<Colon>\<preceq> Class Ca; wf_prog G; G\<turnstile>Ca\<preceq>\<^sub>C C; C \<noteq> Object; 
+  class G C = Some y; (x2,s2)\<Colon>\<preceq>(G, L); s1\<le>|s2; inited C (globs s1); 
+  (if static f then id else np a) x2 = None\<rbrakk> 
+ \<Longrightarrow>  
+  \<exists>obj. globs s2 (if static f then Inr C else Inl (the_Addr a)) = Some obj \<and> 
+  var_tys G (tag obj)  (if static f then Inr C else Inl(the_Addr a)) 
+          (Inl(fn,C)) = Some (type f)"
+apply (drule initedD)
+apply (frule subcls_is_class2, simp (no_asm_simp))
+apply (case_tac "static f")
+apply  clarsimp
+apply  (drule (1) rev_gext_objD, clarsimp)
+apply  (frule fields_declC, erule wf_ws_prog, simp (no_asm_simp))
+apply  (rule var_tys_Some_eq [THEN iffD2])
+apply  clarsimp
+apply  (erule fields_table_SomeI, simp (no_asm))
+apply clarsimp
+apply (drule conf_RefTD, clarsimp, rule var_tys_Some_eq [THEN iffD2])
+apply (auto dest!: widen_Array split add: obj_tag.split)
+apply (rotate_tac -1) (* to front: tag obja = CInst pid_field_type to enable
+                         conditional rewrite *)
+apply (rule fields_table_SomeI)
+apply (auto elim!: fields_mono subcls_is_class2)
+done
+
+lemma FVar_lemma: 
+"\<lbrakk>((v, f), Norm s2') = fvar C (static field) fn a (x2, s2); G\<turnstile>Ca\<preceq>\<^sub>C C;  
+  table_of (DeclConcepts.fields G Ca) (fn, C) = Some field; wf_prog G;   
+  x2 = None \<longrightarrow> G,s2\<turnstile>a\<Colon>\<preceq>Class Ca; C \<noteq> Object; class G C = Some y;   
+  (x2, s2)\<Colon>\<preceq>(G, L); s1\<le>|s2; inited C (globs s1)\<rbrakk> \<Longrightarrow>  
+  G,s2'\<turnstile>v\<Colon>\<preceq>type field \<and> s2'\<le>|f\<preceq>type field\<Colon>\<preceq>(G, L)"
+apply (unfold assign_conforms_def)
+apply (drule sym)
+apply (clarsimp simp add: fvar_def2)
+apply (drule (9) FVar_lemma1)
+apply (clarsimp)
+apply (drule (2) conforms_globsD [THEN oconf_lconf, THEN lconfD])
+apply clarsimp
+apply (drule (1) rev_gext_objD)
+apply (auto elim!: conforms_upd_gobj)
+done
+
+
+lemma AVar_lemma1: "\<lbrakk>globs s (Inl a) = Some obj;tag obj=Arr ty i; 
+ the_Intg i' in_bounds i; wf_prog G; G\<turnstile>ty.[]\<preceq>Tb.[]; Norm s\<Colon>\<preceq>(G, L)
+\<rbrakk> \<Longrightarrow> G,s\<turnstile>the ((values obj) (Inr (the_Intg i')))\<Colon>\<preceq>Tb"
+apply (erule widen_Array_Array [THEN conf_widen])
+apply  (erule_tac [2] wf_ws_prog)
+apply (drule (1) conforms_globsD [THEN oconf_lconf, THEN lconfD])
+defer apply assumption
+apply (force intro: var_tys_Some_eq [THEN iffD2])
+done
+
+lemma obj_split: "\<And> obj. \<exists> t vs. obj = \<lparr>tag=t,values=vs\<rparr>"
+proof record_split
+  fix tag values more
+  show "\<exists>t vs. \<lparr>tag = tag, values = values, \<dots> = more\<rparr> = \<lparr>tag = t, values = vs\<rparr>"
+    by auto
+qed
+ 
+lemma AVar_lemma: "\<lbrakk>wf_prog G; G\<turnstile>(x1, s1) \<midarrow>e2-\<succ>i\<rightarrow> (x2, s2);  
+  ((v,f), Norm s2') = avar G i a (x2, s2); x1 = None \<longrightarrow> G,s1\<turnstile>a\<Colon>\<preceq>Ta.[];  
+  (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)"
+apply (unfold assign_conforms_def)
+apply (drule sym)
+apply (clarsimp simp add: avar_def2)
+apply (drule (1) conf_gext)
+apply (drule conf_RefTD, clarsimp)
+apply (subgoal_tac "\<exists> t vs. obj = \<lparr>tag=t,values=vs\<rparr>")
+defer
+apply   (rule obj_split)
+apply clarify
+apply (frule obj_ty_widenD)
+apply (auto dest!: widen_Class)
+apply  (force dest: AVar_lemma1)
+apply (auto split add: split_if)
+apply (force elim!: fits_Array dest: gext_objD 
+       intro: var_tys_Some_eq [THEN iffD2] conforms_upd_gobj)
+done
+
+
+section "Call"
+lemma conforms_init_lvars_lemma: "\<lbrakk>wf_prog G;  
+  wf_mhead G P sig mh; 
+  Ball (set lvars) (split (\<lambda>vn. is_type G)); 
+  list_all2 (conf G s) pvs pTsa; G\<turnstile>pTsa[\<preceq>](parTs sig)\<rbrakk> \<Longrightarrow>  
+  G,s\<turnstile>init_vals (table_of lvars)(pars mh[\<mapsto>]pvs)
+      [\<Colon>\<preceq>]table_of lvars(pars mh[\<mapsto>]parTs sig)"
+apply (unfold wf_mhead_def)
+apply clarify
+apply (erule (2) Ball_set_table [THEN lconf_init_vals, THEN lconf_ext_list])
+apply (drule wf_ws_prog)
+apply (erule (2) conf_list_widen)
+done
+
+
+lemma lconf_map_lname [simp]: 
+  "G,s\<turnstile>(lname_case l1 l2)[\<Colon>\<preceq>](lname_case L1 L2)
+   =
+  (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit . l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
+apply (unfold lconf_def)
+apply safe
+apply (case_tac "n")
+apply (force split add: lname.split)+
+done
+
+lemma lconf_map_ename [simp]:
+  "G,s\<turnstile>(ename_case l1 l2)[\<Colon>\<preceq>](ename_case L1 L2)
+   =
+  (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit. l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
+apply (unfold lconf_def)
+apply safe
+apply (force split add: ename.split)+
+done
+
+
+lemma defval_conf1 [rule_format (no_asm), elim]: 
+  "is_type G T \<longrightarrow> (\<exists>v\<in>Some (default_val T): G,s\<turnstile>v\<Colon>\<preceq>T)"
+apply (unfold conf_def)
+apply (induct "T")
+apply (auto intro: prim_ty.induct)
+done
+
+
+lemma conforms_init_lvars: 
+"\<lbrakk>wf_mhead G (pid declC) sig (mhead (mthd dm)); wf_prog G;  
+  list_all2 (conf G s) pvs pTsa; G\<turnstile>pTsa[\<preceq>](parTs sig);  
+  (x, s)\<Colon>\<preceq>(G, L); 
+  methd G declC sig = Some dm;  
+  isrtype G statT;
+  G\<turnstile>invC\<preceq>\<^sub>C declC; 
+  G,s\<turnstile>a'\<Colon>\<preceq>RefT statT;  
+  invmode (mhd sm) e = IntVir \<longrightarrow> a' \<noteq> Null; 
+  invmode (mhd sm) e \<noteq> IntVir \<longrightarrow>  
+       (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
+    \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object);
+  invC  = invocation_class (invmode (mhd sm) e) s a' statT;
+  declC = invocation_declclass G (invmode (mhd sm) e) s a' statT sig;
+  Ball (set (lcls (mbody (mthd dm)))) 
+       (split (\<lambda>vn. is_type G)) 
+ \<rbrakk> \<Longrightarrow>  
+  init_lvars G declC sig (invmode (mhd sm) e) a'  
+  pvs (x,s)\<Colon>\<preceq>(G,\<lambda> k. 
+                (case k of
+                   EName e \<Rightarrow> (case e of 
+                                 VNam v 
+                                  \<Rightarrow> (table_of (lcls (mbody (mthd dm)))
+                                        (pars (mthd dm)[\<mapsto>]parTs sig)) v
+                               | Res \<Rightarrow> Some (resTy (mthd dm)))
+                 | This \<Rightarrow> if (static (mthd sm)) 
+                              then None else Some (Class declC)))"
+apply (simp add: init_lvars_def2)
+apply (rule conforms_set_locals)
+apply  (simp (no_asm_simp) split add: split_if)
+apply (drule  (4) DynT_conf)
+apply clarsimp
+(* apply intro *)
+apply (drule (4) conforms_init_lvars_lemma)
+apply (case_tac "dm",simp)
+apply (rule conjI)
+apply (unfold lconf_def, clarify)
+apply (rule defval_conf1)
+apply (clarsimp simp add: wf_mhead_def is_acc_type_def)
+apply (case_tac "is_static sm")
+apply simp_all
+done
+
+
+lemma Call_type_sound: "\<lbrakk>wf_prog G; G\<turnstile>(x1, s1) \<midarrow>ps\<doteq>\<succ>pvs\<rightarrow> (x2, s2);  
+ declC 
+ = invocation_declclass G (invmode (mhd esm) e) s2 a' statT \<lparr>name=mn,parTs=pTs\<rparr>;
+s2'=init_lvars G declC \<lparr>name=mn,parTs=pTs\<rparr> (invmode (mhd esm) e) a' pvs (x2,s2);
+ G\<turnstile>s2' \<midarrow>Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<rightarrow> (x3, s3);    
+ \<forall>L. s2'\<Colon>\<preceq>(G, L) 
+     \<longrightarrow> (\<forall>T. \<lparr>prg=G,cls=declC,lcl=L\<rparr>\<turnstile> Methd declC \<lparr>name=mn,parTs=pTs\<rparr>\<Colon>-T 
+     \<longrightarrow> (x3, s3)\<Colon>\<preceq>(G, L) \<and> (x3 = None \<longrightarrow> G,s3\<turnstile>v\<Colon>\<preceq>T));  
+ Norm s0\<Colon>\<preceq>(G, L); 
+ \<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;  
+ max_spec G C statT \<lparr>name=mn,parTs=pTsa\<rparr> = {(esm, pTs)}; 
+ (x1, s1)\<Colon>\<preceq>(G, L); 
+ x1 = None \<longrightarrow> G,s1\<turnstile>a'\<Colon>\<preceq>RefT statT; (x2, s2)\<Colon>\<preceq>(G, L);  
+ x2 = None \<longrightarrow> list_all2 (conf G s2) pvs pTsa;
+ sm=(mhd esm)\<rbrakk> \<Longrightarrow>     
+ (x3, set_locals (locals s2) s3)\<Colon>\<preceq>(G, L) \<and> 
+ (x3 = None \<longrightarrow> G,s3\<turnstile>v\<Colon>\<preceq>resTy sm)"
+apply clarify
+apply (case_tac "x2")
+defer
+apply  (clarsimp split add: split_if_asm simp add: init_lvars_def2)
+apply (case_tac "a' = Null \<and> \<not> (static (mhd esm)) \<and> e \<noteq> Super")
+apply  (clarsimp simp add: init_lvars_def2)
+apply clarsimp
+apply (drule eval_gext')
+apply (frule (1) conf_gext)
+apply (drule max_spec2mheads, clarsimp)
+apply (subgoal_tac "invmode (mhd esm) e = IntVir \<longrightarrow> a' \<noteq> Null")
+defer  
+apply  (clarsimp simp add: invmode_IntVir_eq)
+apply (frule (6) DynT_mheadsD [OF DynT_propI,of _ "s2"],(rule HOL.refl)+)
+apply clarify
+apply (drule wf_mdeclD1, clarsimp) 
+apply (frule  ty_expr_is_type) apply simp
+apply (frule (2) conforms_init_lvars)
+apply   simp
+apply   assumption+
+apply   simp
+apply   assumption+
+apply   clarsimp
+apply   (rule HOL.refl)
+apply   simp
+apply   (rule Ball_weaken)
+apply     assumption
+apply     (force simp add: is_acc_type_def)
+apply (tactic "smp_tac 1 1")
+apply (frule (2) wt_MethdI, clarsimp)
+apply (subgoal_tac "is_static dm = (static (mthd esm))") 
+apply   (simp only:)
+apply   (tactic "smp_tac 1 1")
+apply   (rule conjI)
+apply     (erule  conforms_return)
+apply     blast
+
+apply     (force dest!: eval_gext del: impCE simp add: init_lvars_def2)
+apply     clarsimp
+apply     (drule (2) widen_trans, erule (1) conf_widen)
+apply     (erule wf_ws_prog)
+
+apply   auto
+done
+
+
+subsection "accessibility"
+
+theorem dynamic_field_access_ok:
+  (assumes wf: "wf_prog G" and
+       eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2" and
+     not_Null: "a\<noteq>Null" and
+    conform_a: "G,(store s2)\<turnstile>a\<Colon>\<preceq> Class statC" and
+   conform_s2: "s2\<Colon>\<preceq>(G, L)" and 
+    normal_s2: "normal s2" and
+         wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>,dt\<Turnstile>e\<Colon>-Class statC" and
+            f: "accfield G accC statC fn = Some f" and
+         dynC: "if stat then dynC=statC  
+                        else dynC=obj_class (lookup_obj (store s2) a)"
+  ) "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f) \<and> 
+     G\<turnstile>Field fn f in dynC dyn_accessible_from accC"
+proof (cases "stat")
+  case True
+  with dynC 
+  have dynC': "dynC=statC" by simp
+  with f
+  have "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)"
+    by (auto simp add: accfield_def Let_def intro!: table_of_remap_SomeD)
+  with dynC' f
+  show ?thesis
+    by (auto intro!: static_to_dynamic_accessible_from
+         dest: accfield_accessibleD accessible_from_commonD)
+next
+  case False
+  with wf conform_a not_Null conform_s2 dynC
+  obtain subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
+    "is_class G dynC"
+    by (auto dest!: conforms_RefTD [of _ _ _ _ "(fst s2)" L]
+              dest: obj_ty_obj_class1
+          simp add: obj_ty_obj_class )
+  with wf f
+  have "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)"
+    by (auto simp add: accfield_def Let_def
+                 dest: fields_mono
+                dest!: table_of_remap_SomeD)
+  moreover
+  from f subclseq
+  have "G\<turnstile>Field fn f in dynC dyn_accessible_from accC"
+    by (auto intro!: static_to_dynamic_accessible_from 
+               dest: accfield_accessibleD)
+  ultimately show ?thesis
+    by blast
+qed
+
+lemma call_access_ok: 
+(assumes invC_prop: "G\<turnstile>invmode (mhd statM) e\<rightarrow>invC\<preceq>statT" 
+     and        wf: "wf_prog G" 
+     and      wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
+     and     statM: "statM \<in> mheads G accC statT sig" 
+     and      invC: "invC = invocation_class (invmode (mhd statM) e) s a statT"
+)"\<exists> dynM. dynlookup G statT invC sig = Some dynM \<and>
+  G\<turnstile>Methd sig dynM in invC dyn_accessible_from accC"
+proof -
+  from wt_e wf have type_statT: "is_type G (RefT statT)"
+    by (auto dest: ty_expr_is_type)
+  from statM have not_Null: "statT \<noteq> NullT" by auto
+  from type_statT wt_e 
+  have wf_I: "(\<forall>I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> 
+                                        invmode (mhd statM) e \<noteq> SuperM)"
+    by (auto dest: invocationTypeExpr_noClassD)
+  from wt_e
+  have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode (mhd statM) e \<noteq> SuperM)"
+    by (auto dest: invocationTypeExpr_noClassD)
+  show ?thesis
+  proof (cases "invmode (mhd statM) e = IntVir")
+    case True
+    with invC_prop not_Null
+    have invC_prop': "is_class G invC \<and>  
+                      (if (\<exists>T. statT=ArrayT T) then invC=Object 
+                                              else G\<turnstile>Class invC\<preceq>RefT statT)"
+      by (auto simp add: DynT_prop_def)
+    with True not_Null
+    have "G,statT \<turnstile> invC valid_lookup_cls_for is_static statM"
+     by (cases statT) (auto simp add: invmode_def 
+                         split: split_if split_if_asm) (*  was deleted above *)
+    with statM type_statT wf 
+    show ?thesis
+      by - (rule dynlookup_access,auto)
+  next
+    case False
+    with type_statT wf invC not_Null wf_I wf_A
+    have invC_prop': " is_class G invC \<and>
+                      ((\<exists> statC. statT=ClassT statC \<and> invC=statC) \<or>
+                      (\<forall> statC. statT\<noteq>ClassT statC \<and> invC=Object)) "
+        by (case_tac "statT") (auto simp add: invocation_class_def 
+                                       split: inv_mode.splits)
+    with not_Null wf
+    have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig"
+      by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_C_C
+                                            dynimethd_def)
+   from statM wf wt_e not_Null False invC_prop' obtain dynM where
+                "accmethd G accC invC sig = Some dynM" 
+     by (auto dest!: static_mheadsD)
+   from invC_prop' False not_Null wf_I
+   have "G,statT \<turnstile> invC valid_lookup_cls_for is_static statM"
+     by (cases statT) (auto simp add: invmode_def
+                        split: split_if split_if_asm) (*  was deleted above *)
+   with statM type_statT wf 
+    show ?thesis
+      by - (rule dynlookup_access,auto)
+  qed
+qed
+
+section "main proof of type safety"
+
+ML {*
+val forward_hyp_tac = EVERY' [smp_tac 1,
+	FIRST'[mp_tac,etac exI,smp_tac 2,smp_tac 1,EVERY'[etac impE,etac exI]],
+	REPEAT o (etac conjE)];
+val typD_tac = eresolve_tac (thms "wt_elim_cases") THEN_ALL_NEW 
+	EVERY' [full_simp_tac (simpset() setloop (K no_tac)), 
+         clarify_tac(claset() addSEs[])]
+*}
+
+lemma conforms_locals [rule_format]: 
+  "(a,b)\<Colon>\<preceq>(G, L) \<longrightarrow> L x = Some T \<longrightarrow> G,b\<turnstile>the (locals b x)\<Colon>\<preceq>T"
+apply (force simp: conforms_def Let_def lconf_def)
+done
+
+lemma eval_type_sound [rule_format (no_asm)]: 
+ "wf_prog G \<Longrightarrow> G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1) \<Longrightarrow> (\<forall>L. s0\<Colon>\<preceq>(G,L) \<longrightarrow>    
+  (\<forall>C T. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<longrightarrow> s1\<Colon>\<preceq>(G,L) \<and>  
+  (let (x,s) = s1 in x = None \<longrightarrow> G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T)))"
+apply (erule eval_induct)
+
+(* 29 subgoals *)
+(* Xcpt, Inst, Methd, Nil, Skip, Expr, Comp *)
+apply         (simp_all (no_asm_use) add: Let_def body_def)
+apply       (tactic "ALLGOALS (EVERY'[Clarify_tac, TRY o typD_tac, 
+                     TRY o forward_hyp_tac])")
+apply      (tactic"ALLGOALS(EVERY'[asm_simp_tac(simpset()),TRY o Clarify_tac])")
+
+(* 20 subgoals *)
+
+(* Break *)
+apply (erule conforms_absorb)
+
+(* Cons *)
+apply (erule_tac V = "G\<turnstile>Norm s0 \<midarrow>?ea\<succ>\<rightarrow> ?vs1" in thin_rl)
+apply (frule eval_gext')
+apply force
+
+(* LVar *)
+apply (force elim: conforms_localD [THEN lconfD] conforms_lupd 
+       simp add: assign_conforms_def lvar_def)
+
+(* Cast *)
+apply (force dest: fits_conf)
+
+(* Lit *)
+apply (rule conf_litval)
+apply (simp add: empty_dt_def)
+
+(* Super *)
+apply (rule conf_widen)
+apply   (erule (1) subcls_direct [THEN widen.subcls])
+apply  (erule (1) conforms_localD [THEN lconfD])
+apply (erule wf_ws_prog)
+
+(* Acc *)
+apply fast
+
+(* Body *)
+apply (rule conjI)
+apply (rule conforms_absorb)
+apply (fast)
+apply (fast intro: conforms_locals)
+
+(* Cond *)
+apply (simp split: split_if_asm)
+apply  (tactic "forward_hyp_tac 1", force)
+apply (tactic "forward_hyp_tac 1", force)
+
+(* If *)
+apply (force split add: split_if_asm)
+
+(* Loop *)
+apply (drule (1) wt.Loop)
+apply (clarsimp split: split_if_asm)
+apply (fast intro: conforms_absorb)
+
+(* Fin *)
+apply (case_tac "x1", force)
+apply (drule spec, erule impE, erule conforms_NormI)
+apply (erule impE)
+apply   blast
+apply (clarsimp)
+apply (erule (3) Fin_lemma)
+
+(* Throw *)
+apply (erule (3) Throw_lemma)
+
+(* NewC *)
+apply (clarsimp simp add: is_acc_class_def)
+apply (drule (1) halloc_type_sound,blast, rule HOL.refl, simp, simp)
+
+(* NewA *)
+apply (tactic "smp_tac 1 1",frule wt_init_comp_ty,erule impE,blast)
+apply (tactic "forward_hyp_tac 1")
+apply (case_tac "check_neg i' ab")
+apply  (clarsimp simp add: is_acc_type_def)
+apply  (drule (2) halloc_type_sound, rule HOL.refl, simp, simp)
+apply force
+
+(* Level 34, 6 subgoals *)
+
+(* Init *)
+apply (case_tac "inited C (globs s0)")
+apply  (clarsimp)
+apply (clarsimp)
+apply (frule (1) wf_prog_cdecl)
+apply (drule spec, erule impE, erule (3) conforms_init_class_obj)
+apply (drule_tac "psi" = "class G C = ?x" in asm_rl,erule impE,
+      force dest!: wf_cdecl_supD split add: split_if simp add: is_acc_class_def)
+apply (drule spec, erule impE, erule conforms_set_locals, rule lconf_empty)
+apply (erule impE) apply (rule exI) apply (erule wf_cdecl_wt_init)
+apply (drule (1) conforms_return, force dest: eval_gext', assumption)
+
+
+(* Ass *)
+apply (tactic "forward_hyp_tac 1")
+apply (rename_tac x1 s1 x2 s2 v va w L C Ta T', case_tac x1)
+prefer 2 apply force
+apply (case_tac x2)
+prefer 2 apply force
+apply (simp, drule conjunct2)
+apply (drule (1) conf_widen)
+apply  (erule wf_ws_prog)
+apply (erule (2) Ass_lemma)
+apply (clarsimp simp add: assign_conforms_def)
+
+(* Try *)
+apply (drule (1) sxalloc_type_sound, simp (no_asm_use))
+apply (case_tac a)
+apply  clarsimp
+apply clarsimp
+apply (tactic "smp_tac 1 1")
+apply (simp split add: split_if_asm)
+apply (fast dest: conforms_deallocL Try_lemma)
+
+(* FVar *)
+
+apply (frule accfield_fields)
+apply (frule ty_expr_is_type [THEN type_is_class],simp)
+apply simp
+apply (frule wf_ws_prog)
+apply (frule (1) fields_declC,simp)
+apply clarsimp 
+(*b y EVERY'[datac cfield_defpl_is_class 2, Clarsimp_tac] 1; not useful here*)
+apply (tactic "smp_tac 1 1")
+apply (tactic "forward_hyp_tac 1")
+apply (rule conjI, force split add: split_if simp add: fvar_def2)
+apply (drule init_yields_initd, frule eval_gext')
+apply clarsimp
+apply (case_tac "C=Object")
+apply  clarsimp
+apply (erule (9) FVar_lemma)
+
+(* AVar *)
+apply (tactic "forward_hyp_tac 1")
+apply (erule_tac V = "G\<turnstile>Norm s0 \<midarrow>?e1-\<succ>?a'\<rightarrow> (?x1 1, ?s1)" in thin_rl, 
+         frule eval_gext')
+apply (rule conjI)
+apply  (clarsimp simp add: avar_def2)
+apply clarsimp
+apply (erule (5) AVar_lemma)
+
+(* Call *)
+apply (tactic "forward_hyp_tac 1")
+apply (rule Call_type_sound)
+apply auto
+done
+
+
+declare fun_upd_apply [simp]
+declare split_paired_All [simp] split_paired_Ex [simp]
+declare split_if     [split] split_if_asm     [split] 
+        option.split [split] option.split_asm [split]
+ML_setup {* 
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac);
+claset_ref()  := claset () addSbefore ("split_all_tac", split_all_tac)
+*}
+
+theorem eval_ts: 
+ "\<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> 
+\<Longrightarrow>  (x',s')\<Colon>\<preceq>(G,L) \<and> (x'=None \<longrightarrow> G,s'\<turnstile>v\<Colon>\<preceq>T)"
+apply (drule (3) eval_type_sound)
+apply (unfold Let_def)
+apply clarsimp
+done
+
+theorem evals_ts: 
+"\<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> 
+\<Longrightarrow>  (x',s')\<Colon>\<preceq>(G,L) \<and> (x'=None \<longrightarrow> list_all2 (conf G s') vs Ts)"
+apply (drule (3) eval_type_sound)
+apply (unfold Let_def)
+apply clarsimp
+done
+
+theorem evar_ts: 
+"\<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>  
+  (x',s')\<Colon>\<preceq>(G,L) \<and> (x'=None \<longrightarrow> G,L,s'\<turnstile>In2 v\<succ>In2 vf\<Colon>\<preceq>Inl T)"
+apply (drule (3) eval_type_sound)
+apply (unfold Let_def)
+apply clarsimp
+done
+
+theorem exec_ts: 
+"\<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)"
+apply (drule (3) eval_type_sound)
+apply (unfold Let_def)
+apply clarsimp
+done
+
+(*
+theorem dyn_methods_understood: 
+ "\<And>s. \<lbrakk>wf_prog G; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>{t,md,IntVir}e..mn({pTs'}ps)\<Colon>-rT;  
+  s\<Colon>\<preceq>(G,L); G\<turnstile>s \<midarrow>e-\<succ>a'\<rightarrow> Norm s'; a' \<noteq> Null\<rbrakk> \<Longrightarrow>  
+  \<exists>a obj. a'=Addr a \<and> heap s' a = Some obj \<and> 
+  cmethd G (obj_class obj) (mn, pTs') \<noteq> None"
+apply (erule wt_elim_cases)
+apply (drule max_spec2mheads)
+apply (drule (3) eval_ts)
+apply (clarsimp split del: split_if split_if_asm)
+apply (drule (2) DynT_propI)
+apply  (simp (no_asm_simp))
+apply (tactic *) (* {* exhaust_cmethd_tac "the (cmethd G (target (invmode m e) s' a' md) (mn, pTs'))" 1 *} *)(*)
+apply (drule (4) DynT_mheadsD [THEN conjunct1], rule HOL.refl)
+apply (drule conf_RefTD)
+apply clarsimp
+done 
+*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/Value.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,51 @@
+(*  Title:      isabelle/Bali/Value.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* Java values *}
+
+
+
+theory Value = Type:
+
+typedecl loc            (* locations, i.e. abstract references on objects *)
+arities	 loc :: "type"
+
+datatype val
+	= Unit          (* dummy result value of void methods *)
+	| Bool bool     (* Boolean value *)
+	| Intg int      (* integer value *)
+	| Null          (* null reference *)
+	| Addr loc      (* addresses, i.e. locations of objects *)
+
+
+translations "val" <= (type) "Term.val"
+             "loc" <= (type) "Term.loc"
+
+consts   the_Bool   :: "val \<Rightarrow> bool"  
+primrec "the_Bool (Bool b) = b"
+consts   the_Intg   :: "val \<Rightarrow> int"
+primrec "the_Intg (Intg i) = i"
+consts   the_Addr   :: "val \<Rightarrow> loc"
+primrec "the_Addr (Addr a) = a"
+
+types	dyn_ty      = "loc \<Rightarrow> ty option"
+consts
+  typeof        :: "dyn_ty \<Rightarrow> val \<Rightarrow> ty option"
+  defpval       :: "prim_ty \<Rightarrow> val"      (* default value for primitive types *)
+  default_val   :: "     ty \<Rightarrow> val"      (* default value for all types *)
+
+primrec "typeof dt  Unit    = Some (PrimT Void)"
+	"typeof dt (Bool b) = Some (PrimT Boolean)"
+	"typeof dt (Intg i) = Some (PrimT Integer)"
+	"typeof dt  Null    = Some NT"
+	"typeof dt (Addr a) = dt a"
+
+primrec	"defpval Void    = Unit"
+	"defpval Boolean = Bool False"
+	"defpval Integer = Intg 0"
+primrec	"default_val (PrimT pt) = defpval pt"
+	"default_val (RefT  r ) = Null"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/WellForm.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,3004 @@
+(*  Title:      isabelle/Bali/WellForm.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+
+header {* Well-formedness of Java programs *}
+
+theory WellForm = WellType:
+
+text {*
+For static checks on expressions and statements, see WellType.thy
+
+improvements over Java Specification 1.0 (cf. 8.4.6.3, 8.4.6.4, 9.4.1):
+\begin{itemize}
+\item a method implementing or overwriting another method may have a result 
+      type that widens to the result type of the other method 
+      (instead of identical type)
+\item if a method hides another method (both methods have to be static!)
+  there are no restrictions to the result type 
+  since the methods have to be static and there is no dynamic binding of 
+  static methods
+\item if an interface inherits more than one method with the same signature, the
+  methods need not have identical return types
+\end{itemize}
+simplifications:
+\begin{itemize}
+\item Object and standard exceptions are assumed to be declared like normal 
+      classes
+\end{itemize}
+*}
+
+section "well-formed field declarations"
+  (* well-formed field declaration (common part for classes and interfaces),
+     cf. 8.3 and (9.3) *)
+
+constdefs
+  wf_fdecl :: "prog \<Rightarrow> pname \<Rightarrow> fdecl \<Rightarrow> bool"
+ "wf_fdecl G P \<equiv> \<lambda>(fn,f). is_acc_type G P (type f)"
+
+lemma wf_fdecl_def2: "\<And>fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))"
+apply (unfold wf_fdecl_def)
+apply simp
+done
+
+
+
+section "well-formed method declarations"
+  (*well-formed method declaration,cf. 8.4, 8.4.1, 8.4.3, 8.4.5, 14.3.2, (9.4)*)
+  (* cf. 14.15, 15.7.2, for scope issues cf. 8.4.1 and 14.3.2 *)
+
+text {*
+A method head is wellformed if:
+\begin{itemize}
+\item the signature and the method head agree in the number of parameters
+\item all types of the parameters are visible
+\item the result type is visible
+\item the parameter names are unique
+\end{itemize} 
+*}
+constdefs
+  wf_mhead :: "prog \<Rightarrow> pname \<Rightarrow> sig \<Rightarrow> mhead \<Rightarrow> bool"
+ "wf_mhead G P \<equiv> \<lambda> sig mh. length (parTs sig) = length (pars mh) \<and>
+			    \<spacespace> ( \<forall>T\<in>set (parTs sig). is_acc_type G P T) \<and> 
+                            is_acc_type G P (resTy mh) \<and>
+			    \<spacespace> nodups (pars mh)"
+
+
+text {*
+A method declaration is wellformed if:
+\begin{itemize}
+\item the method head is wellformed
+\item the names of the local variables are unique
+\item the types of the local variables must be accessible
+\item the local variables don't shadow the parameters
+\item the class of the method is defined
+\item the body statement is welltyped with respect to the
+      modified environment of local names, were the local variables, 
+      the parameters the special result variable (Res) and This are assoziated
+      with there types. 
+\end{itemize}
+*}
+
+constdefs
+  wf_mdecl :: "prog \<Rightarrow> qtname \<Rightarrow> mdecl \<Rightarrow> bool"
+ "wf_mdecl G C \<equiv> 
+      \<lambda>(sig,m).
+	  wf_mhead G (pid C) sig (mhead m) \<and> 
+          unique (lcls (mbody m)) \<and> 
+          (\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T) \<and> 
+	  (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and>
+          is_class G C \<and>
+          \<lparr>prg=G,cls=C,
+           lcl=\<lambda> k. 
+               (case k of
+                  EName e 
+                  \<Rightarrow> (case e of 
+                        VNam v 
+                        \<Rightarrow>(table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
+                      | Res \<Rightarrow> Some (resTy m))
+	        | This \<Rightarrow> if static m then None else Some (Class C))
+          \<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>"
+
+lemma wf_mheadI: 
+"\<lbrakk>length (parTs sig) = length (pars m); \<forall>T\<in>set (parTs sig). is_acc_type G P T;
+  is_acc_type G P (resTy m); nodups (pars m)\<rbrakk> \<Longrightarrow>  
+  wf_mhead G P sig m"
+apply (unfold wf_mhead_def)
+apply (simp (no_asm_simp))
+done
+
+lemma wf_mdeclI: "\<lbrakk>  
+  wf_mhead G (pid C) sig (mhead m); unique (lcls (mbody m));  
+  (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None); 
+  \<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T;
+  is_class G C;
+  \<lparr>prg=G,cls=C,
+   lcl=\<lambda> k. 
+       (case k of
+          EName e 
+          \<Rightarrow> (case e of 
+                VNam v 
+                \<Rightarrow> (table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
+              | Res \<Rightarrow> Some (resTy m))
+        | This \<Rightarrow> if static m then None else Some (Class C))
+  \<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>
+  \<rbrakk> \<Longrightarrow>  
+  wf_mdecl G C (sig,m)"
+apply (unfold wf_mdecl_def)
+apply simp
+done
+
+
+lemma wf_mdeclD1: 
+"wf_mdecl G C (sig,m) \<Longrightarrow>  
+   wf_mhead G (pid C) sig (mhead m) \<and> unique (lcls (mbody m)) \<and>  
+  (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and> 
+  (\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T)"
+apply (unfold wf_mdecl_def)
+apply auto
+done
+
+lemma wf_mdecl_bodyD: 
+"wf_mdecl G C (sig,m) \<Longrightarrow>  
+ (\<exists>T. \<lparr>prg=G,cls=C,
+       lcl = \<lambda> k. 
+         (case k of
+            EName e 
+            \<Rightarrow> (case e of 
+                VNam v \<Rightarrow> (table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
+                | Res  \<Rightarrow> Some (resTy m))
+          | This \<Rightarrow> if static m then None else Some (Class C))
+       \<rparr>\<turnstile>Body C (stmt (mbody m))\<Colon>-T \<and> G\<turnstile>T\<preceq>(resTy m))"
+apply (unfold wf_mdecl_def)
+apply clarify
+apply (rule_tac x="(resTy m)" in exI)
+apply (unfold wf_mhead_def)
+apply (auto simp add: wf_mhead_def is_acc_type_def intro: wt.Body )
+done
+
+
+(*
+lemma static_Object_methodsE [elim!]: 
+ "\<lbrakk>wf_mdecl G Object (sig, m);static m\<rbrakk> \<Longrightarrow> R"
+apply (unfold wf_mdecl_def)
+apply auto
+done
+*)
+
+lemma rT_is_acc_type: 
+  "wf_mhead G P sig m \<Longrightarrow> is_acc_type G P (resTy m)"
+apply (unfold wf_mhead_def)
+apply auto
+done
+
+section "well-formed interface declarations"
+  (* well-formed interface declaration, cf. 9.1, 9.1.2.1, 9.1.3, 9.4 *)
+
+text {*
+A interface declaration is wellformed if:
+\begin{itemize}
+\item the interface hierarchy is wellstructured
+\item there is no class with the same name
+\item the method heads are wellformed and not static and have Public access
+\item the methods are uniquely named
+\item all superinterfaces are accessible
+\item the result type of a method overriding a method of Object widens to the
+      result type of the overridden method.
+      Shadowing static methods is forbidden.
+\item the result type of a method overriding a set of methods defined in the
+      superinterfaces widens to each of the corresponding result types
+\end{itemize}
+*}
+constdefs
+  wf_idecl :: "prog  \<Rightarrow> idecl \<Rightarrow> bool"
+ "wf_idecl G \<equiv> 
+    \<lambda>(I,i). 
+        ws_idecl G I (isuperIfs i) \<and> 
+	\<not>is_class G I \<and>
+	(\<forall>(sig,mh)\<in>set (imethods i). wf_mhead G (pid I) sig mh \<and> 
+                                     \<not>is_static mh \<and>
+                                      accmodi mh = Public) \<and>
+	unique (imethods i) \<and>
+        (\<forall> J\<in>set (isuperIfs i). is_acc_iface G (pid I) J) \<and>
+        (table_of (imethods i)
+          hiding (methd G Object)
+          under  (\<lambda> new old. accmodi old \<noteq> Private)
+          entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old \<and> 
+                             is_static new = is_static old)) \<and> 
+        (o2s \<circ> table_of (imethods i) 
+               hidings Un_tables((\<lambda>J.(imethds G J))`set (isuperIfs i))
+	       entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old))"
+
+lemma wf_idecl_mhead: "\<lbrakk>wf_idecl G (I,i); (sig,mh)\<in>set (imethods i)\<rbrakk> \<Longrightarrow>  
+  wf_mhead G (pid I) sig mh \<and> \<not>is_static mh \<and> accmodi mh = Public"
+apply (unfold wf_idecl_def)
+apply auto
+done
+
+lemma wf_idecl_hidings: 
+"wf_idecl G (I, i) \<Longrightarrow> 
+  (\<lambda>s. o2s (table_of (imethods i) s)) 
+  hidings Un_tables ((\<lambda>J. imethds G J) ` set (isuperIfs i))  
+  entails \<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old"
+apply (unfold wf_idecl_def o_def)
+apply simp
+done
+
+lemma wf_idecl_hiding:
+"wf_idecl G (I, i) \<Longrightarrow> 
+ (table_of (imethods i)
+           hiding (methd G Object)
+           under  (\<lambda> new old. accmodi old \<noteq> Private)
+           entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old \<and> 
+                              is_static new = is_static old))"
+apply (unfold wf_idecl_def)
+apply simp
+done
+
+lemma wf_idecl_supD: 
+"\<lbrakk>wf_idecl G (I,i); J \<in> set (isuperIfs i)\<rbrakk> 
+ \<Longrightarrow> is_acc_iface G (pid I) J \<and> (J, I) \<notin> (subint1 G)^+"
+apply (unfold wf_idecl_def ws_idecl_def)
+apply auto
+done
+
+section "well-formed class declarations"
+  (* well-formed class declaration, cf. 8.1, 8.1.2.1, 8.1.2.2, 8.1.3, 8.1.4 and
+   class method declaration, cf. 8.4.3.3, 8.4.6.1, 8.4.6.2, 8.4.6.3, 8.4.6.4 *)
+
+text {*
+A class declaration is wellformed if:
+\begin{itemize}
+\item there is no interface with the same name
+\item all superinterfaces are accessible and for all methods implementing 
+      an interface method the result type widens to the result type of 
+      the interface method, the method is not static and offers at least 
+      as much access 
+      (this actually means that the method has Public access, since all 
+      interface methods have public access)
+\item all field declarations are wellformed and the field names are unique
+\item all method declarations are wellformed and the method names are unique
+\item the initialization statement is welltyped
+\item the classhierarchy is wellstructured
+\item Unless the class is Object:
+      \begin{itemize}
+      \item the superclass is accessible
+      \item for all methods overriding another method (of a superclass )the
+            result type widens to the result type of the overridden method,
+            the access modifier of the new method provides at least as much
+            access as the overwritten one.
+      \item for all methods hiding a method (of a superclass) the hidden 
+            method must be static and offer at least as much access rights.
+            Remark: In contrast to the Java Language Specification we don't
+            restrict the result types of the method
+            (as in case of overriding), because there seems to be no reason,
+            since there is no dynamic binding of static methods.
+            (cf. 8.4.6.3 vs. 15.12.1).
+            Stricly speaking the restrictions on the access rights aren't 
+            necessary to, since the static type and the access rights 
+            together determine which method is to be called statically. 
+            But if a class gains more then one static method with the 
+            same signature due to inheritance, it is confusing when the 
+            method selection depends on the access rights only: 
+            e.g.
+              Class C declares static public method foo().
+              Class D is subclass of C and declares static method foo()
+              with default package access.
+              D.foo() ? if this call is in the same package as D then
+                        foo of class D is called, otherwise foo of class C.
+      \end{itemize}
+
+\end{itemize}
+*}
+(* to Table *)
+constdefs entails:: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool"
+                                 ("_ entails _" 20)
+"t entails P \<equiv> \<forall>k. \<forall> x \<in> t k: P x"
+
+lemma entailsD:
+ "\<lbrakk>t entails P; t k = Some x\<rbrakk> \<Longrightarrow> P x"
+by (simp add: entails_def)
+
+lemma empty_entails[simp]: "empty entails P"
+by (simp add: entails_def)
+
+constdefs
+ wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool"
+"wf_cdecl G \<equiv> 
+   \<lambda>(C,c).
+      \<not>is_iface G C \<and>
+      (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
+        (\<forall>s. \<forall> im \<in> imethds G I s.
+      	    (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
+      	                             \<not> is_static cm \<and>
+                                     accmodi im \<le> accmodi cm))) \<and>
+      (\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and> 
+      (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and> 
+      \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd> \<and> ws_cdecl G C (super c) \<and>
+      (C \<noteq> Object \<longrightarrow> 
+            (is_acc_class G (pid C) (super c) \<and>
+            (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
+             entails (\<lambda> new. \<forall> old sig. 
+                       (G,sig\<turnstile>new overrides\<^sub>S old 
+                        \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
+                             accmodi old \<le> accmodi new \<and>
+      	                     \<not>is_static old)) \<and>
+                       (G,sig\<turnstile>new hides old 
+                         \<longrightarrow> (accmodi old \<le> accmodi new \<and>
+      	                      is_static old)))) 
+            ))"
+
+(*
+constdefs
+ wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool"
+"wf_cdecl G \<equiv> 
+   \<lambda>(C,c).
+      \<not>is_iface G C \<and>
+      (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
+        (\<forall>s. \<forall> im \<in> imethds G I s.
+      	    (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy (mthd cm)\<preceq>resTy (mthd im) \<and>
+      	                             \<not> is_static cm \<and>
+                                     accmodi im \<le> accmodi cm))) \<and>
+      (\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and> 
+      (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and> 
+      \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd> \<and> ws_cdecl G C (super c) \<and>
+      (C \<noteq> Object \<longrightarrow> 
+            (is_acc_class G (pid C) (super c) \<and>
+            (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
+              hiding methd G (super c)
+              under (\<lambda> new old. G\<turnstile>new overrides old)
+              entails (\<lambda> new old. 
+                           (G\<turnstile>resTy (mthd new)\<preceq>resTy (mthd old) \<and>
+                            accmodi old \<le> accmodi new \<and>
+      	                   \<not> is_static old)))  \<and>
+            (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
+              hiding methd G (super c)
+              under (\<lambda> new old. G\<turnstile>new hides old)
+              entails (\<lambda> new old. is_static old \<and> 
+                                  accmodi old \<le> accmodi new))  \<and>
+            (table_of (cfields c) hiding accfield G C (super c)
+              entails (\<lambda> newF oldF. accmodi oldF \<le> access newF))))"
+*)
+
+lemma wf_cdecl_unique: 
+"wf_cdecl G (C,c) \<Longrightarrow> unique (cfields c) \<and> unique (methods c)"
+apply (unfold wf_cdecl_def)
+apply auto
+done
+
+lemma wf_cdecl_fdecl: 
+"\<lbrakk>wf_cdecl G (C,c); f\<in>set (cfields c)\<rbrakk> \<Longrightarrow> wf_fdecl G (pid C) f"
+apply (unfold wf_cdecl_def)
+apply auto
+done
+
+lemma wf_cdecl_mdecl: 
+"\<lbrakk>wf_cdecl G (C,c); m\<in>set (methods c)\<rbrakk> \<Longrightarrow> wf_mdecl G C m"
+apply (unfold wf_cdecl_def)
+apply auto
+done
+
+lemma wf_cdecl_impD: 
+"\<lbrakk>wf_cdecl G (C,c); I\<in>set (superIfs c)\<rbrakk> 
+\<Longrightarrow> is_acc_iface G (pid C) I \<and>  
+    (\<forall>s. \<forall>im \<in> imethds G I s.  
+        (\<exists>cm \<in> methd G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and> \<not>is_static cm \<and>
+                                   accmodi im \<le> accmodi cm))"
+apply (unfold wf_cdecl_def)
+apply auto
+done
+
+lemma wf_cdecl_supD: 
+"\<lbrakk>wf_cdecl G (C,c); C \<noteq> Object\<rbrakk> \<Longrightarrow>  
+  is_acc_class G (pid C) (super c) \<and> (super c,C) \<notin> (subcls1 G)^+ \<and> 
+   (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
+    entails (\<lambda> new. \<forall> old sig. 
+                 (G,sig\<turnstile>new overrides\<^sub>S old 
+                  \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
+                       accmodi old \<le> accmodi new \<and>
+                       \<not>is_static old)) \<and>
+                 (G,sig\<turnstile>new hides old 
+                   \<longrightarrow> (accmodi old \<le> accmodi new \<and>
+                        is_static old))))"
+apply (unfold wf_cdecl_def ws_cdecl_def)
+apply auto
+done
+
+
+lemma wf_cdecl_overrides_SomeD:
+"\<lbrakk>wf_cdecl G (C,c); C \<noteq> Object; table_of (methods c) sig = Some newM;
+  G,sig\<turnstile>(C,newM) overrides\<^sub>S old
+\<rbrakk> \<Longrightarrow>  G\<turnstile>resTy newM\<preceq>resTy old \<and>
+       accmodi old \<le> accmodi newM \<and>
+       \<not> is_static old" 
+apply (drule (1) wf_cdecl_supD)
+apply (clarify)
+apply (drule entailsD)
+apply   (blast intro: table_of_map_SomeI)
+apply (drule_tac x="old" in spec)
+apply (auto dest: overrides_eq_sigD simp add: msig_def)
+done
+
+lemma wf_cdecl_hides_SomeD:
+"\<lbrakk>wf_cdecl G (C,c); C \<noteq> Object; table_of (methods c) sig = Some newM;
+  G,sig\<turnstile>(C,newM) hides old
+\<rbrakk> \<Longrightarrow>  accmodi old \<le> access newM \<and>
+       is_static old" 
+apply (drule (1) wf_cdecl_supD)
+apply (clarify)
+apply (drule entailsD)
+apply   (blast intro: table_of_map_SomeI)
+apply (drule_tac x="old" in spec)
+apply (auto dest: hides_eq_sigD simp add: msig_def)
+done
+
+lemma wf_cdecl_wt_init: 
+ "wf_cdecl G (C, c) \<Longrightarrow> \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>init c\<Colon>\<surd>"
+apply (unfold wf_cdecl_def)
+apply auto
+done
+
+
+section "well-formed programs"
+  (* well-formed program, cf. 8.1, 9.1 *)
+
+text {*
+A program declaration is wellformed if:
+\begin{itemize}
+\item the class ObjectC of Object is defined
+\item every method of has an access modifier distinct from Package. This is
+      necessary since every interface automatically inherits from Object.  
+      We must know, that every time a Object method is "overriden" by an 
+      interface method this is also overriden by the class implementing the
+      the interface (see @{text "implement_dynmethd and class_mheadsD"})
+\item all standard Exceptions are defined
+\item all defined interfaces are wellformed
+\item all defined classes are wellformed
+\end{itemize}
+*}
+constdefs
+  wf_prog  :: "prog \<Rightarrow> bool"
+ "wf_prog G \<equiv> let is = ifaces G; cs = classes G in
+	         ObjectC \<in> set cs \<and> 
+                (\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package) \<and>
+                (\<forall>xn. SXcptC xn \<in> set cs) \<and>
+		(\<forall>i\<in>set is. wf_idecl G i) \<and> unique is \<and>
+		(\<forall>c\<in>set cs. wf_cdecl G c) \<and> unique cs"
+
+lemma wf_prog_idecl: "\<lbrakk>iface G I = Some i; wf_prog G\<rbrakk> \<Longrightarrow> wf_idecl G (I,i)"
+apply (unfold wf_prog_def Let_def)
+apply simp
+apply (fast dest: map_of_SomeD)
+done
+
+lemma wf_prog_cdecl: "\<lbrakk>class G C = Some c; wf_prog G\<rbrakk> \<Longrightarrow> wf_cdecl G (C,c)"
+apply (unfold wf_prog_def Let_def)
+apply simp
+apply (fast dest: map_of_SomeD)
+done
+
+lemma wf_prog_Object_mdecls:
+"wf_prog G \<Longrightarrow> (\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package)"
+apply (unfold wf_prog_def Let_def)
+apply simp
+done
+
+lemma wf_prog_acc_superD:
+ "\<lbrakk>wf_prog G; class G C = Some c; C \<noteq> Object \<rbrakk> 
+  \<Longrightarrow> is_acc_class G (pid C) (super c)"
+by (auto dest: wf_prog_cdecl wf_cdecl_supD)
+
+lemma wf_ws_prog [elim!,simp]: "wf_prog G \<Longrightarrow> ws_prog G"
+apply (unfold wf_prog_def Let_def)
+apply (rule ws_progI)
+apply  (simp_all (no_asm))
+apply  (auto simp add: is_acc_class_def is_acc_iface_def 
+             dest!: wf_idecl_supD wf_cdecl_supD )+
+done
+
+lemma class_Object [simp]: 
+"wf_prog G \<Longrightarrow> 
+  class G Object = Some \<lparr>access=Public,cfields=[],methods=Object_mdecls,
+                                  init=Skip,super=arbitrary,superIfs=[]\<rparr>"
+apply (unfold wf_prog_def Let_def ObjectC_def)
+apply (fast dest!: map_of_SomeI)
+done
+
+lemma methd_Object[simp]: "wf_prog G \<Longrightarrow> methd G Object =  
+  table_of (map (\<lambda>(s,m). (s, Object, m)) Object_mdecls)"
+apply (subst methd_rec)
+apply (auto simp add: Let_def)
+done
+
+lemma wf_prog_Object_methd:
+"\<lbrakk>wf_prog G; methd G Object sig = Some m\<rbrakk> \<Longrightarrow> accmodi m \<noteq> Package"
+by (auto dest!: wf_prog_Object_mdecls) (auto dest!: map_of_SomeD) 
+
+lemma wf_prog_Object_is_public[intro]:
+ "wf_prog G \<Longrightarrow> is_public G Object"
+by (auto simp add: is_public_def dest: class_Object)
+
+lemma class_SXcpt [simp]: 
+"wf_prog G \<Longrightarrow> 
+  class G (SXcpt xn) = Some \<lparr>access=Public,cfields=[],methods=SXcpt_mdecls,
+                                   init=Skip,
+                                   super=if xn = Throwable then Object 
+                                                           else SXcpt Throwable,
+                                   superIfs=[]\<rparr>"
+apply (unfold wf_prog_def Let_def SXcptC_def)
+apply (fast dest!: map_of_SomeI)
+done
+
+lemma wf_ObjectC [simp]: 
+	"wf_cdecl G ObjectC = (\<not>is_iface G Object \<and> Ball (set Object_mdecls)
+  (wf_mdecl G Object) \<and> unique Object_mdecls)"
+apply (unfold wf_cdecl_def ws_cdecl_def ObjectC_def)
+apply (simp (no_asm))
+done
+
+lemma Object_is_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_class G Object"
+apply (simp (no_asm_simp))
+done
+ 
+lemma Object_is_acc_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_acc_class G S Object"
+apply (simp (no_asm_simp) add: is_acc_class_def is_public_def
+                               accessible_in_RefT_simp)
+done
+
+lemma SXcpt_is_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_class G (SXcpt xn)"
+apply (simp (no_asm_simp))
+done
+
+lemma SXcpt_is_acc_class [simp,elim!]: 
+"wf_prog G \<Longrightarrow> is_acc_class G S (SXcpt xn)"
+apply (simp (no_asm_simp) add: is_acc_class_def is_public_def
+                               accessible_in_RefT_simp)
+done
+
+lemma fields_Object [simp]: "wf_prog G \<Longrightarrow> DeclConcepts.fields G Object = []"
+by (force intro: fields_emptyI)
+
+lemma accfield_Object [simp]: 
+ "wf_prog G \<Longrightarrow> accfield G S Object = empty"
+apply (unfold accfield_def)
+apply (simp (no_asm_simp) add: Let_def)
+done
+
+lemma fields_Throwable [simp]: 
+ "wf_prog G \<Longrightarrow> DeclConcepts.fields G (SXcpt Throwable) = []"
+by (force intro: fields_emptyI)
+
+lemma fields_SXcpt [simp]: "wf_prog G \<Longrightarrow> DeclConcepts.fields G (SXcpt xn) = []"
+apply (case_tac "xn = Throwable")
+apply  (simp (no_asm_simp))
+by (force intro: fields_emptyI)
+
+lemmas widen_trans = ws_widen_trans [OF _ _ wf_ws_prog, elim]
+lemma widen_trans2 [elim]: "\<lbrakk>G\<turnstile>U\<preceq>T; G\<turnstile>S\<preceq>U; wf_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>S\<preceq>T"
+apply (erule (2) widen_trans)
+done
+
+lemma Xcpt_subcls_Throwable [simp]: 
+"wf_prog G \<Longrightarrow> G\<turnstile>SXcpt xn\<preceq>\<^sub>C SXcpt Throwable"
+apply (rule SXcpt_subcls_Throwable_lemma)
+apply auto
+done
+
+lemma unique_fields: 
+ "\<lbrakk>is_class G C; wf_prog G\<rbrakk> \<Longrightarrow> unique (DeclConcepts.fields G C)"
+apply (erule ws_unique_fields)
+apply  (erule wf_ws_prog)
+apply (erule (1) wf_prog_cdecl [THEN wf_cdecl_unique [THEN conjunct1]])
+done
+
+lemma fields_mono: 
+"\<lbrakk>table_of (DeclConcepts.fields G C) fn = Some f; G\<turnstile>D\<preceq>\<^sub>C C; 
+  is_class G D; wf_prog G\<rbrakk> 
+   \<Longrightarrow> table_of (DeclConcepts.fields G D) fn = Some f"
+apply (rule map_of_SomeI)
+apply  (erule (1) unique_fields)
+apply (erule (1) map_of_SomeD [THEN fields_mono_lemma])
+apply (erule wf_ws_prog)
+done
+
+
+lemma fields_is_type [elim]: 
+"\<lbrakk>table_of (DeclConcepts.fields G C) m = Some f; wf_prog G; is_class G C\<rbrakk> \<Longrightarrow> 
+      is_type G (type f)"
+apply (frule wf_ws_prog)
+apply (force dest: fields_declC [THEN conjunct1] 
+                   wf_prog_cdecl [THEN wf_cdecl_fdecl]
+             simp add: wf_fdecl_def2 is_acc_type_def)
+done
+
+lemma imethds_wf_mhead [rule_format (no_asm)]: 
+"\<lbrakk>m \<in> imethds G I sig; wf_prog G; is_iface G I\<rbrakk> \<Longrightarrow>  
+  wf_mhead G (pid (decliface m)) sig (mthd m) \<and> 
+  \<not> is_static m \<and> accmodi m = Public"
+apply (frule wf_ws_prog)
+apply (drule (2) imethds_declI [THEN conjunct1])
+apply clarify
+apply (frule_tac I="(decliface m)" in wf_prog_idecl,assumption)
+apply (drule wf_idecl_mhead)
+apply (erule map_of_SomeD)
+apply (cases m, simp)
+done
+
+lemma methd_wf_mdecl: 
+ "\<lbrakk>methd G C sig = Some m; wf_prog G; class G C = Some y\<rbrakk> \<Longrightarrow>  
+  G\<turnstile>C\<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m) \<and> 
+  wf_mdecl G (declclass m) (sig,(mthd m))"
+apply (frule wf_ws_prog)
+apply (drule (1) methd_declC)
+apply  fast
+apply clarsimp
+apply (frule (1) wf_prog_cdecl, erule wf_cdecl_mdecl, erule map_of_SomeD)
+done
+
+(*
+This lemma doesn't hold!
+lemma methd_rT_is_acc_type: 
+"\<lbrakk>wf_prog G;methd G C C sig = Some (D,m);
+    class G C = Some y\<rbrakk>
+\<Longrightarrow> is_acc_type G (pid C) (resTy m)"
+The result Type is only visible in the scope of defining class D 
+"is_vis_type G (pid D) (resTy m)" but not necessarily in scope of class C!
+(The same is true for the type of pramaters of a method)
+*)
+
+
+lemma methd_rT_is_type: 
+"\<lbrakk>wf_prog G;methd G C sig = Some m;
+    class G C = Some y\<rbrakk>
+\<Longrightarrow> is_type G (resTy m)"
+apply (drule (2) methd_wf_mdecl)
+apply clarify
+apply (drule wf_mdeclD1)
+apply clarify
+apply (drule rT_is_acc_type)
+apply (cases m, simp add: is_acc_type_def)
+done
+
+lemma accmethd_rT_is_type:
+"\<lbrakk>wf_prog G;accmethd G S C sig = Some m;
+    class G C = Some y\<rbrakk>
+\<Longrightarrow> is_type G (resTy m)"
+by (auto simp add: accmethd_def  
+         intro: methd_rT_is_type)
+
+lemma methd_Object_SomeD:
+"\<lbrakk>wf_prog G;methd G Object sig = Some m\<rbrakk> 
+ \<Longrightarrow> declclass m = Object"
+by (auto dest: class_Object simp add: methd_rec )
+
+lemma wf_imethdsD: 
+ "\<lbrakk>im \<in> imethds G I sig;wf_prog G; is_iface G I\<rbrakk> 
+ \<Longrightarrow> \<not>is_static im \<and> accmodi im = Public"
+proof -
+  assume asm: "wf_prog G" "is_iface G I" "im \<in> imethds G I sig"
+  have "wf_prog G \<longrightarrow> 
+         (\<forall> i im. iface G I = Some i \<longrightarrow> im \<in> imethds G I sig
+                  \<longrightarrow> \<not>is_static im \<and> accmodi im = Public)" (is "?P G I")
+  proof (rule iface_rec.induct,intro allI impI)
+    fix G I i im
+    assume hyp: "\<forall> J i. J \<in> set (isuperIfs i) \<and> ws_prog G \<and> iface G I = Some i
+                 \<longrightarrow> ?P G J"
+    assume wf: "wf_prog G" and if_I: "iface G I = Some i" and 
+           im: "im \<in> imethds G I sig" 
+    show "\<not>is_static im \<and> accmodi im = Public" 
+    proof -
+      let ?inherited = "Un_tables (imethds G ` set (isuperIfs i))"
+      let ?new = "(o2s \<circ> table_of (map (\<lambda>(s, mh). (s, I, mh)) (imethods i)))"
+      from if_I wf im have imethds:"im \<in> (?inherited \<oplus>\<oplus> ?new) sig"
+	by (simp add: imethds_rec)
+      from wf if_I have 
+	wf_supI: "\<forall> J. J \<in> set (isuperIfs i) \<longrightarrow> (\<exists> j. iface G J = Some j)"
+	by (blast dest: wf_prog_idecl wf_idecl_supD is_acc_ifaceD)
+      from wf if_I have
+	"\<forall> im \<in> set (imethods i). \<not> is_static im \<and> accmodi im = Public"
+	by (auto dest!: wf_prog_idecl wf_idecl_mhead)
+      then have new_ok: "\<forall> im. table_of (imethods i) sig = Some im 
+                         \<longrightarrow>  \<not> is_static im \<and> accmodi im = Public"
+	by (auto dest!: table_of_Some_in_set)
+      show ?thesis
+	proof (cases "?new sig = {}")
+	  case True
+	  from True wf wf_supI if_I imethds hyp 
+	  show ?thesis by (auto simp del:  split_paired_All)  
+	next
+	  case False
+	  from False wf wf_supI if_I imethds new_ok hyp 
+	  show ?thesis by (auto dest: wf_idecl_hidings hidings_entailsD)
+	qed
+      qed
+    qed
+  with asm show ?thesis by (auto simp del: split_paired_All)
+qed
+
+lemma wf_prog_hidesD:
+ (assumes hides: "G \<turnstile>new hides old" and wf: "wf_prog G") 
+   "accmodi old \<le> accmodi new \<and>
+    is_static old"
+proof -
+  from hides 
+  obtain c where 
+    clsNew: "class G (declclass new) = Some c" and
+    neqObj: "declclass new \<noteq> Object"
+    by (auto dest: hidesD declared_in_classD)
+  with hides obtain newM oldM where
+    newM: "table_of (methods c) (msig new) = Some newM" and 
+     new: "new = (declclass new,(msig new),newM)" and
+     old: "old = (declclass old,(msig old),oldM)" and
+          "msig new = msig old"
+    by (cases new,cases old) 
+       (auto dest: hidesD 
+         simp add: cdeclaredmethd_def declared_in_def)
+  with hides 
+  have hides':
+        "G,(msig new)\<turnstile>(declclass new,newM) hides (declclass old,oldM)"
+    by auto
+  from clsNew wf 
+  have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl)
+  note wf_cdecl_hides_SomeD [OF this neqObj newM hides']
+  with new old 
+  show ?thesis
+    by (cases new, cases old) auto
+qed
+
+text {* Compare this lemma about static  
+overriding @{term "G \<turnstile>new overrides\<^sub>S old"} with the definition of 
+dynamic overriding @{term "G \<turnstile>new overrides old"}. 
+Conforming result types and restrictions on the access modifiers of the old 
+and the new method are not part of the predicate for static overriding. But
+they are enshured in a wellfromed program.  Dynamic overriding has 
+no restrictions on the access modifiers but enforces confrom result types 
+as precondition. But with some efford we can guarantee the access modifier
+restriction for dynamic overriding, too. See lemma 
+@{text wf_prog_dyn_override_prop}.
+*}
+lemma wf_prog_stat_overridesD:
+ (assumes stat_override: "G \<turnstile>new overrides\<^sub>S old" and wf: "wf_prog G") 
+   "G\<turnstile>resTy new\<preceq>resTy old \<and>
+    accmodi old \<le> accmodi new \<and>
+    \<not> is_static old"
+proof -
+  from stat_override 
+  obtain c where 
+    clsNew: "class G (declclass new) = Some c" and
+    neqObj: "declclass new \<noteq> Object"
+    by (auto dest: stat_overrides_commonD declared_in_classD)
+  with stat_override obtain newM oldM where
+    newM: "table_of (methods c) (msig new) = Some newM" and 
+     new: "new = (declclass new,(msig new),newM)" and
+     old: "old = (declclass old,(msig old),oldM)" and
+          "msig new = msig old"
+    by (cases new,cases old) 
+       (auto dest: stat_overrides_commonD 
+         simp add: cdeclaredmethd_def declared_in_def)
+  with stat_override 
+  have stat_override':
+        "G,(msig new)\<turnstile>(declclass new,newM) overrides\<^sub>S (declclass old,oldM)"
+    by auto
+  from clsNew wf 
+  have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl)
+  note wf_cdecl_overrides_SomeD [OF this neqObj newM stat_override']
+  with new old 
+  show ?thesis
+    by (cases new, cases old) auto
+qed
+    
+lemma static_to_dynamic_overriding: 
+  (assumes stat_override: "G\<turnstile>new overrides\<^sub>S old" and wf : "wf_prog G")
+   "G\<turnstile>new overrides old"
+proof -
+  from stat_override 
+  show ?thesis (is "?Overrides new old")
+  proof (induct)
+    case (Direct new old superNew)
+    then have stat_override:"G\<turnstile>new overrides\<^sub>S old" 
+      by (rule stat_overridesR.Direct)
+    from stat_override wf
+    have resTy_widen: "G\<turnstile>resTy new\<preceq>resTy old" and
+      not_static_old: "\<not> is_static old" 
+      by (auto dest: wf_prog_stat_overridesD)  
+    have not_private_new: "accmodi new \<noteq> Private"
+    proof -
+      from stat_override 
+      have "accmodi old \<noteq> Private"
+	by (rule no_Private_stat_override)
+      moreover
+      from stat_override wf
+      have "accmodi old \<le> accmodi new"
+	by (auto dest: wf_prog_stat_overridesD)
+      ultimately
+      show ?thesis
+	by (auto dest: acc_modi_bottom)
+    qed
+    with Direct resTy_widen not_static_old 
+    show "?Overrides new old" 
+      by (auto intro: overridesR.Direct) 
+  next
+    case (Indirect inter new old)
+    then show "?Overrides new old" 
+      by (blast intro: overridesR.Indirect) 
+  qed
+qed
+
+lemma non_Package_instance_method_inheritance:
+(assumes old_inheritable: "G\<turnstile>Method old inheritable_in (pid C)" and
+             accmodi_old: "accmodi old \<noteq> Package" and 
+         instance_method: "\<not> is_static old" and
+                  subcls: "G\<turnstile>C \<prec>\<^sub>C declclass old" and
+            old_declared: "G\<turnstile>Method old declared_in (declclass old)" and
+                      wf: "wf_prog G"
+) "G\<turnstile>Method old member_of C \<or>
+   (\<exists> new. G\<turnstile> new overrides\<^sub>S old \<and>  G\<turnstile>Method new member_of C)"
+proof -
+  from wf have ws: "ws_prog G" by auto
+  from old_declared have iscls_declC_old: "is_class G (declclass old)"
+    by (auto simp add: declared_in_def cdeclaredmethd_def)
+  from subcls have  iscls_C: "is_class G C"
+    by (blast dest:  subcls_is_class)
+  from iscls_C ws old_inheritable subcls 
+  show ?thesis (is "?P C old")
+  proof (induct rule: ws_class_induct')
+    case Object
+    assume "G\<turnstile>Object\<prec>\<^sub>C declclass old"
+    then show "?P Object old"
+      by blast
+  next
+    case (Subcls C c)
+    assume cls_C: "class G C = Some c" and 
+       neq_C_Obj: "C \<noteq> Object" and
+             hyp: "\<lbrakk>G \<turnstile>Method old inheritable_in pid (super c); 
+                   G\<turnstile>super c\<prec>\<^sub>C declclass old\<rbrakk> \<Longrightarrow> ?P (super c) old" and
+     inheritable: "G \<turnstile>Method old inheritable_in pid C" and
+         subclsC: "G\<turnstile>C\<prec>\<^sub>C declclass old"
+    from cls_C neq_C_Obj  
+    have super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c" 
+      by (rule subcls1I)
+    from wf cls_C neq_C_Obj
+    have accessible_super: "G\<turnstile>(Class (super c)) accessible_in (pid C)" 
+      by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
+    {
+      fix old
+      assume    member_super: "G\<turnstile>Method old member_of (super c)"
+      assume     inheritable: "G \<turnstile>Method old inheritable_in pid C"
+      assume instance_method: "\<not> is_static old"
+      from member_super
+      have old_declared: "G\<turnstile>Method old declared_in (declclass old)"
+       by (cases old) (auto dest: member_of_declC)
+      have "?P C old"
+      proof (cases "G\<turnstile>mid (msig old) undeclared_in C")
+	case True
+	with inheritable super accessible_super member_super
+	have "G\<turnstile>Method old member_of C"
+	  by (cases old) (auto intro: members.Inherited)
+	then show ?thesis
+	  by auto
+      next
+	case False
+	then obtain new_member where
+	     "G\<turnstile>new_member declared_in C" and
+             "mid (msig old) = memberid new_member"
+          by (auto dest: not_undeclared_declared)
+	then obtain new where
+	          new: "G\<turnstile>Method new declared_in C" and
+               eq_sig: "msig old = msig new" and
+	    declC_new: "declclass new = C" 
+	  by (cases new_member) auto
+	then have member_new: "G\<turnstile>Method new member_of C"
+	  by (cases new) (auto intro: members.Immediate)
+	from declC_new super member_super
+	have subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old"
+	  by (auto dest!: member_of_subclseq_declC
+	            dest: r_into_trancl intro: trancl_rtrancl_trancl)
+	show ?thesis
+	proof (cases "is_static new")
+	  case False
+	  with eq_sig declC_new new old_declared inheritable
+	       super member_super subcls_new_old
+	  have "G\<turnstile>new overrides\<^sub>S old"
+	    by (auto intro!: stat_overridesR.Direct)
+	  with member_new show ?thesis
+	    by blast
+	next
+	  case True
+	  with eq_sig declC_new subcls_new_old new old_declared inheritable
+	  have "G\<turnstile>new hides old"
+	    by (auto intro: hidesI)    
+	  with wf 
+	  have "is_static old"
+	    by (blast dest: wf_prog_hidesD)
+	  with instance_method
+	  show ?thesis
+	    by (contradiction)
+	qed
+      qed
+    } note hyp_member_super = this
+    from subclsC cls_C 
+    have "G\<turnstile>(super c)\<preceq>\<^sub>C declclass old"
+      by (rule subcls_superD)
+    then
+    show "?P C old"
+    proof (cases rule: subclseq_cases) 
+      case Eq
+      assume "super c = declclass old"
+      with old_declared 
+      have "G\<turnstile>Method old member_of (super c)" 
+	by (cases old) (auto intro: members.Immediate)
+      with inheritable instance_method 
+      show ?thesis
+	by (blast dest: hyp_member_super)
+    next
+      case Subcls
+      assume "G\<turnstile>super c\<prec>\<^sub>C declclass old"
+      moreover
+      from inheritable accmodi_old
+      have "G \<turnstile>Method old inheritable_in pid (super c)"
+	by (cases "accmodi old") (auto simp add: inheritable_in_def)
+      ultimately
+      have "?P (super c) old"
+	by (blast dest: hyp)
+      then show ?thesis
+      proof
+	assume "G \<turnstile>Method old member_of super c"
+	with inheritable instance_method
+	show ?thesis
+	  by (blast dest: hyp_member_super)
+      next
+	assume "\<exists>new. G \<turnstile> new overrides\<^sub>S old \<and> G \<turnstile>Method new member_of super c"
+	then obtain super_new where
+	  super_new_override:  "G \<turnstile> super_new overrides\<^sub>S old" and
+            super_new_member:  "G \<turnstile>Method super_new member_of super c"
+	  by blast
+	from super_new_override wf
+	have "accmodi old \<le> accmodi super_new"
+	  by (auto dest: wf_prog_stat_overridesD)
+	with inheritable accmodi_old
+	have "G \<turnstile>Method super_new inheritable_in pid C"
+	  by (auto simp add: inheritable_in_def 
+	              split: acc_modi.splits
+                       dest: acc_modi_le_Dests)
+	moreover
+	from super_new_override 
+	have "\<not> is_static super_new"
+	  by (auto dest: stat_overrides_commonD)
+	moreover
+	note super_new_member
+	ultimately have "?P C super_new"
+	  by (auto dest: hyp_member_super)
+	then show ?thesis
+	proof 
+	  assume "G \<turnstile>Method super_new member_of C"
+	  with super_new_override
+	  show ?thesis
+	    by blast
+	next
+	  assume "\<exists>new. G \<turnstile> new overrides\<^sub>S super_new \<and> 
+                  G \<turnstile>Method new member_of C"
+	  with super_new_override show ?thesis
+	    by (blast intro: stat_overridesR.Indirect) 
+	qed
+      qed
+    qed
+  qed
+qed
+
+lemma non_Package_instance_method_inheritance_cases [consumes 6,
+         case_names Inheritance Overriding]:
+(assumes old_inheritable: "G\<turnstile>Method old inheritable_in (pid C)" and
+             accmodi_old: "accmodi old \<noteq> Package" and 
+         instance_method: "\<not> is_static old" and
+                  subcls: "G\<turnstile>C \<prec>\<^sub>C declclass old" and
+            old_declared: "G\<turnstile>Method old declared_in (declclass old)" and
+                      wf: "wf_prog G" and
+             inheritance: "G\<turnstile>Method old member_of C \<Longrightarrow> P" and
+             overriding:  "\<And> new. 
+                           \<lbrakk>G\<turnstile> new overrides\<^sub>S old;G\<turnstile>Method new member_of C\<rbrakk>
+                           \<Longrightarrow> P"
+) "P"
+proof -
+  from old_inheritable accmodi_old instance_method subcls old_declared wf 
+       inheritance overriding
+  show ?thesis
+    by (auto dest: non_Package_instance_method_inheritance)
+qed
+
+lemma dynamic_to_static_overriding:
+(assumes dyn_override: "G\<turnstile> new overrides old" and
+          accmodi_old: "accmodi old \<noteq> Package" and
+                   wf: "wf_prog G"
+) "G\<turnstile> new overrides\<^sub>S old"  
+proof - 
+  from dyn_override accmodi_old
+  show ?thesis (is "?Overrides new old")
+  proof (induct rule: overridesR.induct)
+    case (Direct new old)
+    assume   new_declared: "G\<turnstile>Method new declared_in declclass new"
+    assume eq_sig_new_old: "msig new = msig old"
+    assume subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old"
+    assume "G \<turnstile>Method old inheritable_in pid (declclass new)" and
+           "accmodi old \<noteq> Package" and
+           "\<not> is_static old" and
+           "G\<turnstile>declclass new\<prec>\<^sub>C declclass old" and
+           "G\<turnstile>Method old declared_in declclass old" 
+    from this wf
+    show "?Overrides new old"
+    proof (cases rule: non_Package_instance_method_inheritance_cases)
+      case Inheritance
+      assume "G \<turnstile>Method old member_of declclass new"
+      then have "G\<turnstile>mid (msig old) undeclared_in declclass new"
+      proof cases
+	case Immediate 
+	with subcls_new_old wf show ?thesis 	
+	  by (auto dest: subcls_irrefl)
+      next
+	case Inherited
+	then show ?thesis
+	  by (cases old) auto
+      qed
+      with eq_sig_new_old new_declared
+      show ?thesis
+	by (cases old,cases new) (auto dest!: declared_not_undeclared)
+    next
+      case (Overriding new') 
+      assume stat_override_new': "G \<turnstile> new' overrides\<^sub>S old"
+      then have "msig new' = msig old"
+	by (auto dest: stat_overrides_commonD)
+      with eq_sig_new_old have eq_sig_new_new': "msig new=msig new'"
+	by simp
+      assume "G \<turnstile>Method new' member_of declclass new"
+      then show ?thesis
+      proof (cases)
+	case Immediate
+	then have declC_new: "declclass new' = declclass new" 
+	  by auto
+	from Immediate 
+	have "G\<turnstile>Method new' declared_in declclass new"
+	  by (cases new') auto
+	with new_declared eq_sig_new_new' declC_new 
+	have "new=new'"
+	  by (cases new, cases new') (auto dest: unique_declared_in) 
+	with stat_override_new'
+	show ?thesis
+	  by simp
+      next
+	case Inherited
+	then have "G\<turnstile>mid (msig new') undeclared_in declclass new"
+	  by (cases new') (auto)
+	with eq_sig_new_new' new_declared
+	show ?thesis
+	  by (cases new,cases new') (auto dest!: declared_not_undeclared)
+      qed
+    qed
+  next
+    case (Indirect inter new old)
+    assume accmodi_old: "accmodi old \<noteq> Package"
+    assume "accmodi old \<noteq> Package \<Longrightarrow> G \<turnstile> inter overrides\<^sub>S old"
+    with accmodi_old 
+    have stat_override_inter_old: "G \<turnstile> inter overrides\<^sub>S old"
+      by blast
+    moreover 
+    assume hyp_inter: "accmodi inter \<noteq> Package \<Longrightarrow> G \<turnstile> new overrides\<^sub>S inter"
+    moreover
+    have "accmodi inter \<noteq> Package"
+    proof -
+      from stat_override_inter_old wf 
+      have "accmodi old \<le> accmodi inter"
+	by (auto dest: wf_prog_stat_overridesD)
+      with stat_override_inter_old accmodi_old
+      show ?thesis
+	by (auto dest!: no_Private_stat_override
+                 split: acc_modi.splits 
+	         dest: acc_modi_le_Dests)
+    qed
+    ultimately show "?Overrides new old"
+      by (blast intro: stat_overridesR.Indirect)
+  qed
+qed
+
+lemma wf_prog_dyn_override_prop:
+ (assumes dyn_override: "G \<turnstile> new overrides old" and
+                    wf: "wf_prog G"
+ ) "accmodi old \<le> accmodi new"
+proof (cases "accmodi old = Package")
+  case True
+  note old_Package = this
+  show ?thesis
+  proof (cases "accmodi old \<le> accmodi new")
+    case True then show ?thesis .
+  next
+    case False
+    with old_Package 
+    have "accmodi new = Private"
+      by (cases "accmodi new") (auto simp add: le_acc_def less_acc_def)
+    with dyn_override 
+    show ?thesis
+      by (auto dest: overrides_commonD)
+  qed    
+next
+  case False
+  with dyn_override wf
+  have "G \<turnstile> new overrides\<^sub>S old"
+    by (blast intro: dynamic_to_static_overriding)
+  with wf 
+  show ?thesis
+   by (blast dest: wf_prog_stat_overridesD)
+qed 
+
+lemma overrides_Package_old: 
+(assumes dyn_override: "G \<turnstile> new overrides old" and 
+          accmodi_new: "accmodi new = Package" and
+                   wf: "wf_prog G "
+) "accmodi old = Package"
+proof (cases "accmodi old")
+  case Private
+  with dyn_override show ?thesis
+    by (simp add: no_Private_override)
+next
+  case Package
+  then show ?thesis .
+next
+  case Protected
+  with dyn_override wf
+  have "G \<turnstile> new overrides\<^sub>S old"
+    by (auto intro: dynamic_to_static_overriding)
+  with wf 
+  have "accmodi old \<le> accmodi new"
+    by (auto dest: wf_prog_stat_overridesD)
+  with Protected accmodi_new
+  show ?thesis
+    by (simp add: less_acc_def le_acc_def)
+next
+  case Public
+  with dyn_override wf
+  have "G \<turnstile> new overrides\<^sub>S old"
+    by (auto intro: dynamic_to_static_overriding)
+  with wf 
+  have "accmodi old \<le> accmodi new"
+    by (auto dest: wf_prog_stat_overridesD)
+  with Public accmodi_new
+  show ?thesis
+    by (simp add: less_acc_def le_acc_def)
+qed
+
+lemma dyn_override_Package:
+ (assumes dyn_override: "G \<turnstile> new overrides old" and
+          accmodi_old: "accmodi old = Package" and 
+          accmodi_new: "accmodi new = Package" and
+                    wf: "wf_prog G"
+ ) "pid (declclass old) = pid (declclass new)"
+proof - 
+  from dyn_override accmodi_old accmodi_new
+  show ?thesis (is "?EqPid old new")
+  proof (induct rule: overridesR.induct)
+    case (Direct new old)
+    assume "accmodi old = Package"
+           "G \<turnstile>Method old inheritable_in pid (declclass new)"
+    then show "pid (declclass old) =  pid (declclass new)"
+      by (auto simp add: inheritable_in_def)
+  next
+    case (Indirect inter new old)
+    assume accmodi_old: "accmodi old = Package" and
+           accmodi_new: "accmodi new = Package" 
+    assume "G \<turnstile> new overrides inter"
+    with accmodi_new wf
+    have "accmodi inter = Package"
+      by  (auto intro: overrides_Package_old)
+    with Indirect
+    show "pid (declclass old) =  pid (declclass new)"
+      by auto
+  qed
+qed
+
+lemma dyn_override_Package_escape:
+ (assumes dyn_override: "G \<turnstile> new overrides old" and
+          accmodi_old: "accmodi old = Package" and 
+         outside_pack: "pid (declclass old) \<noteq> pid (declclass new)" and
+                    wf: "wf_prog G"
+ ) "\<exists> inter. G \<turnstile> new overrides inter \<and> G \<turnstile> inter overrides old \<and>
+             pid (declclass old) = pid (declclass inter) \<and>
+             Protected \<le> accmodi inter"
+proof -
+  from dyn_override accmodi_old outside_pack
+  show ?thesis (is "?P new old")
+  proof (induct rule: overridesR.induct)
+    case (Direct new old)
+    assume accmodi_old: "accmodi old = Package"
+    assume outside_pack: "pid (declclass old) \<noteq> pid (declclass new)"
+    assume "G \<turnstile>Method old inheritable_in pid (declclass new)"
+    with accmodi_old 
+    have "pid (declclass old) = pid (declclass new)"
+      by (simp add: inheritable_in_def)
+    with outside_pack 
+    show "?P new old"
+      by (contradiction)
+  next
+    case (Indirect inter new old)
+    assume accmodi_old: "accmodi old = Package"
+    assume outside_pack: "pid (declclass old) \<noteq> pid (declclass new)"
+    assume override_new_inter: "G \<turnstile> new overrides inter"
+    assume override_inter_old: "G \<turnstile> inter overrides old"
+    assume hyp_new_inter: "\<lbrakk>accmodi inter = Package; 
+                           pid (declclass inter) \<noteq> pid (declclass new)\<rbrakk>
+                           \<Longrightarrow> ?P new inter"
+    assume hyp_inter_old: "\<lbrakk>accmodi old = Package; 
+                           pid (declclass old) \<noteq> pid (declclass inter)\<rbrakk>
+                           \<Longrightarrow> ?P inter old"
+    show "?P new old"
+    proof (cases "pid (declclass old) = pid (declclass inter)")
+      case True
+      note same_pack_old_inter = this
+      show ?thesis
+      proof (cases "pid (declclass inter) = pid (declclass new)")
+	case True
+	with same_pack_old_inter outside_pack
+	show ?thesis
+	  by auto
+      next
+	case False
+	note diff_pack_inter_new = this
+	show ?thesis
+	proof (cases "accmodi inter = Package")
+	  case True
+	  with diff_pack_inter_new hyp_new_inter  
+	  obtain newinter where
+	    over_new_newinter: "G \<turnstile> new overrides newinter" and
+            over_newinter_inter: "G \<turnstile> newinter overrides inter" and 
+            eq_pid: "pid (declclass inter) = pid (declclass newinter)" and
+            accmodi_newinter: "Protected \<le> accmodi newinter"
+	    by auto
+	  from over_newinter_inter override_inter_old
+	  have "G\<turnstile>newinter overrides old"
+	    by (rule overridesR.Indirect)
+	  moreover
+	  from eq_pid same_pack_old_inter 
+	  have "pid (declclass old) = pid (declclass newinter)"
+	    by simp
+	  moreover
+	  note over_new_newinter accmodi_newinter
+	  ultimately show ?thesis
+	    by blast
+	next
+	  case False
+	  with override_new_inter
+	  have "Protected \<le> accmodi inter"
+	    by (cases "accmodi inter") (auto dest: no_Private_override)
+	  with override_new_inter override_inter_old same_pack_old_inter
+	  show ?thesis
+	    by blast
+	qed
+      qed
+    next
+      case False
+      with accmodi_old hyp_inter_old
+      obtain newinter where
+	over_inter_newinter: "G \<turnstile> inter overrides newinter" and
+          over_newinter_old: "G \<turnstile> newinter overrides old" and 
+                eq_pid: "pid (declclass old) = pid (declclass newinter)" and
+	accmodi_newinter: "Protected \<le> accmodi newinter"
+	by auto
+      from override_new_inter over_inter_newinter 
+      have "G \<turnstile> new overrides newinter"
+	by (rule overridesR.Indirect)
+      with eq_pid over_newinter_old accmodi_newinter
+      show ?thesis
+	by blast
+    qed
+  qed
+qed
+
+declare split_paired_All [simp del] split_paired_Ex [simp del]
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac";
+claset_ref () := claset () delSWrapper "split_all_tac"
+*}
+
+lemma declclass_widen[rule_format]: 
+ "wf_prog G 
+ \<longrightarrow> (\<forall>c m. class G C = Some c \<longrightarrow> methd G C sig = Some m 
+ \<longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m)" (is "?P G C")
+proof (rule class_rec.induct,intro allI impI)
+  fix G C c m
+  assume Hyp: "\<forall>c. C \<noteq> Object \<and> ws_prog G \<and> class G C = Some c 
+               \<longrightarrow> ?P G (super c)"
+  assume wf: "wf_prog G" and cls_C: "class G C = Some c" and
+         m:  "methd G C sig = Some m"
+  show "G\<turnstile>C\<preceq>\<^sub>C declclass m" 
+  proof (cases "C=Object")
+    case True 
+    with wf m show ?thesis by (simp add: methd_Object_SomeD)
+  next
+    let ?filter="filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)"
+    let ?table = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
+    case False 
+    with cls_C wf m
+    have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m "
+      by (simp add: methd_rec)
+    show ?thesis
+    proof (cases "?table sig")
+      case None
+      from this methd_C have "?filter (methd G (super c)) sig = Some m"
+	by simp
+      moreover
+      from wf cls_C False obtain sup where "class G (super c) = Some sup"
+	by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
+      moreover note wf False cls_C Hyp 
+      ultimately have "G\<turnstile>super c \<preceq>\<^sub>C declclass m"  by auto
+      moreover from cls_C False have  "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c" by (rule subcls1I)
+      ultimately show ?thesis by - (rule rtrancl_into_rtrancl2)
+    next
+      case Some
+      from this wf False cls_C methd_C show ?thesis by auto
+    qed
+  qed
+qed
+
+(*
+lemma declclass_widen[rule_format]: 
+ "wf_prog G 
+ \<longrightarrow> (\<forall>c m. class G C = Some c \<longrightarrow> methd G C sig = Some m 
+ \<longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m)" (is "?P G C")
+apply (rule class_rec.induct)
+apply (rule impI)+
+apply (case_tac "C=Object")
+apply   (force simp add: methd_Object_SomeD)
+
+apply   (rule allI)+
+apply   (rule impI)
+apply   (simp (no_asm_simp) add: methd_rec)
+apply   (case_tac "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig")
+apply    (simp add: override_def)
+apply    (frule (1) subcls1I)
+apply    (drule (1) wf_prog_cdecl)
+apply    (drule (1) wf_cdecl_supD)
+apply    clarify
+apply    (drule is_acc_class_is_class)
+apply    clarify
+apply    (blast dest: rtrancl_into_rtrancl2)
+
+apply    auto
+done
+*)
+
+(*
+lemma accessible_public_inheritance_lemma1:
+  "\<lbrakk>wf_prog G; class G C = Some c; C \<noteq> Object; accmodi m = Public;
+    G\<turnstile>m accessible_through_inheritance_in (super c)\<rbrakk> 
+   \<Longrightarrow> G\<turnstile>m accessible_through_inheritance_in C"
+apply   (frule (1) subcls1I)
+apply   (rule accessible_through_inheritance.Indirect)
+apply     (blast)
+apply     (erule accessible_through_inheritance_subclsD)
+apply     (blast dest: wf_prog_acc_superD is_acc_classD)
+apply     assumption
+apply     (force dest: wf_prog_acc_superD is_acc_classD
+                 simp add: accessible_for_inheritance_in_def)
+done
+
+lemma accessible_public_inheritance_lemma[rule_format]:
+  "\<lbrakk>wf_prog G;C \<noteq> Object; class G C = Some c; 
+    accmodi m = Public
+   \<rbrakk> \<Longrightarrow> methd G (super c) sig = Some m 
+        \<longrightarrow> G\<turnstile>m accessible_through_inheritance_in C" 
+apply (frule (2) wf_prog_acc_superD [THEN is_acc_classD])
+apply (erule conjE)
+apply (simp only: not_None_eq)
+apply (erule exE)
+apply (case_tac "(super c) = Object")
+apply   (rule impI)
+apply   (rule accessible_through_inheritance.Direct)
+apply     force
+apply     (force simp add: accessible_for_inheritance_in_def)
+
+apply   (frule wf_ws_prog) 
+apply   (simp add: methd_rec)
+apply   (case_tac "table_of (map (\<lambda>(s, m). (s, super c, m)) (methods y)) sig")
+apply     simp
+apply     (clarify)
+apply     (rule_tac D="super c" in accessible_through_inheritance.Indirect)
+apply       (blast dest: subcls1I)
+apply       (blast)
+apply       simp
+apply       assumption
+apply       (simp add: accessible_for_inheritance_in_def)
+
+apply     clarsimp
+apply     (rule accessible_through_inheritance.Direct)
+apply     (auto dest: subcls1I simp add: accessible_for_inheritance_in_def)
+done
+
+lemma accessible_public_inheritance:
+  "\<lbrakk>wf_prog G; class G D = Some d; G\<turnstile>C \<prec>\<^sub>C D; methd G D sig = Some m; 
+    accmodi m = Public\<rbrakk> 
+   \<Longrightarrow> G \<turnstile> m accessible_through_inheritance_in C"
+apply (erule converse_trancl_induct)
+apply  (blast dest: subcls1D intro: accessible_public_inheritance_lemma)
+
+apply  (frule subcls1D)
+apply  clarify
+apply  (frule  (2) wf_prog_acc_superD [THEN is_acc_classD])
+apply  clarify
+apply  (rule_tac D="super c" in accessible_through_inheritance.Indirect)
+apply   (auto intro:trancl_into_trancl2 
+                    accessible_through_inheritance_subclsD
+              simp add: accessible_for_inheritance_in_def)
+done
+*)
+
+
+lemma declclass_methd_Object: 
+ "\<lbrakk>wf_prog G; methd G Object sig = Some m\<rbrakk> \<Longrightarrow> declclass m = Object"
+by auto
+
+
+lemma methd_declaredD: 
+ "\<lbrakk>wf_prog G; is_class G C;methd G C sig = Some m\<rbrakk> 
+  \<Longrightarrow> G\<turnstile>(mdecl (sig,mthd m)) declared_in (declclass m)"
+proof -
+  assume    wf: "wf_prog G"
+  then have ws: "ws_prog G" ..
+  assume  clsC: "is_class G C"
+  from clsC ws 
+  show "methd G C sig = Some m 
+        \<Longrightarrow> G\<turnstile>(mdecl (sig,mthd m)) declared_in (declclass m)"
+    (is "PROP ?P C") 
+  proof (induct ?P C rule: ws_class_induct')
+    case Object
+    assume "methd G Object sig = Some m" 
+    with wf show ?thesis
+      by - (rule method_declared_inI, auto) 
+  next
+    case Subcls
+    fix C c
+    assume clsC: "class G C = Some c"
+    and       m: "methd G C sig = Some m"
+    and     hyp: "methd G (super c) sig = Some m \<Longrightarrow> ?thesis" 
+    let ?newMethods = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
+    show ?thesis
+    proof (cases "?newMethods sig")
+      case None
+      from None ws clsC m hyp 
+      show ?thesis by (auto intro: method_declared_inI simp add: methd_rec)
+    next
+      case Some
+      from Some ws clsC m 
+      show ?thesis by (auto intro: method_declared_inI simp add: methd_rec) 
+    qed
+  qed
+qed
+
+
+lemma methd_rec_Some_cases [consumes 4, case_names NewMethod InheritedMethod]:
+(assumes methd_C: "methd G C sig = Some m" and
+                    ws: "ws_prog G" and
+                  clsC: "class G C = Some c" and
+             neq_C_Obj: "C\<noteq>Object" )  
+"\<lbrakk>table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = Some m \<Longrightarrow> P;
+  \<lbrakk>G\<turnstile>C inherits (method sig m); methd G (super c) sig = Some m\<rbrakk> \<Longrightarrow> P 
+ \<rbrakk> \<Longrightarrow> P"
+proof -
+  let ?inherited   = "filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m) 
+                              (methd G (super c))"
+  let ?new = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
+  from ws clsC neq_C_Obj methd_C 
+  have methd_unfold: "(?inherited ++ ?new) sig = Some m"
+    by (simp add: methd_rec)
+  assume NewMethod: "?new sig = Some m \<Longrightarrow> P"
+  assume InheritedMethod: "\<lbrakk>G\<turnstile>C inherits (method sig m); 
+                            methd G (super c) sig = Some m\<rbrakk> \<Longrightarrow> P"
+  show P
+  proof (cases "?new sig")
+    case None
+    with methd_unfold have "?inherited sig = Some m"
+      by (auto)
+    with InheritedMethod show P by blast
+  next
+    case Some
+    with methd_unfold have "?new sig = Some m"
+      by auto
+    with NewMethod show P by blast
+  qed
+qed
+
+  
+lemma methd_member_of:
+ (assumes wf: "wf_prog G") 
+  "\<lbrakk>is_class G C; methd G C sig = Some m\<rbrakk> \<Longrightarrow> G\<turnstile>Methd sig m member_of C" 
+  (is "?Class C \<Longrightarrow> ?Method C \<Longrightarrow> ?MemberOf C") 
+proof -
+  from wf   have   ws: "ws_prog G" ..
+  assume defC: "is_class G C"
+  from defC ws 
+  show "?Class C \<Longrightarrow> ?Method C \<Longrightarrow> ?MemberOf C"
+  proof (induct rule: ws_class_induct')  
+    case Object
+    with wf have declC: "declclass m = Object"
+      by (blast intro: declclass_methd_Object)
+    with Object wf have "G\<turnstile>Methd sig m declared_in Object"
+      by (auto dest: methd_declaredD simp del: methd_Object)
+    with declC 
+    show "?MemberOf Object"
+      by (auto intro!: members.Immediate
+                  simp del: methd_Object)
+  next
+    case (Subcls C c)
+    assume  clsC: "class G C = Some c" and
+       neq_C_Obj: "C \<noteq> Object"  
+    assume methd: "?Method C"
+    from methd ws clsC neq_C_Obj
+    show "?MemberOf C"
+    proof (cases rule: methd_rec_Some_cases)
+      case NewMethod
+      with clsC show ?thesis
+	by (auto dest: method_declared_inI intro!: members.Immediate)
+    next
+      case InheritedMethod
+      then show "?thesis"
+	by (blast dest: inherits_member)
+    qed
+  qed
+qed
+
+lemma current_methd: 
+      "\<lbrakk>table_of (methods c) sig = Some new;
+        ws_prog G; class G C = Some c; C \<noteq> Object; 
+        methd G (super c) sig = Some old\<rbrakk> 
+    \<Longrightarrow> methd G C sig = Some (C,new)"
+by (auto simp add: methd_rec
+            intro: filter_tab_SomeI override_find_right table_of_map_SomeI)
+
+lemma wf_prog_staticD:
+ (assumes     wf: "wf_prog G" and 
+            clsC: "class G C = Some c" and
+       neq_C_Obj: "C \<noteq> Object" and 
+             old: "methd G (super c) sig = Some old" and 
+     accmodi_old: "Protected \<le> accmodi old" and
+             new: "table_of (methods c) sig = Some new"
+ ) "is_static new = is_static old"
+proof -
+  from clsC wf 
+  have wf_cdecl: "wf_cdecl G (C,c)" by (rule wf_prog_cdecl)
+  from wf clsC neq_C_Obj
+  have is_cls_super: "is_class G (super c)" 
+    by (blast dest: wf_prog_acc_superD is_acc_classD)
+  from wf is_cls_super old 
+  have old_member_of: "G\<turnstile>Methd sig old member_of (super c)"  
+    by (rule methd_member_of)
+  from old wf is_cls_super 
+  have old_declared: "G\<turnstile>Methd sig old declared_in (declclass old)"
+    by (auto dest: methd_declared_in_declclass)
+  from new clsC 
+  have new_declared: "G\<turnstile>Methd sig (C,new) declared_in C"
+    by (auto intro: method_declared_inI)
+  note trancl_rtrancl_tranc = trancl_rtrancl_trancl [trans] (* ### in Basis *)
+  from clsC neq_C_Obj
+  have subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
+    by (rule subcls1I)
+  then have "G\<turnstile>C \<prec>\<^sub>C super c" ..
+  also from old wf is_cls_super
+  have "G\<turnstile>super c \<preceq>\<^sub>C (declclass old)" by (auto dest: methd_declC)
+  finally have subcls_C_old:  "G\<turnstile>C \<prec>\<^sub>C (declclass old)" .
+  from accmodi_old 
+  have inheritable: "G\<turnstile>Methd sig old inheritable_in pid C"
+    by (auto simp add: inheritable_in_def
+                 dest: acc_modi_le_Dests)
+  show ?thesis
+  proof (cases "is_static new")
+    case True
+    with subcls_C_old new_declared old_declared inheritable
+    have "G,sig\<turnstile>(C,new) hides old"
+      by (auto intro: hidesI)
+    with True wf_cdecl neq_C_Obj new 
+    show ?thesis
+      by (auto dest: wf_cdecl_hides_SomeD)
+  next
+    case False
+    with subcls_C_old new_declared old_declared inheritable subcls1_C_super
+         old_member_of
+    have "G,sig\<turnstile>(C,new) overrides\<^sub>S old"
+      by (auto intro: stat_overridesR.Direct)
+    with False wf_cdecl neq_C_Obj new 
+    show ?thesis
+      by (auto dest: wf_cdecl_overrides_SomeD)
+  qed
+qed
+
+lemma inheritable_instance_methd: 
+ (assumes subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" and
+              is_cls_D: "is_class G D" and
+                    wf: "wf_prog G" and 
+                   old: "methd G D sig = Some old" and
+           accmodi_old: "Protected \<le> accmodi old" and  
+        not_static_old: "\<not> is_static old"
+ )  
+  "\<exists>new. methd G C sig = Some new \<and>
+         (new = old \<or> G,sig\<turnstile>new overrides\<^sub>S old)"
+ (is "(\<exists>new. (?Constraint C new old))")
+proof -
+  from subclseq_C_D is_cls_D 
+  have is_cls_C: "is_class G C" by (rule subcls_is_class2) 
+  from wf 
+  have ws: "ws_prog G" ..
+  from is_cls_C ws subclseq_C_D 
+  show "\<exists>new. ?Constraint C new old"
+  proof (induct rule: ws_class_induct')
+    case (Object co)
+    then have eq_D_Obj: "D=Object" by auto
+    with old 
+    have "?Constraint Object old old"
+      by auto
+    with eq_D_Obj 
+    show "\<exists> new. ?Constraint Object new old" by auto
+  next
+    case (Subcls C c)
+    assume hyp: "G\<turnstile>super c\<preceq>\<^sub>C D \<Longrightarrow> \<exists>new. ?Constraint (super c) new old"
+    assume clsC: "class G C = Some c"
+    assume neq_C_Obj: "C\<noteq>Object"
+    from clsC wf 
+    have wf_cdecl: "wf_cdecl G (C,c)" 
+      by (rule wf_prog_cdecl)
+    from ws clsC neq_C_Obj
+    have is_cls_super: "is_class G (super c)"
+      by (auto dest: ws_prog_cdeclD)
+    from clsC wf neq_C_Obj 
+    have superAccessible: "G\<turnstile>(Class (super c)) accessible_in (pid C)" and
+	 subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
+      by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD
+              intro: subcls1I)
+    show "\<exists>new. ?Constraint C new old"
+    proof (cases "G\<turnstile>super c\<preceq>\<^sub>C D")
+      case False
+      from False Subcls 
+      have eq_C_D: "C=D"
+	by (auto dest: subclseq_superD)
+      with old 
+      have "?Constraint C old old"
+	by auto
+      with eq_C_D 
+      show "\<exists> new. ?Constraint C new old" by auto
+    next
+      case True
+      with hyp obtain super_method
+	where super: "?Constraint (super c) super_method old" by blast
+      from super not_static_old
+      have not_static_super: "\<not> is_static super_method"
+	by (auto dest!: stat_overrides_commonD)
+      from super old wf accmodi_old
+      have accmodi_super_method: "Protected \<le> accmodi super_method"
+	by (auto dest!: wf_prog_stat_overridesD
+                 intro: order_trans)
+      from super accmodi_old wf
+      have inheritable: "G\<turnstile>Methd sig super_method inheritable_in (pid C)"
+	by (auto dest!: wf_prog_stat_overridesD
+                        acc_modi_le_Dests
+              simp add: inheritable_in_def)	           
+      from super wf is_cls_super
+      have member: "G\<turnstile>Methd sig super_method member_of (super c)"
+	by (auto intro: methd_member_of) 
+      from member
+      have decl_super_method:
+        "G\<turnstile>Methd sig super_method declared_in (declclass super_method)"
+	by (auto dest: member_of_declC)
+      from super subcls1_C_super ws is_cls_super 
+      have subcls_C_super: "G\<turnstile>C \<prec>\<^sub>C (declclass super_method)"
+	by (auto intro: rtrancl_into_trancl2 dest: methd_declC) 
+      show "\<exists> new. ?Constraint C new old"
+      proof (cases "methd G C sig")
+	case None
+	have "methd G (super c) sig = None"
+	proof -
+	  from clsC ws None 
+	  have no_new: "table_of (methods c) sig = None" 
+	    by (auto simp add: methd_rec)
+	  with clsC 
+	  have undeclared: "G\<turnstile>mid sig undeclared_in C"
+	    by (auto simp add: undeclared_in_def cdeclaredmethd_def)
+	  with inheritable member superAccessible subcls1_C_super
+	  have inherits: "G\<turnstile>C inherits (method sig super_method)"
+	    by (auto simp add: inherits_def)
+	  with clsC ws no_new super neq_C_Obj
+	  have "methd G C sig = Some super_method"
+	    by (auto simp add: methd_rec override_def
+	                intro: filter_tab_SomeI)
+          with None show ?thesis
+	    by simp
+	qed
+	with super show ?thesis by auto
+      next
+	case (Some new)
+	from this ws clsC neq_C_Obj
+	show ?thesis
+	proof (cases rule: methd_rec_Some_cases)
+	  case InheritedMethod
+	  with super Some show ?thesis 
+	    by auto
+	next
+	  case NewMethod
+	  assume new: "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig 
+                       = Some new"
+	  from new 
+	  have declcls_new: "declclass new = C" 
+	    by auto
+	  from wf clsC neq_C_Obj super new not_static_super accmodi_super_method
+	  have not_static_new: "\<not> is_static new" 
+	    by (auto dest: wf_prog_staticD) 
+	  from clsC new
+	  have decl_new: "G\<turnstile>Methd sig new declared_in C"
+	    by (auto simp add: declared_in_def cdeclaredmethd_def)
+	  from not_static_new decl_new decl_super_method
+	       member subcls1_C_super inheritable declcls_new subcls_C_super 
+	  have "G,sig\<turnstile> new overrides\<^sub>S super_method"
+	    by (auto intro: stat_overridesR.Direct) 
+	  with super Some
+	  show ?thesis
+	    by (auto intro: stat_overridesR.Indirect)
+	qed
+      qed
+    qed
+  qed
+qed
+
+lemma inheritable_instance_methd_cases [consumes 6
+                                       , case_names Inheritance Overriding]: 
+ (assumes subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" and
+              is_cls_D: "is_class G D" and
+                    wf: "wf_prog G" and 
+                   old: "methd G D sig = Some old" and
+           accmodi_old: "Protected \<le> accmodi old" and  
+        not_static_old: "\<not> is_static old" and
+           inheritance:  "methd G C sig = Some old \<Longrightarrow> P" and
+            overriding:  "\<And> new. \<lbrakk>methd G C sig = Some new;
+                                   G,sig\<turnstile>new overrides\<^sub>S old\<rbrakk> \<Longrightarrow> P"
+        
+ ) "P"
+proof -
+from subclseq_C_D is_cls_D wf old accmodi_old not_static_old 
+     inheritance overriding
+show ?thesis
+  by (auto dest: inheritable_instance_methd)
+qed
+
+lemma inheritable_instance_methd_props: 
+ (assumes subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" and
+              is_cls_D: "is_class G D" and
+                    wf: "wf_prog G" and 
+                   old: "methd G D sig = Some old" and
+           accmodi_old: "Protected \<le> accmodi old" and  
+        not_static_old: "\<not> is_static old"
+ )  
+  "\<exists>new. methd G C sig = Some new \<and>
+          \<not> is_static new \<and> G\<turnstile>resTy new\<preceq>resTy old \<and> accmodi old \<le>accmodi new"
+ (is "(\<exists>new. (?Constraint C new old))")
+proof -
+  from subclseq_C_D is_cls_D wf old accmodi_old not_static_old 
+  show ?thesis
+  proof (cases rule: inheritable_instance_methd_cases)
+    case Inheritance
+    with not_static_old accmodi_old show ?thesis by auto
+  next
+    case (Overriding new)
+    then have "\<not> is_static new" by (auto dest: stat_overrides_commonD)
+    with Overriding not_static_old accmodi_old wf 
+    show ?thesis 
+      by (auto dest!: wf_prog_stat_overridesD
+               intro: order_trans)
+  qed
+qed
+ 	  
+(* ### Probleme: Die tollen methd_subcls_cases Lemma wird warscheinlich
+  kaum gebraucht: 
+Redundanz: stat_overrides.Direct old declared in declclass old folgt aus
+           member of 
+   Problem: Predikate wie overrides, sind global üper die Hierarchie hinweg
+            definiert, aber oft barucht man eben zusätlich Induktion
+            : von super c auf C; Dann ist aber auss dem Kontext
+            die Unterscheidung in die 5 fälle overkill,
+            da man dann warscheinlich meistens eh in einem speziellen
+            Fall kommt (durch die Hypothesen)
+*)
+    
+(* local lemma *)
+ML {* bind_thm("bexI'",permute_prems 0 1 bexI) *}
+ML {* bind_thm("ballE'",permute_prems 1 1 ballE) *}
+lemma subint_widen_imethds: 
+ "\<lbrakk>G\<turnstile>I\<preceq>I J; wf_prog G; is_iface G J; jm \<in> imethds G J sig\<rbrakk> \<Longrightarrow>   
+  \<exists> im \<in> imethds G I sig. is_static im = is_static jm \<and> 
+                          accmodi im = accmodi jm \<and>
+                          G\<turnstile>resTy im\<preceq>resTy jm"
+proof -
+  assume irel: "G\<turnstile>I\<preceq>I J" and
+           wf: "wf_prog G" and
+     is_iface: "is_iface G J"
+  from irel show "jm \<in> imethds G J sig \<Longrightarrow> ?thesis" 
+               (is "PROP ?P I" is "PROP ?Prem J \<Longrightarrow> ?Concl I")
+  proof (induct ?P I rule: converse_rtrancl_induct) 
+    case Id
+    assume "jm \<in> imethds G J sig"
+    then show "?Concl J" by  (blast elim: bexI')
+  next
+    case Step
+    fix I SI
+    assume subint1_I_SI: "G\<turnstile>I \<prec>I1 SI" and 
+            subint_SI_J: "G\<turnstile>SI \<preceq>I J" and
+                    hyp: "PROP ?P SI" and
+                     jm: "jm \<in> imethds G J sig"
+    from subint1_I_SI 
+    obtain i where
+      ifI: "iface G I = Some i" and
+       SI: "SI \<in> set (isuperIfs i)"
+      by (blast dest: subint1D)
+
+    let ?newMethods 
+          = "(o2s \<circ> table_of (map (\<lambda>(sig, mh). (sig, I, mh)) (imethods i)))"
+    show "?Concl I"
+    proof (cases "?newMethods sig = {}")
+      case True
+      with ifI SI hyp wf jm 
+      show "?thesis" 
+	by (auto simp add: imethds_rec) 
+    next
+      case False
+      from ifI wf False
+      have imethds: "imethds G I sig = ?newMethods sig"
+	by (simp add: imethds_rec)
+      from False
+      obtain im where
+        imdef: "im \<in> ?newMethods sig" 
+	by (blast)
+      with imethds 
+      have im: "im \<in> imethds G I sig"
+	by (blast)
+      with im wf ifI 
+      obtain
+	 imStatic: "\<not> is_static im" and
+         imPublic: "accmodi im = Public"
+	by (auto dest!: imethds_wf_mhead)	
+      from ifI wf 
+      have wf_I: "wf_idecl G (I,i)" 
+	by (rule wf_prog_idecl)
+      with SI wf  
+      obtain si where
+	 ifSI: "iface G SI = Some si" and
+	wf_SI: "wf_idecl G (SI,si)" 
+	by (auto dest!: wf_idecl_supD is_acc_ifaceD
+                  dest: wf_prog_idecl)
+      from jm hyp 
+      obtain sim::"qtname \<times> mhead"  where
+                      sim: "sim \<in> imethds G SI sig" and
+         eq_static_sim_jm: "is_static sim = is_static jm" and 
+         eq_access_sim_jm: "accmodi sim = accmodi jm" and 
+        resTy_widen_sim_jm: "G\<turnstile>resTy sim\<preceq>resTy jm"
+	by blast
+      with wf_I SI imdef sim 
+      have "G\<turnstile>resTy im\<preceq>resTy sim"   
+	by (auto dest!: wf_idecl_hidings hidings_entailsD)
+      with wf resTy_widen_sim_jm 
+      have resTy_widen_im_jm: "G\<turnstile>resTy im\<preceq>resTy jm"
+	by (blast intro: widen_trans)
+      from sim wf ifSI  
+      obtain
+	simStatic: "\<not> is_static sim" and
+        simPublic: "accmodi sim = Public"
+	by (auto dest!: imethds_wf_mhead)
+      from im 
+           imStatic simStatic eq_static_sim_jm
+           imPublic simPublic eq_access_sim_jm
+           resTy_widen_im_jm
+      show ?thesis 
+	by auto 
+    qed
+  qed
+qed
+     
+(* Tactical version *)
+(* 
+lemma subint_widen_imethds: "\<lbrakk>G\<turnstile>I\<preceq>I J; wf_prog G; is_iface G J\<rbrakk> \<Longrightarrow>  
+  \<forall> jm \<in> imethds G J sig.  
+  \<exists> im \<in> imethds G I sig. static (mthd im)=static (mthd jm) \<and> 
+                          access (mthd im)= access (mthd jm) \<and>
+                          G\<turnstile>resTy (mthd im)\<preceq>resTy (mthd jm)"
+apply (erule converse_rtrancl_induct)
+apply  (clarsimp elim!: bexI')
+apply (frule subint1D)
+apply clarify
+apply (erule ballE')
+apply  fast
+apply (erule_tac V = "?x \<in> imethds G J sig" in thin_rl)
+apply clarsimp
+apply (subst imethds_rec, assumption, erule wf_ws_prog)
+apply (unfold overrides_t_def)
+apply (drule (1) wf_prog_idecl)
+apply (frule (3) imethds_wf_mhead [OF _ _ wf_idecl_supD [THEN conjunct1 
+                                       [THEN is_acc_ifaceD [THEN conjunct1]]]])
+apply (case_tac "(o2s \<circ> table_of (map (\<lambda>(s, mh). (s, y, mh)) (imethods i)))
+                  sig ={}")
+apply   force
+
+apply   (simp only:)
+apply   (simp)
+apply   clarify
+apply   (frule wf_idecl_hidings [THEN hidings_entailsD])
+apply     blast
+apply     blast
+apply   (rule bexI')
+apply     simp
+apply     (drule table_of_map_SomeI [of _ "sig"])
+apply     simp
+
+apply     (frule wf_idecl_mhead [of _ _ _ "sig"])
+apply       (rule table_of_Some_in_set)
+apply       assumption
+apply     auto
+done
+*)
+    
+
+(* local lemma *)
+lemma implmt1_methd: 
+ "\<And>sig. \<lbrakk>G\<turnstile>C\<leadsto>1I; wf_prog G; im \<in> imethds G I sig\<rbrakk> \<Longrightarrow>  
+  \<exists>cm \<in>methd G C sig: \<not> is_static cm \<and> \<not> is_static im \<and> 
+                       G\<turnstile>resTy cm\<preceq>resTy im \<and>
+                       accmodi im = Public \<and> accmodi cm = Public"
+apply (drule implmt1D)
+apply clarify
+apply (drule (2) wf_prog_cdecl [THEN wf_cdecl_impD])
+apply (frule (1) imethds_wf_mhead)
+apply  (simp add: is_acc_iface_def)
+apply (force)
+done
+
+
+(* local lemma *)
+lemma implmt_methd [rule_format (no_asm)]: 
+"\<lbrakk>wf_prog G; G\<turnstile>C\<leadsto>I\<rbrakk> \<Longrightarrow> is_iface G I \<longrightarrow>  
+ (\<forall> im    \<in>imethds G I   sig.  
+  \<exists> cm\<in>methd G C sig: \<not>is_static cm \<and> \<not> is_static im \<and> 
+                      G\<turnstile>resTy cm\<preceq>resTy im \<and>
+                      accmodi im = Public \<and> accmodi cm = Public)"
+apply (frule implmt_is_class)
+apply (erule implmt.induct)
+apply   safe
+apply   (drule (2) implmt1_methd)
+apply   fast
+apply  (drule (1) subint_widen_imethds)
+apply   simp
+apply   assumption
+apply  clarify
+apply  (drule (2) implmt1_methd)
+apply  (force)
+apply (frule subcls1D)
+apply (drule (1) bspec)
+apply clarify
+apply (drule (3) r_into_rtrancl [THEN inheritable_instance_methd_props, 
+                                 OF _ implmt_is_class])
+apply auto 
+done
+
+
+declare split_paired_All [simp] split_paired_Ex [simp]
+ML_setup {*
+claset_ref()  := claset() addSbefore ("split_all_tac", split_all_tac);
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
+*}
+lemma mheadsD [rule_format (no_asm)]: 
+"emh \<in> mheads G S t sig \<longrightarrow> wf_prog G \<longrightarrow>
+ (\<exists>C D m. t = ClassT C \<and> declrefT emh = ClassT D \<and> 
+          accmethd G S C sig = Some m \<and>
+          (declclass m = D) \<and> mhead (mthd m) = (mhd emh)) \<or>
+ (\<exists>I. t = IfaceT I \<and> ((\<exists>im. im  \<in> accimethds G (pid S) I sig \<and> 
+          mthd im = mhd emh) \<or> 
+  (\<exists>m. G\<turnstile>Iface I accessible_in (pid S) \<and> accmethd G S Object sig = Some m \<and> 
+       accmodi m \<noteq> Private \<and> 
+       declrefT emh = ClassT Object \<and> mhead (mthd m) = mhd emh))) \<or>
+ (\<exists>T m. t = ArrayT T \<and> G\<turnstile>Array T accessible_in (pid S) \<and>
+        accmethd G S Object sig = Some m \<and> accmodi m \<noteq> Private \<and> 
+        declrefT emh = ClassT Object \<and> mhead (mthd m) = mhd emh)"
+apply (rule_tac "ref_ty1"="t" in ref_ty_ty.induct [THEN conjunct1])
+apply auto
+apply (auto simp add: cmheads_def accObjectmheads_def Objectmheads_def)
+apply (auto  dest!: accmethd_SomeD)
+done
+
+lemma mheads_cases [consumes 2, case_names Class_methd 
+                    Iface_methd Iface_Object_methd Array_Object_methd]: 
+"\<lbrakk>emh \<in> mheads G S t sig; wf_prog G;
+ \<And> C D m. \<lbrakk>t = ClassT C;declrefT emh = ClassT D; accmethd G S C sig = Some m;
+           (declclass m = D); mhead (mthd m) = (mhd emh)\<rbrakk> \<Longrightarrow> P emh; 
+ \<And> I im. \<lbrakk>t = IfaceT I; im  \<in> accimethds G (pid S) I sig; mthd im = mhd emh\<rbrakk>
+          \<Longrightarrow> P emh;
+ \<And> I m. \<lbrakk>t = IfaceT I; G\<turnstile>Iface I accessible_in (pid S);
+          accmethd G S Object sig = Some m; accmodi m \<noteq> Private;
+         declrefT emh = ClassT Object; mhead (mthd m) = mhd emh\<rbrakk> \<Longrightarrow> P emh;
+ \<And> T m. \<lbrakk>t = ArrayT T;G\<turnstile>Array T accessible_in (pid S);
+          accmethd G S Object sig = Some m; accmodi m \<noteq> Private; 
+          declrefT emh = ClassT Object; mhead (mthd m) = mhd emh\<rbrakk> \<Longrightarrow>  P emh 
+\<rbrakk> \<Longrightarrow> P emh"
+by (blast dest!: mheadsD)
+
+lemma declclassD[rule_format]:
+ "\<lbrakk>wf_prog G;class G C = Some c; methd G C sig = Some m; 
+   class G (declclass m) = Some d\<rbrakk>
+  \<Longrightarrow> table_of (methods d) sig  = Some (mthd m)"
+proof -
+  assume    wf: "wf_prog G"
+  then have ws: "ws_prog G" ..
+  assume  clsC: "class G C = Some c"
+  from clsC ws 
+  show "\<And> m d. \<lbrakk>methd G C sig = Some m; class G (declclass m) = Some d\<rbrakk>
+        \<Longrightarrow> table_of (methods d) sig  = Some (mthd m)" 
+         (is "PROP ?P C") 
+  proof (induct ?P C rule: ws_class_induct)
+    case Object
+    fix m d
+    assume "methd G Object sig = Some m" 
+           "class G (declclass m) = Some d"
+    with wf show "?thesis m d" by auto
+  next
+    case Subcls
+    fix C c m d
+    assume hyp: "PROP ?P (super c)"
+    and      m: "methd G C sig = Some m"
+    and  declC: "class G (declclass m) = Some d"
+    and   clsC: "class G C = Some c"
+    and   nObj: "C \<noteq> Object"
+    let ?newMethods = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig"
+    show "?thesis m d" 
+    proof (cases "?newMethods")
+      case None
+      from None clsC nObj ws m declC hyp  
+      show "?thesis" by (auto simp add: methd_rec)
+    next
+      case Some
+      from Some clsC nObj ws m declC hyp  
+      show "?thesis" 
+	by (auto simp add: methd_rec
+                     dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
+    qed
+  qed
+qed
+
+(* Tactical version *)
+(*
+lemma declclassD[rule_format]:
+ "wf_prog G \<longrightarrow>  
+ (\<forall> c d m. class G C = Some c \<longrightarrow> methd G C sig = Some m \<longrightarrow> 
+  class G (declclass m) = Some d
+ \<longrightarrow> table_of (methods d) sig  = Some (mthd m))"
+apply (rule class_rec.induct)
+apply (rule impI)
+apply (rule allI)+
+apply (rule impI)
+apply (case_tac "C=Object")
+apply   (force simp add: methd_rec)
+
+apply   (subst methd_rec)
+apply     (blast dest: wf_ws_prog)+
+apply   (case_tac "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig")
+apply     (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
+done
+*)
+
+lemma dynmethd_Object:
+ (assumes statM: "methd G Object sig = Some statM" and
+        private: "accmodi statM = Private" and 
+       is_cls_C: "is_class G C" and
+             wf: "wf_prog G"
+ ) "dynmethd G Object C sig = Some statM"
+proof -
+  from is_cls_C wf 
+  have subclseq: "G\<turnstile>C \<preceq>\<^sub>C Object" 
+    by (auto intro: subcls_ObjectI)
+  from wf have ws: "ws_prog G" 
+    by simp
+  from wf 
+  have is_cls_Obj: "is_class G Object" 
+    by simp
+  from statM subclseq is_cls_Obj ws private
+  show ?thesis
+  proof (cases rule: dynmethd_cases)
+    case Static then show ?thesis .
+  next
+    case Overrides 
+    with private show ?thesis 
+      by (auto dest: no_Private_override)
+  qed
+qed
+
+lemma wf_imethds_hiding_objmethdsD: 
+  (assumes     old: "methd G Object sig = Some old" and
+           is_if_I: "is_iface G I" and
+                wf: "wf_prog G" and    
+       not_private: "accmodi old \<noteq> Private" and
+               new: "new \<in> imethds G I sig" 
+  ) "G\<turnstile>resTy new\<preceq>resTy old \<and> is_static new = is_static old" (is "?P new")
+proof -
+  from wf have ws: "ws_prog G" by simp
+  {
+    fix I i new
+    assume ifI: "iface G I = Some i"
+    assume new: "table_of (imethods i) sig = Some new" 
+    from ifI new not_private wf old  
+    have "?P (I,new)"
+      by (auto dest!: wf_prog_idecl wf_idecl_hiding cond_hiding_entailsD
+            simp del: methd_Object)
+  } note hyp_newmethod = this  
+  from is_if_I ws new 
+  show ?thesis
+  proof (induct rule: ws_interface_induct)
+    case (Step I i)
+    assume ifI: "iface G I = Some i" 
+    assume new: "new \<in> imethds G I sig" 
+    from Step
+    have hyp: "\<forall> J \<in> set (isuperIfs i). (new \<in> imethds G J sig \<longrightarrow> ?P new)"
+      by auto 
+    from new ifI ws
+    show "?P new"
+    proof (cases rule: imethds_cases)
+      case NewMethod
+      with ifI hyp_newmethod
+      show ?thesis
+	by auto
+    next
+      case (InheritedMethod J)
+      assume "J \<in> set (isuperIfs i)" 
+             "new \<in> imethds G J sig"
+      with hyp 
+      show "?thesis"
+	by auto
+    qed
+  qed
+qed
+
+text {*
+Which dynamic classes are valid to look up a member of a distinct static type?
+We have to distinct class members (named static members in Java) 
+from instance members. Class members are global to all Objects of a class,
+instance members are local to a single Object instance. If a member is
+equipped with the static modifier it is a class member, else it is an 
+instance member.
+The following table gives an overview of the current framework. We assume
+to have a reference with static type statT and a dynamic class dynC. Between
+both of these types the widening relation holds 
+@{term "G\<turnstile>Class dynC\<preceq> statT"}. Unfortunately this ordinary widening relation 
+isn't enough to describe the valid lookup classes, since we must cope the
+special cases of arrays and interfaces,too. If we statically expect an array or
+inteface we may lookup a field or a method in Object which isn't covered in 
+the widening relation.
+\begin{verbatim}
+statT      field         instance method       static (class) method
+------------------------------------------------------------------------
+ NullT      /                  /                   /
+ Iface      /                dynC                Object
+ Class    dynC               dynC                 dynC
+ Array      /                Object              Object
+\end{verbatim}
+In most cases we con lookup the member in the dynamic class. But as an
+interface can't declare new static methods, nor an array can define new
+methods at all, we have to lookup methods in the base class Object.
+
+The limitation to classes in the field column is artificial  and comes out
+of the typing rule for the field access (see rule @{text "FVar"} in the 
+welltyping relation @{term "wt"} in theory WellType). 
+I stems out of the fact, that Object
+indeed has no non private fields. So interfaces and arrays can actually
+have no fields at all and a field access would be senseless. (In Java
+interfaces are allowed to declare new fields but in current Bali not!).
+So there is no principal reason why we should not allow Objects to declare
+non private fields. Then we would get the following column:
+\begin{verbatim}
+ statT    field
+----------------- 
+ NullT      /  
+ Iface    Object 
+ Class    dynC 
+ Array    Object
+\end{verbatim}
+*}
+consts valid_lookup_cls:: "prog \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> bool \<Rightarrow> bool"
+                        ("_,_ \<turnstile> _ valid'_lookup'_cls'_for _" [61,61,61,61] 60)
+primrec
+"G,NullT    \<turnstile> dynC valid_lookup_cls_for static_membr = False"
+"G,IfaceT I \<turnstile> dynC valid_lookup_cls_for static_membr 
+              = (if static_membr 
+                    then dynC=Object 
+                    else G\<turnstile>Class dynC\<preceq> Iface I)"
+"G,ClassT C \<turnstile> dynC valid_lookup_cls_for static_membr = G\<turnstile>Class dynC\<preceq> Class C"
+"G,ArrayT T \<turnstile> dynC valid_lookup_cls_for static_membr = (dynC=Object)"
+
+lemma valid_lookup_cls_is_class:
+ (assumes dynC: "G,statT \<turnstile> dynC valid_lookup_cls_for static_membr" and
+      ty_statT: "isrtype G statT" and
+            wf: "wf_prog G"
+ ) "is_class G dynC"
+proof (cases statT)
+  case NullT
+  with dynC ty_statT show ?thesis
+    by (auto dest: widen_NT2)
+next
+  case (IfaceT I)
+  with dynC wf show ?thesis
+    by (auto dest: implmt_is_class)
+next
+  case (ClassT C)
+  with dynC ty_statT show ?thesis
+    by (auto dest: subcls_is_class2)
+next
+  case (ArrayT T)
+  with dynC wf show ?thesis
+    by (auto)
+qed
+
+declare split_paired_All [simp del] split_paired_Ex [simp del]
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac";
+claset_ref () := claset () delSWrapper "split_all_tac"
+*}
+
+lemma dynamic_mheadsD:   
+"\<lbrakk>emh \<in> mheads G S statT sig;    
+  G,statT \<turnstile> dynC valid_lookup_cls_for (is_static emh);
+  isrtype G statT; wf_prog G
+ \<rbrakk> \<Longrightarrow> \<exists>m \<in> dynlookup G statT dynC sig: 
+          is_static m=is_static emh \<and> G\<turnstile>resTy m\<preceq>resTy emh"
+proof - 
+  assume      emh: "emh \<in> mheads G S statT sig"
+  and          wf: "wf_prog G"
+  and   dynC_Prop: "G,statT \<turnstile> dynC valid_lookup_cls_for (is_static emh)"
+  and      istype: "isrtype G statT"
+  from dynC_Prop istype wf 
+  obtain y where
+    dynC: "class G dynC = Some y" 
+    by (auto dest: valid_lookup_cls_is_class)
+  from emh wf show ?thesis
+  proof (cases rule: mheads_cases)
+    case Class_methd
+    fix statC statDeclC sm
+    assume     statC: "statT = ClassT statC"
+    assume            "accmethd G S statC sig = Some sm"
+    then have     sm: "methd G statC sig = Some sm" 
+      by (blast dest: accmethd_SomeD)  
+    assume eq_mheads: "mhead (mthd sm) = mhd emh"
+    from statC 
+    have dynlookup: "dynlookup G statT dynC sig = dynmethd G statC dynC sig"
+      by (simp add: dynlookup_def)
+    from wf statC istype dynC_Prop sm 
+    obtain dm where
+      "dynmethd G statC dynC sig = Some dm"
+      "is_static dm = is_static sm" 
+      "G\<turnstile>resTy dm\<preceq>resTy sm"  
+      by (auto dest!: ws_dynmethd accmethd_SomeD)
+    with dynlookup eq_mheads 
+    show ?thesis 
+      by (cases emh type: *) (auto)
+  next
+    case Iface_methd
+    fix I im
+    assume    statI: "statT = IfaceT I" and
+          eq_mheads: "mthd im = mhd emh" and
+                     "im \<in> accimethds G (pid S) I sig" 
+    then have im: "im \<in> imethds G I sig" 
+      by (blast dest: accimethdsD)
+    with istype statI eq_mheads wf 
+    have not_static_emh: "\<not> is_static emh"
+      by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead)
+    from statI im
+    have dynlookup: "dynlookup G statT dynC sig = methd G dynC sig"
+      by (auto simp add: dynlookup_def dynimethd_def) 
+    from wf dynC_Prop statI istype im not_static_emh 
+    obtain dm where
+      "methd G dynC sig = Some dm"
+      "is_static dm = is_static im" 
+      "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd im)" 
+      by (auto dest: implmt_methd)
+    with dynlookup eq_mheads
+    show ?thesis 
+      by (cases emh type: *) (auto)
+  next
+    case Iface_Object_methd
+    fix I sm
+    assume   statI: "statT = IfaceT I" and
+                sm: "accmethd G S Object sig = Some sm" and 
+         eq_mheads: "mhead (mthd sm) = mhd emh" and
+             nPriv: "accmodi sm \<noteq> Private"
+     show ?thesis 
+     proof (cases "imethds G I sig = {}")
+       case True
+       with statI 
+       have dynlookup: "dynlookup G statT dynC sig = dynmethd G Object dynC sig"
+	 by (simp add: dynlookup_def dynimethd_def)
+       from wf dynC 
+       have subclsObj: "G\<turnstile>dynC \<preceq>\<^sub>C Object"
+	 by (auto intro: subcls_ObjectI)
+       from wf dynC dynC_Prop istype sm subclsObj 
+       obtain dm where
+	 "dynmethd G Object dynC sig = Some dm"
+	 "is_static dm = is_static sm" 
+	 "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd sm)"  
+	 by (auto dest!: ws_dynmethd accmethd_SomeD 
+                  intro: class_Object [OF wf])
+       with dynlookup eq_mheads
+       show ?thesis 
+	 by (cases emh type: *) (auto)
+     next
+       case False
+       with statI
+       have dynlookup: "dynlookup G statT dynC sig = methd G dynC sig"
+	 by (simp add: dynlookup_def dynimethd_def)
+       from istype statI
+       have "is_iface G I"
+	 by auto
+       with wf sm nPriv False 
+       obtain im where
+	      im: "im \<in> imethds G I sig" and
+	 eq_stat: "is_static im = is_static sm" and
+         resProp: "G\<turnstile>resTy (mthd im)\<preceq>resTy (mthd sm)"
+	 by (auto dest: wf_imethds_hiding_objmethdsD accmethd_SomeD)
+       from im wf statI istype eq_stat eq_mheads
+       have not_static_sm: "\<not> is_static emh"
+	 by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead)
+       from im wf dynC_Prop dynC istype statI not_static_sm
+       obtain dm where
+	 "methd G dynC sig = Some dm"
+	 "is_static dm = is_static im" 
+	 "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd im)" 
+	 by (auto dest: implmt_methd)
+       with wf eq_stat resProp dynlookup eq_mheads
+       show ?thesis 
+	 by (cases emh type: *) (auto intro: widen_trans)
+     qed
+  next
+    case Array_Object_methd
+    fix T sm
+    assume statArr: "statT = ArrayT T" and
+                sm: "accmethd G S Object sig = Some sm" and 
+         eq_mheads: "mhead (mthd sm) = mhd emh" 
+    from statArr dynC_Prop wf
+    have dynlookup: "dynlookup G statT dynC sig = methd G Object sig"
+      by (auto simp add: dynlookup_def dynmethd_C_C)
+    with sm eq_mheads sm 
+    show ?thesis 
+      by (cases emh type: *) (auto dest: accmethd_SomeD)
+  qed
+qed
+
+(* Tactical version *)
+(*
+lemma dynamic_mheadsD: "  
+ \<lbrakk>emh \<in> mheads G S statT sig; wf_prog G; class G dynC = Some y;  
+   if (\<exists>T. statT=ArrayT T) then dynC=Object else G\<turnstile>Class dynC\<preceq>RefT statT; 
+   isrtype G statT\<rbrakk> \<Longrightarrow>  
+  \<exists>m \<in> dynlookup G statT dynC sig: 
+     static (mthd m)=static (mhd emh) \<and> G\<turnstile>resTy (mthd m)\<preceq>resTy (mhd emh)"
+apply (drule mheadsD)
+apply safe
+       -- reftype statT is a class  
+apply  (case_tac "\<exists>T. ClassT C = ArrayT T")
+apply    (simp)
+
+apply    (clarsimp simp add: dynlookup_def )
+apply    (frule_tac statC="C" and dynC="dynC"  and sig="sig"  
+         in ws_dynmethd)
+apply      assumption+
+apply    (case_tac "emh")  
+apply    (force dest: accmethd_SomeD)
+
+       -- reftype statT is a interface, method defined in interface 
+apply    (clarsimp simp add: dynlookup_def)
+apply    (drule (1) implmt_methd)
+apply      blast
+apply      blast
+apply    (clarify)  
+apply    (unfold dynimethd_def)
+apply    (rule_tac x="cm" in bexI)
+apply      (case_tac "emh")
+apply      force
+
+apply      force
+
+        -- reftype statT is a interface, method defined in Object 
+apply    (simp add: dynlookup_def)
+apply    (simp only: dynimethd_def)
+apply    (case_tac "imethds G I sig = {}")
+apply       simp
+apply       (frule_tac statC="Object" and dynC="dynC"  and sig="sig"  
+             in ws_dynmethd)
+apply          (blast intro: subcls_ObjectI wf_ws_prog) 
+apply          (blast dest: class_Object)
+apply       (case_tac "emh") 
+apply       (force dest: accmethd_SomeD)
+
+apply       simp
+apply       (subgoal_tac "\<exists> im. im \<in> imethds G I sig") 
+prefer 2      apply blast
+apply       clarify
+apply       (frule (1) implmt_methd)
+apply         simp
+apply         blast  
+apply       (clarify dest!: accmethd_SomeD)
+apply       (frule (4) iface_overrides_Object)
+apply       clarify
+apply       (case_tac emh)
+apply       force
+
+        -- reftype statT is a array
+apply    (simp add: dynlookup_def)
+apply    (case_tac emh)
+apply    (force dest: accmethd_SomeD simp add: dynmethd_def)
+done
+*)
+
+(* ### auf ws_class_induct umstellen *)
+lemma methd_declclass:
+"\<lbrakk>class G C = Some c; wf_prog G; methd G C sig = Some m\<rbrakk> 
+ \<Longrightarrow> methd G (declclass m) sig = Some m"
+proof -
+  assume asm: "class G C = Some c" "wf_prog G" "methd G C sig = Some m"
+  have "wf_prog G  \<longrightarrow> 
+	   (\<forall> c m. class G C = Some c \<longrightarrow>  methd G C sig = Some m 
+                   \<longrightarrow>  methd G (declclass m) sig = Some m)"      (is "?P G C") 
+  proof (rule class_rec.induct,intro allI impI)
+    fix G C c m
+    assume hyp: "\<forall>c. C \<noteq> Object \<and> ws_prog G \<and> class G C = Some c \<longrightarrow>
+                     ?P G (super c)"
+    assume wf: "wf_prog G" and cls_C: "class G C = Some c" and
+            m: "methd G C sig = Some m"
+    show "methd G (declclass m) sig = Some m"
+    proof (cases "C=Object")
+      case True
+      with wf m show ?thesis by (auto intro: table_of_map_SomeI)
+    next
+      let ?filter="filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)"
+      let ?table = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
+      case False
+      with cls_C wf m
+      have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m "
+	by (simp add: methd_rec)
+      show ?thesis
+      proof (cases "?table sig")
+	case None
+	from this methd_C have "?filter (methd G (super c)) sig = Some m"
+	  by simp
+	moreover
+	from wf cls_C False obtain sup where "class G (super c) = Some sup"
+	  by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
+	moreover note wf False cls_C hyp
+	ultimately show ?thesis by auto
+      next
+	case Some
+	from this methd_C m show ?thesis by auto 
+      qed
+    qed
+  qed	
+  with asm show ?thesis by auto
+qed
+
+lemma dynmethd_declclass:
+ "\<lbrakk>dynmethd G statC dynC sig = Some m;
+   wf_prog G; is_class G statC
+  \<rbrakk> \<Longrightarrow> methd G (declclass m) sig = Some m"
+by (auto dest: dynmethd_declC)
+
+lemma dynlookup_declC:
+ "\<lbrakk>dynlookup G statT dynC sig = Some m; wf_prog G;
+   is_class G dynC;isrtype G statT
+  \<rbrakk> \<Longrightarrow> G\<turnstile>dynC \<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m)"
+by (cases "statT")
+   (auto simp add: dynlookup_def dynimethd_def 
+             dest: methd_declC dynmethd_declC)
+
+lemma dynlookup_Array_declclassD [simp]:
+"\<lbrakk>dynlookup G (ArrayT T) Object sig = Some dm;wf_prog G\<rbrakk> 
+ \<Longrightarrow> declclass dm = Object"
+proof -
+  assume dynL: "dynlookup G (ArrayT T) Object sig = Some dm"
+  assume wf: "wf_prog G"
+  from wf have ws: "ws_prog G" by auto
+  from wf have is_cls_Obj: "is_class G Object" by auto
+  from dynL wf
+  show ?thesis
+    by (auto simp add: dynlookup_def dynmethd_C_C [OF is_cls_Obj ws]
+                 dest: methd_Object_SomeD)
+qed   
+  
+declare split_paired_All [simp del] split_paired_Ex [simp del]
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac";
+claset_ref () := claset () delSWrapper "split_all_tac"
+*}
+
+lemma wt_is_type: "E,dt\<Turnstile>v\<Colon>T \<Longrightarrow>  wf_prog (prg E) \<longrightarrow> 
+  dt=empty_dt \<longrightarrow> (case T of 
+                     Inl T \<Rightarrow> is_type (prg E) T 
+                   | Inr Ts \<Rightarrow> Ball (set Ts) (is_type (prg E)))"
+apply (unfold empty_dt_def)
+apply (erule wt.induct)
+apply (auto split del: split_if_asm simp del: snd_conv 
+            simp add: is_acc_class_def is_acc_type_def)
+apply    (erule typeof_empty_is_type)
+apply   (frule (1) wf_prog_cdecl [THEN wf_cdecl_supD], rotate_tac -1, 
+        force simp del: snd_conv, clarsimp simp add: is_acc_class_def)
+apply  (drule (1) max_spec2mheads [THEN conjunct1, THEN mheadsD])
+apply  (drule_tac [2] accfield_fields) 
+apply  (frule class_Object)
+apply  (auto dest: accmethd_rT_is_type 
+                   imethds_wf_mhead [THEN conjunct1, THEN rT_is_acc_type]
+             dest!:accimethdsD
+             simp del: class_Object
+             simp add: is_acc_type_def
+    )
+done
+declare split_paired_All [simp] split_paired_Ex [simp]
+ML_setup {*
+claset_ref()  := claset() addSbefore ("split_all_tac", split_all_tac);
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
+*}
+
+lemma ty_expr_is_type: 
+"\<lbrakk>E\<turnstile>e\<Colon>-T; wf_prog (prg E)\<rbrakk> \<Longrightarrow> is_type (prg E) T"
+by (auto dest!: wt_is_type)
+lemma ty_var_is_type: 
+"\<lbrakk>E\<turnstile>v\<Colon>=T; wf_prog (prg E)\<rbrakk> \<Longrightarrow> is_type (prg E) T"
+by (auto dest!: wt_is_type)
+lemma ty_exprs_is_type: 
+"\<lbrakk>E\<turnstile>es\<Colon>\<doteq>Ts; wf_prog (prg E)\<rbrakk> \<Longrightarrow> Ball (set Ts) (is_type (prg E))"
+by (auto dest!: wt_is_type)
+
+
+lemma static_mheadsD: 
+ "\<lbrakk> emh \<in> mheads G S t sig; wf_prog G; E\<turnstile>e\<Colon>-RefT t; prg E=G ; 
+   invmode (mhd emh) e \<noteq> IntVir 
+  \<rbrakk> \<Longrightarrow> \<exists>m. (   (\<exists> C. t = ClassT C \<and> accmethd G S C sig = Some m)
+               \<or> (\<forall> C. t \<noteq> ClassT C \<and> accmethd G S Object sig = Some m )) \<and> 
+          declrefT emh = ClassT (declclass m) \<and>  mhead (mthd m) = (mhd emh)"
+apply (subgoal_tac "is_static emh \<or> e = Super")
+defer apply (force simp add: invmode_def)
+apply (frule  ty_expr_is_type)
+apply   simp
+apply (case_tac "is_static emh")
+apply  (frule (1) mheadsD)
+apply  clarsimp
+apply  safe
+apply    blast
+apply   (auto dest!: imethds_wf_mhead
+                     accmethd_SomeD 
+                     accimethdsD
+              simp add: accObjectmheads_def Objectmheads_def)
+
+apply  (erule wt_elim_cases)
+apply  (force simp add: cmheads_def)
+done
+
+lemma wt_MethdI:  
+"\<lbrakk>methd G C sig = Some m; wf_prog G;  
+  class G C = Some c\<rbrakk> \<Longrightarrow>  
+ \<exists>T. \<lparr>prg=G,cls=(declclass m),
+      lcl=\<lambda> k. 
+          (case k of
+             EName e 
+             \<Rightarrow> (case e of 
+                   VNam v 
+                   \<Rightarrow> (table_of (lcls (mbody (mthd m)))
+                                ((pars (mthd m))[\<mapsto>](parTs sig))) v
+                 | Res \<Rightarrow> Some (resTy (mthd m)))
+           | This \<Rightarrow> if is_static m then None else Some (Class (declclass m)))
+     \<rparr>\<turnstile> Methd C sig\<Colon>-T \<and> G\<turnstile>T\<preceq>resTy m"
+apply (frule (2) methd_wf_mdecl, clarify)
+apply (force dest!: wf_mdecl_bodyD intro!: wt.Methd)
+done
+
+subsection "accessibility concerns"
+
+lemma mheads_type_accessible:
+ "\<lbrakk>emh \<in> mheads G S T sig; wf_prog G\<rbrakk>
+ \<Longrightarrow> G\<turnstile>RefT T accessible_in (pid S)"
+by (erule mheads_cases)
+   (auto dest: accmethd_SomeD accessible_from_commonD accimethdsD)
+
+lemma static_to_dynamic_accessible_from:
+"\<lbrakk>G\<turnstile>m of C accessible_from accC;wf_prog G\<rbrakk> 
+ \<Longrightarrow> G\<turnstile>m in C dyn_accessible_from accC"
+proof (induct rule: accessible_fromR.induct)
+qed (auto intro: dyn_accessible_fromR.intros 
+                 member_of_to_member_in
+                 static_to_dynamic_overriding)
+
+lemma static_to_dynamic_accessible_from:
+ (assumes stat_acc: "G\<turnstile>m of statC accessible_from accC" and
+          subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
+                wf: "wf_prog G"
+ ) "G\<turnstile>m in dynC dyn_accessible_from accC"
+proof - 
+  from stat_acc subclseq 
+  show ?thesis (is "?Dyn_accessible m")
+  proof (induct rule: accessible_fromR.induct)
+    case (immediate statC m)
+    then show "?Dyn_accessible m"
+      by (blast intro: dyn_accessible_fromR.immediate
+                       member_inI
+                       permits_acc_inheritance)
+  next
+    case (overriding _ _ m)
+    with wf show "?Dyn_accessible m"
+      by (blast intro: dyn_accessible_fromR.overriding
+                       member_inI
+                       static_to_dynamic_overriding  
+                       rtrancl_trancl_trancl 
+                       static_to_dynamic_accessible_from)
+  qed
+qed
+
+lemma dynmethd_member_in:
+ (assumes    m: "dynmethd G statC dynC sig = Some m" and
+   iscls_statC: "is_class G statC" and
+            wf: "wf_prog G"
+ ) "G\<turnstile>Methd sig m member_in dynC"
+proof -
+  from m 
+  have subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC"
+    by (auto simp add: dynmethd_def)
+  from subclseq iscls_statC 
+  have iscls_dynC: "is_class G dynC"
+    by (rule subcls_is_class2)
+  from  iscls_dynC iscls_statC wf m
+  have "G\<turnstile>dynC \<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m) \<and>
+        methd G (declclass m) sig = Some m" 
+    by - (drule dynmethd_declC, auto)
+  with wf 
+  show ?thesis
+    by (auto intro: member_inI dest: methd_member_of)
+qed
+
+lemma dynmethd_access_prop:
+  (assumes statM: "methd G statC sig = Some statM" and
+        stat_acc: "G\<turnstile>Methd sig statM of statC accessible_from accC" and
+            dynM: "dynmethd G statC dynC sig = Some dynM" and
+              wf: "wf_prog G" 
+   ) "G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
+proof -
+  from wf have ws: "ws_prog G" ..
+  from dynM 
+  have subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC"
+    by (auto simp add: dynmethd_def)
+  from stat_acc 
+  have is_cls_statC: "is_class G statC"
+    by (auto dest: accessible_from_commonD member_of_is_classD)
+  with subclseq 
+  have is_cls_dynC: "is_class G dynC"
+    by (rule subcls_is_class2)
+  from is_cls_statC statM wf 
+  have member_statC: "G\<turnstile>Methd sig statM member_of statC"
+    by (auto intro: methd_member_of)
+  from stat_acc 
+  have statC_acc: "G\<turnstile>Class statC accessible_in (pid accC)"
+    by (auto dest: accessible_from_commonD)
+  from statM subclseq is_cls_statC ws 
+  show ?thesis
+  proof (cases rule: dynmethd_cases)
+    case Static
+    assume dynmethd: "dynmethd G statC dynC sig = Some statM"
+    with dynM have eq_dynM_statM: "dynM=statM" 
+      by simp
+    with stat_acc subclseq wf 
+    show ?thesis
+      by (auto intro: static_to_dynamic_accessible_from)
+  next
+    case (Overrides newM)
+    assume dynmethd: "dynmethd G statC dynC sig = Some newM"
+    assume override: "G,sig\<turnstile>newM overrides statM"
+    assume      neq: "newM\<noteq>statM"
+    from dynmethd dynM 
+    have eq_dynM_newM: "dynM=newM" 
+      by simp
+    from dynmethd eq_dynM_newM wf is_cls_statC
+    have "G\<turnstile>Methd sig dynM member_in dynC"
+      by (auto intro: dynmethd_member_in)
+    moreover
+    from subclseq
+    have "G\<turnstile>dynC\<prec>\<^sub>C statC"
+    proof (cases rule: subclseq_cases)
+      case Eq
+      assume "dynC=statC"
+      moreover
+      from is_cls_statC obtain c
+	where "class G statC = Some c"
+	by auto
+      moreover 
+      note statM ws dynmethd 
+      ultimately
+      have "newM=statM" 
+	by (auto simp add: dynmethd_C_C)
+      with neq show ?thesis 
+	by (contradiction)
+    next
+      case Subcls show ?thesis .
+    qed 
+    moreover
+    from stat_acc wf 
+    have "G\<turnstile>Methd sig statM in statC dyn_accessible_from accC"
+      by (blast intro: static_to_dynamic_accessible_from)
+    moreover
+    note override eq_dynM_newM
+    ultimately show ?thesis
+      by (cases dynM,cases statM) (auto intro: dyn_accessible_fromR.overriding)
+  qed
+qed
+
+lemma implmt_methd_access:
+(fixes accC::qtname
+ assumes iface_methd: "imethds G I sig \<noteq> {}" and
+           implements: "G\<turnstile>dynC\<leadsto>I"  and
+               isif_I: "is_iface G I" and
+                   wf: "wf_prog G" 
+ ) "\<exists> dynM. methd G dynC sig = Some dynM \<and> 
+            G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
+proof -
+  from implements 
+  have iscls_dynC: "is_class G dynC" by (rule implmt_is_class)
+  from iface_methd
+  obtain im
+    where "im \<in> imethds G I sig"
+    by auto
+  with wf implements isif_I 
+  obtain dynM 
+    where dynM: "methd G dynC sig = Some dynM" and
+           pub: "accmodi dynM = Public"
+    by (blast dest: implmt_methd)
+  with iscls_dynC wf
+  have "G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
+    by (auto intro!: dyn_accessible_fromR.immediate 
+              intro: methd_member_of member_of_to_member_in
+                     simp add: permits_acc_def)
+  with dynM    
+  show ?thesis
+    by blast
+qed
+
+corollary implmt_dynimethd_access:
+(fixes accC::qtname
+ assumes iface_methd: "imethds G I sig \<noteq> {}" and
+           implements: "G\<turnstile>dynC\<leadsto>I"  and
+               isif_I: "is_iface G I" and
+                   wf: "wf_prog G" 
+ ) "\<exists> dynM. dynimethd G I dynC sig = Some dynM \<and> 
+            G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
+proof -
+  from iface_methd
+  have "dynimethd G I dynC sig = methd G dynC sig"
+    by (simp add: dynimethd_def)
+  with iface_methd implements isif_I wf 
+  show ?thesis
+    by (simp only:)
+       (blast intro: implmt_methd_access)
+qed
+
+lemma dynlookup_access_prop:
+ (assumes emh: "emh \<in> mheads G accC statT sig" and
+         dynM: "dynlookup G statT dynC sig = Some dynM" and
+    dynC_prop: "G,statT \<turnstile> dynC valid_lookup_cls_for is_static emh" and
+    isT_statT: "isrtype G statT" and
+           wf: "wf_prog G"
+ ) "G \<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
+proof -
+  from emh wf
+  have statT_acc: "G\<turnstile>RefT statT accessible_in (pid accC)"
+    by (rule mheads_type_accessible)
+  from dynC_prop isT_statT wf
+  have iscls_dynC: "is_class G dynC"
+    by (rule valid_lookup_cls_is_class)
+  from emh dynC_prop isT_statT wf dynM
+  have eq_static: "is_static emh = is_static dynM"
+    by (auto dest: dynamic_mheadsD)
+  from emh wf show ?thesis
+  proof (cases rule: mheads_cases)
+    case (Class_methd statC _ statM)
+    assume statT: "statT = ClassT statC"
+    assume "accmethd G accC statC sig = Some statM"
+    then have    statM: "methd G statC sig = Some statM" and
+              stat_acc: "G\<turnstile>Methd sig statM of statC accessible_from accC"
+      by (auto dest: accmethd_SomeD)
+    from dynM statT
+    have dynM': "dynmethd G statC dynC sig = Some dynM"
+      by (simp add: dynlookup_def) 
+    from statM stat_acc wf dynM'
+    show ?thesis
+      by (auto dest!: dynmethd_access_prop)
+  next
+    case (Iface_methd I im)
+    then have iface_methd: "imethds G I sig \<noteq> {}" and
+                 statT_acc: "G\<turnstile>RefT statT accessible_in (pid accC)" 
+      by (auto dest: accimethdsD)
+    assume   statT: "statT = IfaceT I"
+    assume      im: "im \<in>  accimethds G (pid accC) I sig"
+    assume eq_mhds: "mthd im = mhd emh"
+    from dynM statT
+    have dynM': "dynimethd G I dynC sig = Some dynM"
+      by (simp add: dynlookup_def)
+    from isT_statT statT 
+    have isif_I: "is_iface G I"
+      by simp
+    show ?thesis
+    proof (cases "is_static emh")
+      case False
+      with statT dynC_prop 
+      have widen_dynC: "G\<turnstile>Class dynC \<preceq> RefT statT"
+	by simp
+      from statT widen_dynC
+      have implmnt: "G\<turnstile>dynC\<leadsto>I"
+	by auto    
+      from eq_static False 
+      have not_static_dynM: "\<not> is_static dynM" 
+	by simp
+      from iface_methd implmnt isif_I wf dynM'
+      show ?thesis
+	by - (drule implmt_dynimethd_access, auto)
+    next
+      case True
+      assume "is_static emh"
+      moreover
+      from wf isT_statT statT im 
+      have "\<not> is_static im"
+	by (auto dest: accimethdsD wf_prog_idecl imethds_wf_mhead)
+      moreover note eq_mhds
+      ultimately show ?thesis
+	by (cases emh) auto
+    qed
+  next
+    case (Iface_Object_methd I statM)
+    assume statT: "statT = IfaceT I"
+    assume "accmethd G accC Object sig = Some statM"
+    then have    statM: "methd G Object sig = Some statM" and
+              stat_acc: "G\<turnstile>Methd sig statM of Object accessible_from accC"
+      by (auto dest: accmethd_SomeD)
+    assume not_Private_statM: "accmodi statM \<noteq> Private"
+    assume eq_mhds: "mhead (mthd statM) = mhd emh"
+    from iscls_dynC wf
+    have widen_dynC_Obj: "G\<turnstile>dynC \<preceq>\<^sub>C Object"
+      by (auto intro: subcls_ObjectI)
+    show ?thesis
+    proof (cases "imethds G I sig = {}")
+      case True
+      from dynM statT True
+      have dynM': "dynmethd G Object dynC sig = Some dynM"
+	by (simp add: dynlookup_def dynimethd_def)
+      from statT  
+      have "G\<turnstile>RefT statT \<preceq>Class Object"
+	by auto
+      with statM statT_acc stat_acc widen_dynC_Obj statT isT_statT 
+        wf dynM' eq_static dynC_prop  
+      show ?thesis
+	by - (drule dynmethd_access_prop,force+) 
+    next
+      case False
+      then obtain im where
+	im: "im \<in>  imethds G I sig"
+	by auto
+      have not_static_emh: "\<not> is_static emh"
+      proof -
+	from im statM statT isT_statT wf not_Private_statM 
+	have "is_static im = is_static statM"
+	  by (auto dest: wf_imethds_hiding_objmethdsD)
+	with wf isT_statT statT im 
+	have "\<not> is_static statM"
+	  by (auto dest: wf_prog_idecl imethds_wf_mhead)
+	with eq_mhds
+	show ?thesis  
+	  by (cases emh) auto
+      qed
+      with statT dynC_prop
+      have implmnt: "G\<turnstile>dynC\<leadsto>I"
+	by simp
+      with isT_statT statT
+      have isif_I: "is_iface G I"
+	by simp
+      from dynM statT
+      have dynM': "dynimethd G I dynC sig = Some dynM"
+	by (simp add: dynlookup_def) 
+      from False implmnt isif_I wf dynM'
+      show ?thesis
+	by - (drule implmt_dynimethd_access, auto)
+    qed
+  next
+    case (Array_Object_methd T statM)
+    assume statT: "statT = ArrayT T"
+    assume "accmethd G accC Object sig = Some statM"
+    then have    statM: "methd G Object sig = Some statM" and
+              stat_acc: "G\<turnstile>Methd sig statM of Object accessible_from accC"
+      by (auto dest: accmethd_SomeD)
+    from statT dynC_prop
+    have dynC_Obj: "dynC = Object" 
+      by simp
+    then
+    have widen_dynC_Obj: "G\<turnstile>Class dynC \<preceq> Class Object"
+      by simp
+    from dynM statT    
+    have dynM': "dynmethd G Object dynC sig = Some dynM"
+      by (simp add: dynlookup_def)
+    from statM statT_acc stat_acc dynM' wf widen_dynC_Obj  
+         statT isT_statT  
+    show ?thesis   
+      by - (drule dynmethd_access_prop, simp+) 
+  qed
+qed
+
+lemma dynlookup_access:
+ (assumes emh: "emh \<in> mheads G accC statT sig" and
+    dynC_prop: "G,statT \<turnstile> dynC valid_lookup_cls_for (is_static emh) " and
+    isT_statT: "isrtype G statT" and
+           wf: "wf_prog G"
+ ) "\<exists> dynM. dynlookup G statT dynC sig = Some dynM \<and> 
+            G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
+proof - 
+  from dynC_prop isT_statT wf
+  have is_cls_dynC: "is_class G dynC"
+    by (auto dest: valid_lookup_cls_is_class)
+  with emh wf dynC_prop isT_statT
+  obtain dynM where 
+    "dynlookup G statT dynC sig = Some dynM"
+    by - (drule dynamic_mheadsD,auto)
+  with  emh dynC_prop isT_statT wf
+  show ?thesis 
+    by (blast intro: dynlookup_access_prop)
+qed
+
+lemma stat_overrides_Package_old: 
+(assumes stat_override: "G \<turnstile> new overrides\<^sub>S old" and 
+          accmodi_new: "accmodi new = Package" and
+                   wf: "wf_prog G "
+) "accmodi old = Package"
+proof -
+  from stat_override wf 
+  have "accmodi old \<le> accmodi new"
+    by (auto dest: wf_prog_stat_overridesD)
+  with stat_override accmodi_new show ?thesis
+    by (cases "accmodi old") (auto dest: no_Private_stat_override 
+                                   dest: acc_modi_le_Dests)
+qed
+
+text {* @{text dyn_accessible_Package} only works with the @{text wf_prog} assumption. 
+Without it. it is easy to leaf the Package!
+*}
+lemma dyn_accessible_Package:
+ "\<lbrakk>G \<turnstile> m in C dyn_accessible_from accC; accmodi m = Package;
+   wf_prog G\<rbrakk>
+  \<Longrightarrow> pid accC = pid (declclass m)"
+proof -
+  assume wf: "wf_prog G "
+  assume accessible: "G \<turnstile> m in C dyn_accessible_from accC"
+  then show "accmodi m = Package 
+            \<Longrightarrow> pid accC = pid (declclass m)"
+    (is "?Pack m \<Longrightarrow> ?P m")
+  proof (induct rule: dyn_accessible_fromR.induct)
+    case (immediate C m)
+    assume "G\<turnstile>m member_in C"
+           "G \<turnstile> m in C permits_acc_to accC"
+           "accmodi m = Package"      
+    then show "?P m"
+      by (auto simp add: permits_acc_def)
+  next
+    case (overriding declC C new newm old Sup)
+    assume member_new: "G \<turnstile> new member_in C" and
+                  new: "new = (declC, mdecl newm)" and
+             override: "G \<turnstile> (declC, newm) overrides old" and
+         subcls_C_Sup: "G\<turnstile>C \<prec>\<^sub>C Sup" and
+              acc_old: "G \<turnstile> methdMembr old in Sup dyn_accessible_from accC" and
+                  hyp: "?Pack (methdMembr old) \<Longrightarrow> ?P (methdMembr old)" and
+          accmodi_new: "accmodi new = Package"
+    from override accmodi_new new wf 
+    have accmodi_old: "accmodi old = Package"  
+      by (auto dest: overrides_Package_old)
+    with hyp 
+    have P_sup: "?P (methdMembr old)"
+      by (simp)
+    from wf override new accmodi_old accmodi_new
+    have eq_pid_new_old: "pid (declclass new) = pid (declclass old)"
+      by (auto dest: dyn_override_Package)
+    with eq_pid_new_old P_sup show "?P new"
+      by auto
+  qed
+qed
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/WellType.thy	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,598 @@
+(*  Title:      isabelle/Bali/WellType.thy
+    ID:         $Id$
+    Author:     David von Oheimb
+    Copyright   1997 Technische Universitaet Muenchen
+*)
+header {* Well-typedness of Java programs *}
+
+theory WellType = DeclConcepts:
+
+text {*
+improvements over Java Specification 1.0:
+\begin{itemize}
+\item methods of Object can be called upon references of interface or array type
+\end{itemize}
+simplifications:
+\begin{itemize}
+\item the type rules include all static checks on statements and expressions, 
+      e.g. definedness of names (of parameters, locals, fields, methods)
+\end{itemize}
+design issues:
+\begin{itemize}
+\item unified type judgment for statements, variables, expressions, 
+      expression lists
+\item statements are typed like expressions with dummy type Void
+\item the typing rules take an extra argument that is capable of determining 
+  the dynamic type of objects. Therefore, they can be used for both 
+  checking static types and determining runtime types in transition semantics.
+\end{itemize}
+*}
+
+types	lenv
+	= "(lname, ty) table"   (* local variables, including This and Result *)
+
+record env = 
+         prg:: "prog"    (* program *)
+         cls:: "qtname"  (* current package and class name *)
+         lcl:: "lenv"    (* local environment *)     
+  
+translations
+  "lenv" <= (type) "(lname, ty) table"
+  "lenv" <= (type) "lname \<Rightarrow> ty option"
+  "env" <= (type) "\<lparr>prg::prog,cls::qtname,lcl::lenv\<rparr>"
+  "env" <= (type) "\<lparr>prg::prog,cls::qtname,lcl::lenv,\<dots>::'a\<rparr>"
+
+
+
+syntax
+  pkg :: "env \<Rightarrow> pname" (* select the current package from an environment *)
+translations
+  "pkg e" == "pid (cls e)"
+
+section "Static overloading: maximally specific methods "
+
+types
+  emhead = "ref_ty \<times> mhead"
+
+(* Some mnemotic selectors for emhead *)
+constdefs 
+  "declrefT" :: "emhead \<Rightarrow> ref_ty"
+  "declrefT \<equiv> fst"
+
+  "mhd"     :: "emhead \<Rightarrow> mhead"
+  "mhd \<equiv> snd"
+
+lemma declrefT_simp[simp]:"declrefT (r,m) = r"
+by (simp add: declrefT_def)
+
+lemma mhd_simp[simp]:"mhd (r,m) = m"
+by (simp add: mhd_def)
+
+lemma static_mhd_simp[simp]: "static (mhd m) = is_static m"
+by (cases m) (simp add: member_is_static_simp mhd_def)
+
+lemma mhd_resTy_simp [simp]: "resTy (mhd m) = resTy m"
+by (cases m) simp
+
+lemma mhd_is_static_simp [simp]: "is_static (mhd m) = is_static m"
+by (cases m) simp
+
+lemma mhd_accmodi_simp [simp]: "accmodi (mhd m) = accmodi m"
+by (cases m) simp
+
+consts
+  cmheads        :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> emhead set"
+  Objectmheads   :: "prog \<Rightarrow> qtname           \<Rightarrow> sig \<Rightarrow> emhead set"
+  accObjectmheads:: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> emhead set"
+  mheads         :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> emhead set"
+defs
+ cmheads_def:
+"cmheads G S C 
+  \<equiv> \<lambda>sig. (\<lambda>(Cls,mthd). (ClassT Cls,(mhead mthd))) ` o2s (accmethd G S C sig)"
+  Objectmheads_def:
+"Objectmheads G S  
+  \<equiv> \<lambda>sig. (\<lambda>(Cls,mthd). (ClassT Cls,(mhead mthd))) 
+    ` o2s (filter_tab (\<lambda>sig m. accmodi m \<noteq> Private) (accmethd G S Object) sig)"
+  accObjectmheads_def:
+"accObjectmheads G S T
+   \<equiv> if G\<turnstile>RefT T accessible_in (pid S)
+        then Objectmheads G S
+        else \<lambda>sig. {}"
+primrec
+"mheads G S  NullT     = (\<lambda>sig. {})"
+"mheads G S (IfaceT I) = (\<lambda>sig. (\<lambda>(I,h).(IfaceT I,h)) 
+                         ` accimethds G (pid S) I sig \<union> 
+                           accObjectmheads G S (IfaceT I) sig)"
+"mheads G S (ClassT C) = cmheads G S C"
+"mheads G S (ArrayT T) = accObjectmheads G S (ArrayT T)"
+
+
+(* more detailed than necessary for type-safety, see below. *)
+constdefs
+  (* applicable methods, cf. 15.11.2.1 *)
+  appl_methds   :: "prog \<Rightarrow> qtname \<Rightarrow>  ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
+ "appl_methds G S rt \<equiv> \<lambda> sig. 
+      {(mh,pTs') |mh pTs'. mh \<in> mheads G S rt \<lparr>name=name sig,parTs=pTs'\<rparr> \<and> 
+                           G\<turnstile>(parTs sig)[\<preceq>]pTs'}"
+
+  (* more specific methods, cf. 15.11.2.2 *)
+  more_spec     :: "prog \<Rightarrow> emhead \<times> ty list \<Rightarrow> emhead \<times> ty list \<Rightarrow> bool"
+ "more_spec G \<equiv> \<lambda>(mh,pTs). \<lambda>(mh',pTs'). G\<turnstile>pTs[\<preceq>]pTs'"
+(*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'*)
+
+  (* maximally specific methods, cf. 15.11.2.2 *)
+   max_spec     :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
+
+ "max_spec G S rt sig \<equiv>{m. m \<in>appl_methds G S rt sig \<and>
+		      (\<forall>m'\<in>appl_methds G S rt sig. more_spec G m' m \<longrightarrow> m'=m)}"
+(*
+rules (* all properties of more_spec, appl_methods and max_spec we actually need
+         these can easily be proven from the above definitions *)
+
+max_spec2mheads "max_spec G S rt (mn, pTs) = insert (mh, pTs') A \<Longrightarrow>
+      mh\<in>mheads G S rt (mn, pTs') \<and> G\<turnstile>pTs[\<preceq>]pTs'"
+*)
+
+lemma max_spec2appl_meths: 
+  "x \<in> max_spec G S T sig \<Longrightarrow> x \<in> appl_methds G S T sig" 
+by (auto simp: max_spec_def)
+
+lemma appl_methsD: "(mh,pTs')\<in>appl_methds G S T \<lparr>name=mn,parTs=pTs\<rparr> \<Longrightarrow>  
+   mh \<in> mheads G S T \<lparr>name=mn,parTs=pTs'\<rparr> \<and> G\<turnstile>pTs[\<preceq>]pTs'"
+by (auto simp: appl_methds_def)
+
+lemma max_spec2mheads: 
+"max_spec G S rt \<lparr>name=mn,parTs=pTs\<rparr> = insert (mh, pTs') A 
+ \<Longrightarrow> mh \<in> mheads G S rt \<lparr>name=mn,parTs=pTs'\<rparr> \<and> G\<turnstile>pTs[\<preceq>]pTs'"
+apply (auto dest: equalityD2 subsetD max_spec2appl_meths appl_methsD)
+done
+
+
+constdefs
+  empty_dt :: "dyn_ty"
+ "empty_dt \<equiv> \<lambda>a. None"
+
+  invmode :: "('a::type)member_scheme \<Rightarrow> expr \<Rightarrow> inv_mode"
+"invmode m e \<equiv> if static m then Static else if e=Super then SuperM else IntVir"
+
+lemma invmode_nonstatic [simp]: 
+  "invmode \<lparr>access=a,static=False,\<dots>=x\<rparr> (Acc (LVar e)) = IntVir"
+apply (unfold invmode_def)
+apply (simp (no_asm))
+done
+
+
+lemma invmode_Static_eq [simp]: "(invmode m e = Static) = static m"
+apply (unfold invmode_def)
+apply (simp (no_asm))
+done
+
+
+lemma invmode_IntVir_eq: "(invmode m e = IntVir) = (\<not>(static m) \<and> e\<noteq>Super)"
+apply (unfold invmode_def)
+apply (simp (no_asm))
+done
+
+lemma Null_staticD: 
+  "a'=Null \<longrightarrow> (static m) \<Longrightarrow> invmode m e = IntVir \<longrightarrow> a' \<noteq> Null"
+apply (clarsimp simp add: invmode_IntVir_eq)
+done
+
+
+types tys  =        "ty + ty list"
+translations
+  "tys"   <= (type) "ty + ty list"
+
+consts
+  wt    :: "(env\<spacespace> \<times> dyn_ty\<spacespace> \<times>  term \<times> tys) set"
+(*wt    :: " env \<Rightarrow> dyn_ty \<Rightarrow> (term \<times> tys) set" not feasible because of 
+					      \<spacespace>  changing env in Try stmt *)
+
+syntax
+wt      :: "env \<Rightarrow> dyn_ty \<Rightarrow> [term,tys]  \<Rightarrow> bool" ("_,_|=_::_" [51,51,51,51] 50)
+wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow>  stmt       \<Rightarrow> bool" ("_,_|=_:<>" [51,51,51   ] 50)
+ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_,_|=_:-_" [51,51,51,51] 50)
+ty_var  :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_,_|=_:=_" [51,51,51,51] 50)
+ty_exprs:: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr list,
+	         \<spacespace>        \<spacespace>  ty   list] \<Rightarrow> bool" ("_,_|=_:#_" [51,51,51,51] 50)
+
+syntax (xsymbols)
+wt      :: "env \<Rightarrow> dyn_ty \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>_"  [51,51,51,51] 50)
+wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow>  stmt       \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>\<surd>"  [51,51,51   ] 50)
+ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>-_" [51,51,51,51] 50)
+ty_var  :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>=_" [51,51,51,51] 50)
+ty_exprs:: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr list,
+		  \<spacespace>        \<spacespace>  ty   list] \<Rightarrow> bool"("_,_\<Turnstile>_\<Colon>\<doteq>_" [51,51,51,51] 50)
+
+translations
+	"E,dt\<Turnstile>t\<Colon>T" == "(E,dt,t,T) \<in> wt"
+	"E,dt\<Turnstile>s\<Colon>\<surd>"  == "E,dt\<Turnstile>In1r s\<Colon>Inl (PrimT Void)"
+	"E,dt\<Turnstile>e\<Colon>-T" == "E,dt\<Turnstile>In1l e\<Colon>Inl T"
+	"E,dt\<Turnstile>e\<Colon>=T" == "E,dt\<Turnstile>In2  e\<Colon>Inl T"
+	"E,dt\<Turnstile>e\<Colon>\<doteq>T" == "E,dt\<Turnstile>In3  e\<Colon>Inr T"
+
+syntax (* for purely static typing *)
+  wt_      :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_|-_::_" [51,51,51] 50)
+  wt_stmt_ :: "env \<Rightarrow>  stmt       \<Rightarrow> bool" ("_|-_:<>" [51,51   ] 50)
+  ty_expr_ :: "env \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_|-_:-_" [51,51,51] 50)
+  ty_var_  :: "env \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_|-_:=_" [51,51,51] 50)
+  ty_exprs_:: "env \<Rightarrow> [expr list,
+		     \<spacespace> ty   list] \<Rightarrow> bool" ("_|-_:#_" [51,51,51] 50)
+
+syntax (xsymbols)
+  wt_      :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>_"  [51,51,51] 50)
+  wt_stmt_ ::  "env \<Rightarrow>  stmt       \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<surd>"  [51,51   ] 50)
+  ty_expr_ :: "env \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>-_" [51,51,51] 50)
+  ty_var_  :: "env \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>=_" [51,51,51] 50)
+  ty_exprs_ :: "env \<Rightarrow> [expr list,
+		    \<spacespace>  ty   list] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<doteq>_" [51,51,51] 50)
+
+translations
+	"E\<turnstile>t\<Colon> T" == "E,empty_dt\<Turnstile>t\<Colon> T"
+	"E\<turnstile>s\<Colon>\<surd>" == "E,empty_dt\<Turnstile>s\<Colon>\<surd>"
+	"E\<turnstile>e\<Colon>-T" == "E,empty_dt\<Turnstile>e\<Colon>-T"
+	"E\<turnstile>e\<Colon>=T" == "E,empty_dt\<Turnstile>e\<Colon>=T"
+	"E\<turnstile>e\<Colon>\<doteq>T" == "E,empty_dt\<Turnstile>e\<Colon>\<doteq>T"
+
+
+inductive wt intros 
+
+
+(* well-typed statements *)
+
+  Skip:					"E,dt\<Turnstile>Skip\<Colon>\<surd>"
+
+  Expr:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Expr e\<Colon>\<surd>"
+  (* cf. 14.6 *)
+  Lab:  "E,dt\<Turnstile>c\<Colon>\<surd> \<Longrightarrow>                   
+                                         E,dt\<Turnstile>l\<bullet> c\<Colon>\<surd>" 
+
+  Comp:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; 
+	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>c1;; c2\<Colon>\<surd>"
+
+  (* cf. 14.8 *)
+  If:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
+	  E,dt\<Turnstile>c1\<Colon>\<surd>;
+	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>If(e) c1 Else c2\<Colon>\<surd>"
+
+  (* cf. 14.10 *)
+  Loop:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
+	  E,dt\<Turnstile>c\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>l\<bullet> While(e) c\<Colon>\<surd>"
+  (* cf. 14.13, 14.15, 14.16 *)
+  Do:                                   "E,dt\<Turnstile>Do jump\<Colon>\<surd>"
+
+  (* cf. 14.16 *)
+  Throw: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class tn;
+	  prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Throw e\<Colon>\<surd>"
+  (* cf. 14.18 *)
+  Try:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable;
+	  lcl E (VName vn)=None; E \<lparr>lcl := lcl E(VName vn\<mapsto>Class tn)\<rparr>,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk>
+          \<Longrightarrow>
+					 E,dt\<Turnstile>Try c1 Catch(tn vn) c2\<Colon>\<surd>"
+
+  (* cf. 14.18 *)
+  Fin:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>c1 Finally c2\<Colon>\<surd>"
+
+  Init:	"\<lbrakk>is_class (prg E) C\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Init C\<Colon>\<surd>"
+  (* Init is created on the fly during evaluation (see Eval.thy). The class
+   * isn't necessarily accessible from the points Init is called. Therefor
+   * we only demand is_class and not is_acc_class here 
+   *)
+
+(* well-typed expressions *)
+
+  (* cf. 15.8 *)
+  NewC:	"\<lbrakk>is_acc_class (prg E) (pkg E) C\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>NewC C\<Colon>-Class C"
+  (* cf. 15.9 *)
+  NewA:	"\<lbrakk>is_acc_type (prg E) (pkg E) T;
+	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>New T[i]\<Colon>-T.[]"
+
+  (* cf. 15.15 *)
+  Cast:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T; is_acc_type (prg E) (pkg E) T';
+	  prg E\<turnstile>T\<preceq>? T'\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Cast T' e\<Colon>-T'"
+
+  (* cf. 15.19.2 *)
+  Inst:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT T; is_acc_type (prg E) (pkg E) (RefT T');
+	  prg E\<turnstile>RefT T\<preceq>? RefT T'\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>e InstOf T'\<Colon>-PrimT Boolean"
+
+  (* cf. 15.7.1 *)
+  Lit:	"\<lbrakk>typeof dt x = Some T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Lit x\<Colon>-T"
+
+  (* cf. 15.10.2, 15.11.1 *)
+  Super: "\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object;
+	  class (prg E) C = Some c\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Super\<Colon>-Class (super c)"
+
+  (* cf. 15.13.1, 15.10.1, 15.12 *)
+  Acc:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Acc va\<Colon>-T"
+
+  (* cf. 15.25, 15.25.1 *)
+  Ass:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T; va \<noteq> LVar This;
+	  E,dt\<Turnstile>v \<Colon>-T';
+	  prg E\<turnstile>T'\<preceq>T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>va:=v\<Colon>-T'"
+
+  (* cf. 15.24 *)
+  Cond:	"\<lbrakk>E,dt\<Turnstile>e0\<Colon>-PrimT Boolean;
+	  E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2;
+	  prg E\<turnstile>T1\<preceq>T2 \<and> T = T2  \<or>  prg E\<turnstile>T2\<preceq>T1 \<and> T = T1\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>e0 ? e1 : e2\<Colon>-T"
+
+  (* cf. 15.11.1, 15.11.2, 15.11.3 *)
+  Call:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT;
+	  E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;
+	  max_spec (prg E) (cls E) statT \<lparr>name=mn,parTs=pTs\<rparr> 
+            = {((statDeclT,m),pTs')}
+         \<rbrakk> \<Longrightarrow>
+		   E,dt\<Turnstile>{statT,invmode m e}e\<cdot>mn({pTs'}ps)\<Colon>-(resTy m)"
+
+  Methd: "\<lbrakk>is_class (prg E) C;
+	  methd (prg E) C sig = Some m;
+	  E,dt\<Turnstile>Body (declclass m) (stmt (mbody (mthd m)))\<Colon>-T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Methd C sig\<Colon>-T"
+ (* The class C is the dynamic class of the method call (cf. Eval.thy). 
+  * It hasn't got to be directly accessible from the current package (pkg E). 
+  * Only the static class must be accessible (enshured indirectly by Call). 
+  *)
+
+  Body:	"\<lbrakk>is_class (prg E) D;
+	  E,dt\<Turnstile>blk\<Colon>\<surd>;
+	  (lcl E) Result = Some T;
+          is_type (prg E) T\<rbrakk> \<Longrightarrow>
+   					 E,dt\<Turnstile>Body D blk\<Colon>-T"
+  (* the class D implementing the method must not directly be accessible 
+   * from the current package (pkg E), but can also be indirectly accessible 
+   * due to inheritance (enshured in Call)
+   * The result type hasn't got to be accessible in Java! (If it is not 
+   * accessible you can only assign it to Object) 
+   *)
+
+(* well-typed variables *)
+
+  (* cf. 15.13.1 *)
+  LVar:	"\<lbrakk>lcl E vn = Some T; is_acc_type (prg E) (pkg E) T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>LVar vn\<Colon>=T"
+  (* cf. 15.10.1 *)
+  FVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; 
+	  accfield (prg E) (cls E) C fn = Some (fd,f)\<rbrakk> \<Longrightarrow>
+			                 E,dt\<Turnstile>{fd,static f}e..fn\<Colon>=(type f)"
+  (* cf. 15.12 *)
+  AVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T.[]; 
+	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>e.[i]\<Colon>=T"
+
+
+(* well-typed expression lists *)
+
+  (* cf. 15.11.??? *)
+  Nil:					"E,dt\<Turnstile>[]\<Colon>\<doteq>[]"
+
+  (* cf. 15.11.??? *)
+  Cons:	"\<lbrakk>E,dt\<Turnstile>e \<Colon>-T;
+	  E,dt\<Turnstile>es\<Colon>\<doteq>Ts\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>e#es\<Colon>\<doteq>T#Ts"
+
+
+declare not_None_eq [simp del] 
+declare split_if [split del] split_if_asm [split del]
+declare split_paired_All [simp del] split_paired_Ex [simp del]
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac"
+*}
+inductive_cases wt_stmt_cases: "E,dt\<Turnstile>c\<Colon>\<surd>"
+inductive_cases wt_elim_cases:
+	"E,dt\<Turnstile>In2  (LVar vn)               \<Colon>T"
+	"E,dt\<Turnstile>In2  ({fd,s}e..fn)           \<Colon>T"
+	"E,dt\<Turnstile>In2  (e.[i])                 \<Colon>T"
+	"E,dt\<Turnstile>In1l (NewC C)                \<Colon>T"
+	"E,dt\<Turnstile>In1l (New T'[i])             \<Colon>T"
+	"E,dt\<Turnstile>In1l (Cast T' e)             \<Colon>T"
+	"E,dt\<Turnstile>In1l (e InstOf T')           \<Colon>T"
+	"E,dt\<Turnstile>In1l (Lit x)                 \<Colon>T"
+	"E,dt\<Turnstile>In1l (Super)                 \<Colon>T"
+	"E,dt\<Turnstile>In1l (Acc va)                \<Colon>T"
+	"E,dt\<Turnstile>In1l (Ass va v)              \<Colon>T"
+	"E,dt\<Turnstile>In1l (e0 ? e1 : e2)          \<Colon>T"
+	"E,dt\<Turnstile>In1l ({statT,mode}e\<cdot>mn({pT'}p))\<Colon>T"
+	"E,dt\<Turnstile>In1l (Methd C sig)           \<Colon>T"
+	"E,dt\<Turnstile>In1l (Body D blk)            \<Colon>T"
+	"E,dt\<Turnstile>In3  ([])                    \<Colon>Ts"
+	"E,dt\<Turnstile>In3  (e#es)                  \<Colon>Ts"
+	"E,dt\<Turnstile>In1r  Skip                   \<Colon>x"
+	"E,dt\<Turnstile>In1r (Expr e)                \<Colon>x"
+	"E,dt\<Turnstile>In1r (c1;; c2)               \<Colon>x"
+        "E,dt\<Turnstile>In1r (l\<bullet> c)                  \<Colon>x" 
+	"E,dt\<Turnstile>In1r (If(e) c1 Else c2)      \<Colon>x"
+	"E,dt\<Turnstile>In1r (l\<bullet> While(e) c)         \<Colon>x"
+        "E,dt\<Turnstile>In1r (Do jump)               \<Colon>x"
+	"E,dt\<Turnstile>In1r (Throw e)               \<Colon>x"
+	"E,dt\<Turnstile>In1r (Try c1 Catch(tn vn) c2)\<Colon>x"
+	"E,dt\<Turnstile>In1r (c1 Finally c2)         \<Colon>x"
+	"E,dt\<Turnstile>In1r (Init C)                \<Colon>x"
+declare not_None_eq [simp] 
+declare split_if [split] split_if_asm [split]
+declare split_paired_All [simp] split_paired_Ex [simp]
+ML_setup {*
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
+*}
+
+lemma is_acc_class_is_accessible: 
+  "is_acc_class G P C \<Longrightarrow> G\<turnstile>(Class C) accessible_in P"
+by (auto simp add: is_acc_class_def)
+
+lemma is_acc_iface_is_iface: "is_acc_iface G P I \<Longrightarrow> is_iface G I"
+by (auto simp add: is_acc_iface_def)
+
+lemma is_acc_iface_is_accessible: 
+  "is_acc_iface G P I \<Longrightarrow> G\<turnstile>(Iface I) accessible_in P"
+by (auto simp add: is_acc_iface_def)
+
+lemma is_acc_type_is_type: "is_acc_type G P T \<Longrightarrow> is_type G T"
+by (auto simp add: is_acc_type_def)
+
+lemma is_acc_iface_is_accessible: "is_acc_type G P T \<Longrightarrow> G\<turnstile>T accessible_in P"
+by (auto simp add: is_acc_type_def)
+
+lemma wt_Methd_is_methd: 
+  "E\<turnstile>In1l (Methd C sig)\<Colon>T \<Longrightarrow> is_methd (prg E) C sig"
+apply (erule_tac wt_elim_cases)
+apply clarsimp
+apply (erule is_methdI, assumption)
+done
+
+(* Special versions of some typing rules, better suited to pattern match the
+ * conclusion (no selectors in the conclusion\<dots>)
+ *)
+
+lemma wt_Super:
+"\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object;
+  class (prg E) C = Some c;D=super c\<rbrakk> 
+ \<Longrightarrow> E,dt\<Turnstile>Super\<Colon>-Class D"
+by (auto elim: wt.Super)
+ 
+lemma wt_Call: 
+"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT; E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;  
+  max_spec (prg E) (cls E) statT \<lparr>name=mn,parTs=pTs\<rparr> 
+    = {((statDeclC,m),pTs')};rT=(resTy m);   
+ mode = invmode m e\<rbrakk> \<Longrightarrow> E,dt\<Turnstile>{statT,mode}e\<cdot>mn({pTs'}ps)\<Colon>-rT"
+by (auto elim: wt.Call)
+
+
+
+lemma invocationTypeExpr_noClassD: 
+"\<lbrakk> E\<turnstile>e\<Colon>-RefT statT\<rbrakk>
+ \<Longrightarrow> (\<forall> statC. statT \<noteq> ClassT statC) \<longrightarrow> invmode m e \<noteq> SuperM"
+proof -
+  assume wt: "E\<turnstile>e\<Colon>-RefT statT"
+  show ?thesis
+  proof (cases "e=Super")
+    case True
+    with wt obtain "C" where "statT = ClassT C" by (blast elim: wt_elim_cases)
+    then show ?thesis by blast
+  next 
+    case False then show ?thesis 
+      by (auto simp add: invmode_def split: split_if_asm)
+  qed
+qed
+
+lemma wt_Super: 
+"\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object; class (prg E) C = Some c; D=super c\<rbrakk> 
+\<Longrightarrow> E,dt\<Turnstile>Super\<Colon>-Class D"
+by (auto elim: wt.Super)
+
+
+lemma wt_FVar:	
+"\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; accfield (prg E) (cls E) C fn = Some (fd,f);
+  sf=static f; fT=(type f)\<rbrakk> \<Longrightarrow> E,dt\<Turnstile>{fd,sf}e..fn\<Colon>=fT"
+by (auto elim: wt.FVar)
+
+lemma wt_init [iff]: "E,dt\<Turnstile>Init C\<Colon>\<surd> = is_class (prg E) C"
+by (auto elim: wt_elim_cases intro: "wt.Init")
+
+declare wt.Skip [iff]
+
+lemma wt_StatRef: 
+  "is_acc_type (prg E) (pkg E) (RefT rt) \<Longrightarrow> E\<turnstile>StatRef rt\<Colon>-RefT rt"
+apply (rule wt.Cast)
+apply  (rule wt.Lit)
+apply   (simp (no_asm))
+apply  (simp (no_asm_simp))
+apply (rule cast.widen)
+apply (simp (no_asm))
+done
+
+lemma wt_Inj_elim: 
+  "\<And>E. E,dt\<Turnstile>t\<Colon>U \<Longrightarrow> case t of 
+                       In1 ec \<Rightarrow> (case ec of 
+                                    Inl e \<Rightarrow> \<exists>T. U=Inl T  
+                                  | Inr s \<Rightarrow> U=Inl (PrimT Void))  
+                     | In2 e \<Rightarrow> (\<exists>T. U=Inl T) 
+                     | In3 e \<Rightarrow> (\<exists>T. U=Inr T)"
+apply (erule wt.induct)
+apply auto
+done
+
+ML {*
+fun wt_fun name inj rhs =
+let
+  val lhs = "E,dt\<Turnstile>" ^ inj ^ " t\<Colon>U"
+  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
+	(K [Auto_tac, ALLGOALS (ftac (thm "wt_Inj_elim")) THEN Auto_tac])
+  fun is_Inj (Const (inj,_) $ _) = true
+    | is_Inj _                   = false
+  fun pred (t as (_ $ (Const ("Pair",_) $
+     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
+       x))) $ _ )) = is_Inj x
+in
+  make_simproc name lhs pred (thm name)
+end
+
+val wt_expr_proc  = wt_fun "wt_expr_eq"  "In1l" "\<exists>T.  U=Inl T  \<and> E,dt\<Turnstile>t\<Colon>-T"
+val wt_var_proc   = wt_fun "wt_var_eq"   "In2"  "\<exists>T.  U=Inl T  \<and> E,dt\<Turnstile>t\<Colon>=T"
+val wt_exprs_proc = wt_fun "wt_exprs_eq" "In3"  "\<exists>Ts. U=Inr Ts \<and> E,dt\<Turnstile>t\<Colon>\<doteq>Ts"
+val wt_stmt_proc  = wt_fun "wt_stmt_eq"  "In1r" "U=Inl(PrimT Void)\<and>E,dt\<Turnstile>t\<Colon>\<surd>"
+*}
+
+ML {*
+Addsimprocs [wt_expr_proc,wt_var_proc,wt_exprs_proc,wt_stmt_proc]
+*}
+
+lemma Inj_eq_lemma [simp]: 
+  "(\<forall>T. (\<exists>T'. T = Inj T' \<and> P T') \<longrightarrow> Q T) = (\<forall>T'. P T' \<longrightarrow> Q (Inj T'))"
+by auto
+
+(* unused *)
+lemma single_valued_tys_lemma [rule_format (no_asm)]: 
+  "\<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>  
+     G = prg E \<longrightarrow> (\<forall>T'. E,dt\<Turnstile>t\<Colon>T' \<longrightarrow> T  = T')"
+apply (cases "E", erule wt.induct)
+apply (safe del: disjE)
+apply (simp_all (no_asm_use) split del: split_if_asm)
+apply (safe del: disjE)
+(* 17 subgoals *)
+apply (tactic {* ALLGOALS (fn i => if i = 9 then EVERY'[thin_tac "?E,dt\<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) *})
+(*apply (safe del: disjE elim!: wt_elim_cases)*)
+apply (tactic {*ALLGOALS (eresolve_tac (thms "wt_elim_cases"))*})
+apply (simp_all (no_asm_use) split del: split_if_asm)
+apply (erule_tac [10] V = "All ?P" in thin_rl) (* Call *)
+apply ((blast del: equalityCE dest: sym [THEN trans])+)
+done
+
+(* unused *)
+lemma single_valued_tys: 
+"ws_prog (prg E) \<Longrightarrow> single_valued {(t,T). E,dt\<Turnstile>t\<Colon>T}"
+apply (unfold single_valued_def)
+apply clarsimp
+apply (rule single_valued_tys_lemma)
+apply (auto intro!: widen_antisym)
+done
+
+lemma typeof_empty_is_type [rule_format (no_asm)]: 
+  "typeof (\<lambda>a. None) v = Some T \<longrightarrow> is_type G T"
+apply (rule val.induct)
+apply    	auto
+done
+
+(* unused *)
+lemma typeof_is_type [rule_format (no_asm)]: 
+ "(\<forall>a. v \<noteq> Addr a) \<longrightarrow> (\<exists>T. typeof dt v = Some T \<and> is_type G T)"
+apply (rule val.induct)
+prefer 5 
+apply     fast
+apply  (simp_all (no_asm))
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/document/root.tex	Mon Jan 28 17:00:19 2002 +0100
@@ -0,0 +1,238 @@
+
+\documentclass[11pt,a4paper]{book}
+\usepackage{isabelle,isabellesym}
+
+% further packages required for unusual symbols (see also isabellesym.sty)
+\usepackage{latexsym}
+%\usepackage{amssymb}
+%\usepackage[english]{babel}
+%\usepackage[latin1]{inputenc}
+%\usepackage[only,bigsqcap]{stmaryrd}
+%\usepackage{wasysym}
+%\usepackage{eufrak}
+%\usepackage{textcomp}
+%\usepackage{marvosym}
+
+% this should be the last package used
+\usepackage{pdfsetup}
+
+% proper setup for best-style documents
+\urlstyle{rm}
+\isabellestyle{it}
+
+\pagestyle{myheadings}
+
+\addtolength{\hoffset}{-1,5cm}
+\addtolength{\textwidth}{4cm}
+\addtolength{\voffset}{-2cm}
+\addtolength{\textheight}{4cm}
+
+%subsection instead of section to make the toc readable
+\renewcommand{\thesubsection}{\arabic{subsection}}
+\renewcommand{\isamarkupheader}[1]{\newpage\markright{Theory~\isabellecontext}\subsection{#1}}
+\renewcommand{\isamarkupsection}[1]{\subsubsection{#1}}
+
+%remove spaces from the isabelle environment (trivlist makes them too large)
+\renewenvironment{isabelle}
+{\begin{isabellebody}}
+{\end{isabellebody}}
+
+\begin{document}
+
+\title{Java Source and Bytecode Formalizations in Isabelle: Bali\\
+  {\large -- VerifiCard Project Deliverables -- }}
+\author{Gerwin Klein \and Tobias Nipkow \and David von Oheimb \and
+  Leonor Prensa Nieto \and Norbert Schirmer \and Martin Strecker}
+\maketitle
+
+\tableofcontents
+
+\parindent 0pt\parskip 0.5ex
+\chapter{Overview}
+These theories, called Bali,  model and analyse different aspects of the 
+JavaCard \textbf{source language}. 
+The basis is an abstract model of the JavaCard source language. 
+On it, a type system, an operational semantics and an axiomatic semantics 
+(Hoare logic) are built. The execution of a wellformed program (with respect to
+the type system) according to the operational semantics is proved to be 
+typesafe. The axiomatic semantics is proved to be sound and relative complete 
+with respect to the operational semantics.
+
+We have modelled large parts of the original JavaCard source language. It models
+features such as:
+\begin{itemize}
+\item The basic ``primitive types'' of Java 
+\item Classes and related concepts 
+\item Class fields and methods
+\item Instance fields and methods
+\item Interfaces and related concepts 
+\item Arrays
+\item Static initialisation
+\item Static overloading of fields and methods
+\item Inheritance, overriding and hiding of methods, dynamic binding
+\item All cases of abrupt termination
+      \begin{itemize}
+        \item Exception throwing and handling
+        \item \texttt{break}, \texttt{continue} and \texttt{return} 
+      \end{itemize}
+\item Packages
+\item Access Modifiers (\texttt{private}, \texttt{protected}, \texttt{public})
+\end{itemize}
+
+The following features are missing in Bali wrt.{} JavaCard:
+\begin{itemize}
+\item Some primitive types (\texttt{byte, short})
+\item Most numeric operations, syntactic variants of statements
+  (\texttt{do}-loop, \texttt{for}-loop)
+\item A ``definite assignment'' check
+\end{itemize}
+
+In addition, features are missing that are not part of the JavaCard
+language, such as multithreading and garbage collection. No attempt
+has been made to model peculiarities of JavaCard such as the applet
+firewall or the transaction mechanism.
+
+
+Overview of the theories:
+\begin{description}
+\item [Basis.thy] 
+Some basic definitions and settings not specific to JavaCard but missing in HOL.
+
+\item [Table.thy]
+Definition and some properties of a lookup table to map various names 
+(like class names or method names) to some content (like classes or methods).
+
+\item[Name.thy]
+Definition of various names (class names, variable names, package names,...)
+
+\item[Value.thy]
+JavaCard expression values (Boolean, Integer, Addresses,...)
+
+\item[Type.thy]
+JavaCard types. Primitive types (Boolean, Integer,...) and reference types 
+(Classes, Interfaces, Arrays,...)
+
+\item[Term.thy]
+JavaCard terms. Variables, expressions and statements.
+
+\item[Decl.thy]
+Class, interface and program declarations. Recursion operators for the
+class and the interface hierarchy. 
+
+\item[TypeRel.thy]
+Various relations on types like the subclass-, subinterface-, widening-, 
+narrowing- and casting-relation.
+
+\item[DeclConcepts.thy]
+Advanced concepts on the class and interface hierarchy like inheritance, 
+overriding, hiding, accessibility of types and members according to the access 
+modifiers, method lookup.
+
+\item[WellType.thy]
+Typesystem on the JavaCard term level.
+
+\item[WellForm.thy]
+Typesystem on the JavaCard class, interface and program level.
+
+\item[State.thy]
+The program state (like object store) for the execution of JavaCard.
+Abrupt completion (exceptions, break, continue, return) is modelled as flag
+inside the state.
+
+\item[Eval.thy]
+Operational (big step) semantics for JavaCard.
+
+\item[Example.thy]
+An concrete example of a JavaCard program to validate the typesystem and the
+operational semantics.
+
+\item[Conform.thy]
+Conformance predicate for states. When does an execution state conform to the
+static types of the program given by the typesystem.
+
+\item[TypeSafe.thy]
+Typesafety proof of the execution of JavaCard. ''Welltyped programs don't go
+wrong'' or more technical: The execution of a welltyped JavaCard program 
+preserves the conformance of execution states.
+
+\item[Evaln.thy]
+Copy of the operational semantics given in Eval.thy expanded with an annotation
+for the maximal recursive depth. The semantics is not altered. The annotation
+is needed for the soundness proof of the axiomatic semantics.
+
+\item[AxSem.thy]
+An axiomatic semantics (Hoare logic) for JavaCard.
+
+\item[AxSound.thy]
+The soundness proof of the axiomatic semantics with respect to the operational
+semantics.
+
+\item[AxCompl.thy]
+The proof of (relative) completeness of the axiomatic semantics with respect
+to the operational semantics. 
+
+\item[AxExample.thy]
+An concrete example of the axiomatic semantics at work, applied to prove 
+some properties of the JavaCard example given in Example.thy.
+\end{description}
+
+% include generated text of all theories
+\chapter{Basis.thy}
+\input{../generated/Basis.tex}
+\chapter{Table.thy}
+\input{../generated/Table.tex}    
+
+\chapter{Name.thy}
+\input{../generated/Name.tex}
+\chapter{Value.thy}     
+\input{../generated/Value.tex}
+\chapter{Type.thy}
+\input{../generated/Type.tex}      
+\chapter{Term.thy}
+\input{../generated/Term.tex}     
+\chapter{Decl.thy}
+\input{../generated/Decl.tex}          
+\chapter{TypeRel.thy}
+\input{../generated/TypeRel.tex}   
+\chapter{DeclConcepts.thy}
+\input{../generated/DeclConcepts.tex}  
+
+\chapter{WellType.thy}
+\input{../generated/WellType.tex}
+\chapter{WellForm.thy}
+\input{../generated/WellForm.tex}
+
+\chapter{State.thy}
+\input{../generated/State.tex}    
+\chapter{Eval.thy}
+\input{../generated/Eval.tex}          
+
+\chapter{Example.thy}
+\input{../generated/Example.tex}  
+
+\chapter{Conform.thy}
+\input{../generated/Conform.tex}       
+\chapter{TypeSafe.thy}
+\input{../generated/TypeSafe.tex}
+
+\chapter{Evaln.thy}
+\input{../generated/Evaln.tex}         
+\chapter{AxSem.thy}
+\input{../generated/AxSem.tex}      
+\chapter{AxSound.thy}
+\input{../generated/AxSound.tex}    
+\chapter{AxCompl.thy}
+\input{../generated/AxCompl.tex}    
+\chapter{AxExample.thy}
+\input{../generated/AxExample.tex}  
+
+
+
+
+
+
+
+%\bibliographystyle{abbrv}
+%\bibliography{root}
+
+\end{document}