Added check for field/method access to operational semantics and proved the acesses valid.
authorschirmer
Fri, 22 Feb 2002 11:26:44 +0100
changeset 12925 99131847fb93
parent 12924 02eb40cde931
child 12926 cd0dd6e0bf5c
Added check for field/method access to operational semantics and proved the acesses valid.
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/State.thy
src/HOL/Bali/Table.thy
src/HOL/Bali/Term.thy
src/HOL/Bali/Trans.thy
src/HOL/Bali/TypeSafe.thy
src/HOL/Bali/WellForm.thy
src/HOL/Bali/WellType.thy
--- a/src/HOL/Bali/AxCompl.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/AxCompl.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -16,6 +16,9 @@
 \item proof structured by Most General Formulas (-> Thomas Kleymann)
 \end{itemize}
 *}
+
+
+
 section "set of not yet initialzed classes"
 
 constdefs
@@ -155,11 +158,14 @@
   "{=: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>}"
+
+lemma MGF_valid: "wf_prog G \<Longrightarrow> G,{}\<Turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
 apply (unfold MGF_def)
-apply (force dest!: evaln_eval simp add: ax_valids_def triple_valid_def2)
+apply (simp add:  ax_valids_def triple_valid_def2)
+apply (auto elim: evaln_eval)
 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
@@ -228,6 +234,25 @@
 apply (auto elim: conseq12 simp add: MGFn_def MGF_def)
 done
 
+lemma MGFn_free_wt_NormalConformI: 
+"(\<forall> T L C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T 
+  \<longrightarrow> G,(A::state triple set)
+      \<turnstile>{Normal((\<lambda>Y' s' s. s'=s \<and> normal s) \<and>. G\<turnstile>init\<le>n) \<and>. (\<lambda> s. s\<Colon>\<preceq>(G, L))}
+      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 (rule MGFn_NormalI)
+apply (rule ax_no_hazard)
+apply (rule ax_escape)
+apply (intro strip)
+apply (simp only: type_ok_def peek_and_def)
+apply (erule conjE)+
+apply (erule exE,erule exE, erule exE,erule conjE,drule (1) mp)
+apply (drule spec,drule spec, drule spec, drule (1) mp)
+apply (erule conseq12)
+apply blast
+done
+
 
 section "main lemmas"
 
@@ -250,6 +275,9 @@
     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 wt_conf_prepare_tac = EVERY'[
+    rtac (thm "MGFn_free_wt_NormalConformI"),
+    clarsimp_tac (claset() addSEs (thms "wt_elim_cases"), simpset())]
 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]
@@ -290,22 +318,64 @@
 done
 lemmas MGFn_InitD = MGFn_Init [THEN MGFnD, THEN ax_NormalD]
 
+text {* For @{text MGFn_Call} we need the wellformedness of the program to
+switch from the evaln-semantics to the eval-semantics *}
 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")
+  G,A\<turnstile>{=:n} In1l e\<succ> {G\<rightarrow>}; G,A\<turnstile>{=:n} In3 ps\<succ> {G\<rightarrow>};wf_prog G\<rbrakk> \<Longrightarrow>  
+  G,A\<turnstile>{=:n} In1l ({accC,statT,mode}e\<cdot>mn({pTs'}ps))\<succ> {G\<rightarrow>}"
+apply (tactic "wt_conf_prepare_tac 1")
+apply (rule_tac  
+  Q="(\<lambda>Y s1 (x,s) . x = None \<and> 
+        (\<exists>a. G\<turnstile>Norm s \<midarrow>e-\<succ>a\<rightarrow> s1 \<and> (normal s1 \<longrightarrow> G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT)
+         \<and> Y = In1 a)) 
+    \<and>. G\<turnstile>init\<le>n \<and>. (\<lambda> s. s\<Colon>\<preceq>(G, L))" and 
+ 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> 
+                       (normal s1 \<longrightarrow> G, store s1\<turnstile>a'\<Colon>\<preceq>RefT statT)\<and> 
+                       Y = In3 pvs \<and> G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>pvs\<rightarrow> (x2,s2))) 
+            \<and>. G\<turnstile>init\<le>n \<and>. (\<lambda> s. s\<Colon>\<preceq>(G, L))" in ax_derivs.Call)
+apply   (tactic "forw_hyp_tac 1")
+apply   (tactic "clarsimp_tac eval_css 1")
+apply   (frule (3) eval_type_sound)
+apply   force
+
+apply   safe
+apply   (tactic "forw_hyp_tac 1")
+apply   (tactic "clarsimp_tac eval_css 1")
+apply   (frule (3) eval_type_sound)
+apply     (rule conjI)
+apply       (rule exI,rule conjI)
+apply         (assumption)
+
+apply         (rule conjI)
+apply           simp
+apply           assumption
+apply      blast
+
 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))+
+apply   (rule HOL.refl)+
+apply   (subgoal_tac "check_method_access G C statT (invmode m e)
+             \<lparr>name = mn, parTs = pTs'\<rparr> a
+             (init_lvars G
+               (invocation_declclass G (invmode m e) (snd (ab, ba)) a statT
+                 \<lparr>name = mn, parTs = pTs'\<rparr>)
+               \<lparr>name = mn, parTs = pTs'\<rparr> (invmode m e) a vs
+               (ab,
+                ba)) = (init_lvars G
+               (invocation_declclass G (invmode m e) (snd (ab, ba)) a statT
+                 \<lparr>name = mn, parTs = pTs'\<rparr>)
+               \<lparr>name = mn, parTs = pTs'\<rparr> (invmode m e) a vs
+               (ab,
+                ba))")
+apply    simp
+defer 
+apply simp
+apply (erule (3) error_free_call_access) (* now showing the subgoal *)
+apply auto
 done
 
 lemma MGF_altern: "G,A\<turnstile>{Normal (\<doteq> \<and>. p)} t\<succ> {G\<rightarrow>} =  
@@ -356,9 +426,54 @@
   simpset() addsimps [split_paired_all] addsimprocs [eval_stmt_proc]) 1*})+)
 done
 
+text {* For @{text MGFn_FVar} we need the wellformedness of the program to
+switch from the evaln-semantics to the eval-semantics *}
+lemma MGFn_FVar:
+ "\<lbrakk>G,A\<turnstile>{=:n} In1r (Init statDeclC)\<succ> {G\<rightarrow>}; G,A\<turnstile>{=:n} In1l e\<succ> {G\<rightarrow>}; wf_prog G\<rbrakk>
+   \<Longrightarrow> G,(A\<Colon>state triple set)\<turnstile>{=:n} In2 ({accC,statDeclC,stat}e..fn)\<succ> {G\<rightarrow>}"
+apply (tactic "wt_conf_prepare_tac 1")
+apply (rule_tac  
+  Q="(\<lambda>Y s1 (x,s) . x = None \<and> 
+        (G\<turnstile>Norm s \<midarrow>Init statDeclC\<rightarrow> s1 
+         )) \<and>. G\<turnstile>init\<le>n \<and>. (\<lambda> s. s\<Colon>\<preceq>(G, L))"  
+ in ax_derivs.FVar)
+apply (tactic "forw_hyp_tac 1")
+apply (tactic "clarsimp_tac eval_css 1")
+apply (subgoal_tac "is_class G statDeclC")
+apply   (force dest: eval_type_sound)
+apply   (force dest: ty_expr_is_type [THEN type_is_class] 
+                      accfield_fields [THEN fields_declC])
+apply (tactic "forw_hyp_tac 1")
+apply (tactic "clarsimp_tac eval_css 1")
+apply (subgoal_tac "(\<exists> v' s2' s3.   
+        ( fvar statDeclC (is_static f) fn v (aa, ba) = (v',s2') ) \<and>
+            (s3  = check_field_access G C statDeclC fn (is_static f) v s2') \<and>
+            (s3 = s2'))")
+apply   (erule exE)+
+apply   (erule conjE)+
+apply   (erule (1) eval.FVar)
+apply     simp
+apply     simp
+
+apply   (case_tac "fvar statDeclC (is_static f) fn v (aa, ba)")
+apply   (rule exI)+
+apply   (rule context_conjI)
+apply      force
+
+apply   (rule context_conjI)
+apply     simp
+
+apply     (erule (3) error_free_field_access)
+apply       (auto dest: eval_type_sound)
+done
+
+text {* For @{text MGFn_lemma} we need the wellformedness of the program to
+switch from the evaln-semantics to the eval-semantics cf. @{text MGFn_call}, 
+@{text MGFn_FVar}*}
 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>}"
+ "\<lbrakk>\<forall>n C sig. G,(A::state triple set)\<turnstile>{=:n} In1l (Methd C sig)\<succ> {G\<rightarrow>};
+   wf_prog G\<rbrakk> 
+  \<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)
@@ -371,18 +486,15 @@
 apply (rule var_expr_stmt.induct)
 (* 28 subgoals *)
 prefer 14 apply fast (* Methd *)
-prefer 13 apply (erule (2) MGFn_Call)
+prefer 13 apply (erule (3) MGFn_Call)
+prefer 2  apply (drule MGFn_Init,erule (2) MGFn_FVar)
 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 (erule_tac [23] MGFn_Init)
+prefer 18 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")
@@ -480,14 +592,16 @@
 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>}"
+lemma MGF_asm: 
+"\<lbrakk>\<forall>C sig. is_methd G C sig \<longrightarrow> G,A\<turnstile>{\<doteq>} In1l (Methd C sig)\<succ> {G\<rightarrow>}; wf_prog G\<rbrakk>
+ \<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)
+apply assumption (* wf_prog G *)
 done
 
 declare splitI2 [intro!]
@@ -563,7 +677,7 @@
 apply (erule MethdI)
 done
 
-lemma MGF_deriv: "ws_prog G \<Longrightarrow> G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
+lemma MGF_deriv: "wf_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 
@@ -571,11 +685,9 @@
 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  (erule_tac [2] finite_is_methd [OF wf_ws_prog])
 apply (rule MGF_asm [THEN MGFNormalD])
-apply clarify
-apply (rule MGFNormalI)
-apply force
+apply (auto intro: MGFNormalI)
 done
 
 
@@ -598,10 +710,10 @@
 apply (erule eval_Methd)
 done
 
-lemma MGF_simult_Methd: "ws_prog G \<Longrightarrow> 
+lemma MGF_simult_Methd: "wf_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 (frule finite_is_methd [OF wf_ws_prog])
 apply (rule MGF_simult_Methd_lemma)
 apply  assumption
 apply (erule ax_finite_pointwise)
@@ -610,36 +722,42 @@
 apply  blast
 apply clarsimp
 apply (rule MGF_asm [THEN MGFNormalD])
-apply clarify
-apply (rule MGFNormalI)
-apply force
+apply   (auto intro: MGFNormalI)
 done
 
-lemma MGF_deriv: "ws_prog G \<Longrightarrow> G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
+lemma MGF_deriv: "wf_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
+apply auto
 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}"
+lemma eval_to_evaln: "\<lbrakk>G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y', s');type_ok G t s; wf_prog G\<rbrakk>
+ \<Longrightarrow>   \<exists>n. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s')"
+apply (cases "normal s")
+apply   (force simp add: type_ok_def intro: eval_evaln)
+apply   (force intro: evaln.Abrupt)
+done
+
+lemma MGF_complete: 
+ "\<lbrakk>G,{}\<Turnstile>{P} t\<succ> {Q}; G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}; wf_prog G\<rbrakk> 
+  \<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)
+apply (blast dest: eval_to_evaln)
 done
 
-theorem ax_complete: "ws_prog G \<Longrightarrow>  
+theorem ax_complete: "wf_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)
+apply (erule (1) MGF_deriv)
 done
 
 end
--- a/src/HOL/Bali/AxExample.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/AxExample.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -3,6 +3,7 @@
     Author:     David von Oheimb
     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 *)
+
 header {* Example of a proof based on the Bali axiomatic semantics *}
 
 theory AxExample = AxSem + Example:
--- a/src/HOL/Bali/AxSem.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/AxSem.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -39,8 +39,6 @@
 \end{itemize}
 *}
 
-
-
 types  res = vals (* result entry *)
 syntax
   Val  :: "val      \<Rightarrow> res"
@@ -514,7 +512,7 @@
 
   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}"
+                                 G,A\<turnstile>{Normal P} {accC,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>
@@ -560,7 +558,7 @@
       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}"
+         G,A\<turnstile>{Normal P} {accC,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}"
@@ -1044,7 +1042,7 @@
                      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}"
+   \<Longrightarrow> G,A\<turnstile>{Normal P} {accC,statT,IntVir}e\<cdot>mn({pTs}args)-\<succ> {S}"
 apply (erule ax_derivs.Call)
 apply  safe
 apply  (erule spec)
@@ -1062,7 +1060,7 @@
   \<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}"
+\<rbrakk>  \<Longrightarrow>  G,A\<turnstile>{Normal P} {accC,statT,Static}e\<cdot>mn({pTs}args)-\<succ> {S}"
 apply (erule ax_derivs.Call)
 apply  safe
 apply  (erule spec)
--- a/src/HOL/Bali/AxSound.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/AxSound.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -185,18 +185,18 @@
    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}}"
+  G,A|\<Turnstile>\<Colon>{ {Normal P} {accC,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 (3) evaln_eval, drule (3) eval_ts)
+apply (drule (3) 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")
+apply (case_tac "if is_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 
@@ -246,6 +246,19 @@
              del: impCE simp add: init_lvars_def2)
 done
 
+corollary evaln_type_sound:
+      (assumes evaln: "G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)" and
+                  wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T" and
+             conf_s0: "s0\<Colon>\<preceq>(G,L)" and
+                  wf: "wf_prog G"                         
+      ) "s1\<Colon>\<preceq>(G,L) \<and>  (normal s1 \<longrightarrow> G,L,store s1\<turnstile>t\<succ>v\<Colon>\<preceq>T) \<and> 
+         (error_free s0 = error_free s1)"
+proof -
+  from evaln wt conf_s0 wf
+  show ?thesis
+    by (blast dest: evaln_eval eval_type_sound)
+qed
+
 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}};  
@@ -260,7 +273,7 @@
 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 (frule (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
@@ -272,9 +285,25 @@
 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 (erule impE,rule impI,rule exI,erule wf_cdecl_wt_init)
 apply clarsimp
-apply (erule (1) conforms_return, force dest: evaln_eval eval_gext')
+apply (erule (1) conforms_return)
+apply (frule wf_cdecl_wt_init)
+apply (subgoal_tac "(a, set_locals empty b)\<Colon>\<preceq>(G, empty)")
+apply   (frule (3) evaln_eval)
+apply   (drule eval_gext') 
+apply   force
+
+        (* refer to case Init in eval_type_sound proof, to see whats going on*)
+apply   (subgoal_tac "(a,b)\<Colon>\<preceq>(G, L)")
+apply     (blast intro: conforms_set_locals)
+
+apply     (drule evaln_type_sound)
+apply       (cases "C=Object") 
+apply         force 
+apply         (force dest: wf_cdecl_supD is_acc_classD)
+apply     (erule (4) conforms_init_class_obj)
+apply     blast
 done
 
 lemma all_conjunct2: "\<forall>l. P' l \<and> P l \<Longrightarrow> \<forall>l. P l"
@@ -301,10 +330,10 @@
 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    (tactic "smp_tac 3 1", clarify, tactic "smp_tac 1 1",
+          case_tac"fst s",clarsimp,erule (3) evaln_type_sound [THEN conjunct1],
+          clarsimp) (* conseq *)
+apply   (simp (no_asm_use) add: type_ok_def,fast)(* hazard *)
 apply  force (* Abrupt *)
 
 (* 25 subgoals *)
@@ -359,17 +388,17 @@
 apply (case_tac "aa")
 prefer 2
 apply  clarsimp
-apply (drule evaln_eval)+
-apply (frule (3) eval_ts)
+apply (drule (3) evaln_type_sound)
+apply (drule (3) evaln_eval)
+apply (frule (3) eval_type_sound)
 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) 
@@ -391,7 +420,7 @@
 apply (force split add: split_if)
 
 (* Throw *)
-apply (drule evaln_eval, drule (3) eval_ts)
+apply (drule (3) evaln_type_sound)
 apply clarsimp
 apply (drule (3) Throw_lemma)
 apply clarsimp
@@ -416,7 +445,7 @@
 apply (tactic "sound_forw_hyp_tac 1")
 apply (case_tac "x1", force)
 apply clarsimp
-apply (drule evaln_eval, drule (4) Fin_lemma)
+apply (drule (3) evaln_eval, drule (4) Fin_lemma)
 done
 
 
--- a/src/HOL/Bali/Basis.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/Basis.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -51,7 +51,7 @@
 done
 
 syntax
-  "3" :: nat   ("3")
+  "3" :: nat   ("3") 
   "4" :: nat   ("4")
 translations
  "3" == "Suc 2"
@@ -75,7 +75,7 @@
 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^+"
+"\<lbrakk>(a,b)\<in>r^*; a\<noteq>b\<rbrakk> \<Longrightarrow> (a,b)\<in>r^+" 
 apply (drule rtranclD)
 apply auto
 done
--- a/src/HOL/Bali/Conform.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/Conform.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -23,11 +23,19 @@
 
 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"
 
+text {* For the the proof of type soundness we will need the 
+property that during execution, objects are not lost and moreover retain the 
+values of their tags. So the object store grows conservatively. Note that if 
+we considered garbage collection, we would have to restrict this property to 
+accessible objects.
+*}
+
 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"
@@ -348,15 +356,21 @@
   "((Some (Xcpt (Std xn)), s)\<Colon>\<preceq>(G, L)) = (Norm s\<Colon>\<preceq>(G, L))"
 by (auto simp: conforms_def)
 
+lemma conforms_Err [iff]:
+   "((Some (Error e), 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_error_if [iff]: 
+  "((error_if c err x, s)\<Colon>\<preceq>(G, L)) = ((x, s)\<Colon>\<preceq>(G, L))"
+by (auto simp: abrupt_if_def split: split_if)
 
 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)
@@ -436,6 +450,11 @@
             elim!: conforms_XcptLocD simp add: oconf_def)
 done
 
+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 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)
@@ -444,4 +463,5 @@
 apply (force dest: conforms_globsD)+
 done
 
+
 end
--- a/src/HOL/Bali/Decl.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/Decl.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Bali/Decl.thy
     ID:         $Id$
-    Author:     David von Oheimb
+    Author:     David von Oheimb and Norbert Schirmer
     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 *)
 header {* Field, method, interface, and class declarations, whole Java programs
--- a/src/HOL/Bali/DeclConcepts.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/DeclConcepts.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -1,3 +1,8 @@
+(*  Title:      HOL/Bali/DeclConcepts.thy
+    ID:         $Id$
+    Author:     Norbert Schirmer
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+*)
 header {* Advanced concepts on Java declarations like overriding, inheritance,
 dynamic method lookup*}
 
@@ -57,7 +62,7 @@
 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"
+ "is_acc_type  G P T \<Longrightarrow> is_type G T  \<and> G\<turnstile>T accessible_in P"
 by (simp add: is_acc_type_def)
 
 lemma is_acc_reftypeD:
@@ -815,7 +820,8 @@
      | 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) 
+                    (G\<turnstile>accclass \<prec>\<^sub>C declclass membr 
+                     \<and> (G\<turnstile>class \<preceq>\<^sub>C accclass \<or> is_static membr)) 
      | Public    \<Rightarrow> True)"
 text {*
 The subcondition of the @{term "Protected"} case: 
@@ -881,12 +887,12 @@
  \<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;
+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;
+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; 
@@ -934,16 +940,12 @@
 "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;
+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;
+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;
@@ -955,10 +957,22 @@
  \<Longrightarrow> G\<turnstile>m member_of C \<and> G\<turnstile>(Class C) accessible_in (pid S)"
 by (auto elim: accessible_fromR.induct)
 
+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 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 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 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)"
@@ -1115,86 +1129,30 @@
   qed
 qed
 
+lemma member_in_declC: "G\<turnstile>m member_in C\<Longrightarrow> G\<turnstile>m member_in (declclass m)"
+proof -
+  assume member_in_C:  "G\<turnstile>m member_in C"
+  from member_in_C
+  obtain provC where
+    subclseq_C_provC: "G\<turnstile> C \<preceq>\<^sub>C provC" and
+     member_of_provC: "G \<turnstile> m member_of provC"
+    by (auto simp add: member_in_def)
+  from member_of_provC
+  have "G \<turnstile> m member_of declclass m"
+    by (rule member_of_member_of_declC)
+  moreover
+  from member_in_C
+  have "G\<turnstile>C \<preceq>\<^sub>C declclass m"
+    by (rule member_in_class_relation)
+  ultimately
+  show ?thesis
+    by (auto simp add: member_in_def)
+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)
@@ -1209,6 +1167,34 @@
    (auto simp add: permits_acc_def
             intro: subclseq_trans) 
 
+lemma permits_acc_static_declC:
+ "\<lbrakk>G\<turnstile>m in C permits_acc_to accC; G\<turnstile>m member_in C; is_static m
+  \<rbrakk> \<Longrightarrow> G\<turnstile>m in (declclass m) permits_acc_to accC"
+by (cases "accmodi m") (auto simp add: permits_acc_def)
+
+lemma dyn_accessible_from_static_declC: 
+  (assumes  acc_C: "G\<turnstile>m in C dyn_accessible_from accC" and
+           static: "is_static m"
+  ) "G\<turnstile>m in (declclass m) dyn_accessible_from accC"
+proof -
+  from acc_C static
+  show "G\<turnstile>m in (declclass m) dyn_accessible_from accC"
+  proof (induct)
+    case (Immediate C m)
+    then show ?case 
+      by (auto intro!: dyn_accessible_fromR.Immediate
+                 dest: member_in_declC permits_acc_static_declC) 
+  next 
+    case (Overriding declCNew C m new old sup)
+    then have "\<not> is_static m"
+      by (auto dest: overrides_commonD)
+    moreover
+    assume "is_static m"
+    ultimately show ?case 
+      by contradiction
+  qed
+qed
+
 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>
@@ -1445,7 +1431,7 @@
   from stat_acc is_field subclseq 
   show ?thesis
     by (auto dest: accessible_fieldD 
-            intro: dyn_accessible_fromR.immediate 
+            intro: dyn_accessible_fromR.Immediate 
                    member_inI 
                    permits_acc_inheritance)
 qed
@@ -1463,15 +1449,15 @@
   from stat_acc
   show ?thesis
   proof (cases)
-    case immediate
+    case Immediate
     with member_dynC member_statC subclseq dynC_acc
     show ?thesis
-      by (auto intro: accessible_fromR.immediate permits_acc_inheritance)
+      by (auto intro: accessible_fromR.Immediate permits_acc_inheritance)
   next
-    case overriding
+    case Overriding
     with member_dynC subclseq dynC_acc
     show ?thesis
-      by (auto intro: accessible_fromR.overriding rtrancl_trancl_trancl)
+      by (auto intro: accessible_fromR.Overriding rtrancl_trancl_trancl)
   qed
 qed
 
--- a/src/HOL/Bali/Eval.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/Eval.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -370,7 +370,14 @@
 	          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)"
-
+(*
+ "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'); 
@@ -412,7 +419,32 @@
 apply (simp (no_asm) add: Let_def split_beta)
 done
 
+constdefs
+check_field_access::
+"prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> vname \<Rightarrow> bool \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state"
+"check_field_access G accC statDeclC fn stat a' s
+ \<equiv> let oref = if stat then Stat statDeclC
+                      else Heap (the_Addr a');
+       dynC = case oref of
+                   Heap a \<Rightarrow> obj_class (the (globs (store s) oref))
+                 | Stat C \<Rightarrow> C;
+       f    = (the (table_of (DeclConcepts.fields G dynC) (fn,statDeclC)))
+   in abupd 
+        (error_if (\<not> G\<turnstile>Field fn (statDeclC,f) in dynC dyn_accessible_from accC)
+                  AccessViolation)
+        s"
 
+constdefs
+check_method_access:: 
+  "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> inv_mode \<Rightarrow>  sig \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state"
+"check_method_access G accC statT mode sig  a' s
+ \<equiv> let invC = invocation_class mode (store s) a' statT;
+       dynM = the (dynlookup G statT invC sig)
+   in abupd 
+        (error_if (\<not> G\<turnstile>Methd sig dynM in invC dyn_accessible_from accC)
+                  AccessViolation)
+        s"
+       
 section "evaluation judgments"
 
 consts
@@ -552,7 +584,27 @@
 	       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"
-
+   (* This class initialisation rule is a little bit inaccurate. Look at the
+      exact sequence:
+      1. The current class object (the static fields) are initialised
+         (init_class_obj)
+      2. Then the superclasses are initialised
+      3. The static initialiser of the current class is invoked
+      More precisely we should expect another ordering, namely 2 1 3.
+      But we can't just naively toggle 1 and 2. By calling init_class_obj 
+      before initialising the superclasses we also implicitly record that
+      we have started to initialise the current class (by setting an 
+      value for the class object). This becomes 
+      crucial for the completeness proof of the axiomatic semantics 
+      (AxCompl.thy). Static initialisation requires an induction on the number 
+      of classes not yet initialised (or to be more precise, classes where the
+      initialisation has not yet begun). 
+      So we could first assign a dummy value to the class before
+      superclass initialisation and afterwards set the correct values.
+      But as long as we don't take memory overflow into account 
+      when allocating class objects, and don't model definite assignment in
+      the static initialisers, we can leave things as they are for convenience. 
+   *)
 (* evaluation of expressions *)
 
   (* cf. 15.8.1, 12.4.1 *)
@@ -602,10 +654,15 @@
   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>
+    s3=init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs s2;
+    s3' = check_method_access G accC statT mode \<lparr>name=mn,parTs=pTs\<rparr> a' s3;
+    G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<rightarrow> s4\<rbrakk>
    \<Longrightarrow>
-       G\<turnstile>Norm s0 \<midarrow>{statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<rightarrow> (restore_lvars s2 s3)"
+       G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<rightarrow> (restore_lvars s2 s4)"
+(* The accessibility check is after init_lvars, to keep it simple. Init_lvars 
+   already tests for the absence of a null-pointer reference in case of an
+   instance method invocation
+*)
 
   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"
@@ -620,9 +677,14 @@
   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'"
+  FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2;
+	  (v,s2') = fvar statDeclC stat fn a s2;
+          s3 = check_field_access G accC statDeclC fn stat a s2' \<rbrakk> \<Longrightarrow>
+	  G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v\<rightarrow> s3"
+ (* The accessibility check is after fvar, to keep it simple. Fvar already
+    tests for the absence of a null-pointer reference in case of an instance
+    field
+  *)
 
   (* cf. 15.12.1, 15.25.1 *)
   AVar:	"\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2;
@@ -688,35 +750,35 @@
 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'"
+        "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  ({accC,statDeclC,stat}e..fn)   \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                      \<succ>\<rightarrow> vs'"
+	"G\<turnstile>Norm s \<midarrow>In1l ({accC,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 {*
@@ -851,12 +913,14 @@
 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'"
+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>;
+     s3 = init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' pvs s2;
+     s3' = check_method_access G accC statT mode \<lparr>name=mn,parTs=pTs\<rparr> a' s3;
+     G\<turnstile>s3'\<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> v\<rightarrow> s4; 
+       s4' = restore_lvars s2 s4\<rbrakk> \<Longrightarrow>  
+       G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn({pTs}ps)-\<succ>v\<rightarrow> s4'"
 apply (drule eval.Call, assumption)
 apply (rule HOL.refl)
 apply simp+
--- a/src/HOL/Bali/Evaln.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/Evaln.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -1,17 +1,19 @@
 (*  Title:      HOL/Bali/Evaln.thy
     ID:         $Id$
-    Author:     David von Oheimb
+    Author:     David von Oheimb and Norbert Schirmer
     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 *)
 header {* Operational evaluation (big-step) semantics of Java expressions and 
           statements
 *}
 
-theory Evaln = Eval:
+theory Evaln = Eval + TypeSafe:
 
 text {*
-Variant of eval relation with counter for bounded recursive depth
-Evaln could completely replace Eval.
+Variant of eval relation with counter for bounded recursive depth.
+Evaln omits the technical accessibility tests @{term check_field_access}
+and @{term check_method_access}, since we proved the absence of errors for
+wellformed programs.
 *}
 
 consts
@@ -70,9 +72,9 @@
 
   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'"
+  FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s2;
+	  (v,s2') = fvar statDeclC stat fn a' s2\<rbrakk> \<Longrightarrow>
+	  G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,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>
@@ -119,7 +121,8 @@
     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)"
+   \<Longrightarrow> 
+    G\<turnstile>Norm s0 \<midarrow>{accC,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"
@@ -187,56 +190,6 @@
 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]
@@ -268,9 +221,9 @@
 	"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  ({accC,statDeclC,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>In1l ({accC,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]
@@ -370,4 +323,1401 @@
 apply auto
 done
 
+lemma evaln_eval:  
+ (assumes evaln: "G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)" and
+             wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T" and  
+        conf_s0: "s0\<Colon>\<preceq>(G, L)" and
+             wf: "wf_prog G" 
+       
+ )  "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
+proof -
+  from evaln 
+  show "\<And> L accC T. \<lbrakk>s0\<Colon>\<preceq>(G, L);\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T\<rbrakk>
+                    \<Longrightarrow> G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
+       (is "PROP ?EqEval s0 s1 t v")
+  proof (induct)
+    case Abrupt
+    show ?case by (rule eval.Abrupt)
+  next
+    case LVar
+    show ?case by (rule eval.LVar)
+  next
+    case (FVar a accC' e fn n s0 s1 s2 s2' stat statDeclC v L accC T)
+    have eval_initn: "G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<midarrow>n\<rightarrow> s1" .
+    have eval_en: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s2" .
+    have hyp_init: "PROP ?EqEval (Norm s0) s1 (In1r (Init statDeclC)) \<diamondsuit>" .
+    have hyp_e: "PROP ?EqEval s1 s2 (In1l e) (In1 a)" .
+    have fvar: "(v, s2') = fvar statDeclC stat fn a s2" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have wt: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>In2 ({accC',statDeclC,stat}e..fn)\<Colon>T" .
+    then obtain statC f where
+                wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-Class statC" and
+            accfield: "accfield G accC statC fn = Some (statDeclC,f)" and
+                stat: "stat=is_static f" and
+               accC': "accC'=accC" and
+	           T: "T=(Inl (type f))"
+       by (rule wt_elim_cases) (auto simp add: member_is_static_simp)
+    from wf wt_e 
+    have iscls_statC: "is_class G statC"
+      by (auto dest: ty_expr_is_type type_is_class)
+    with wf accfield 
+    have iscls_statDeclC: "is_class G statDeclC"
+      by (auto dest!: accfield_fields dest: fields_declC)
+    then 
+    have wt_init: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
+      by simp
+    from conf_s0 wt_init
+    have eval_init: "G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1"
+      by (rule hyp_init)
+    with wt_init conf_s0 wf 
+    have conf_s1: "s1\<Colon>\<preceq>(G, L)"
+      by (blast dest: exec_ts)
+    with hyp_e wt_e
+    have eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2"
+      by blast
+    with wf conf_s1 wt_e
+    obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+            conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+      by (auto dest!: eval_type_sound)
+    obtain s3 where
+      check: "s3 = check_field_access G accC statDeclC fn stat a s2'"
+      by simp
+    from accfield wt_e eval_init eval_e conf_s2 conf_a fvar stat check  wf
+    have eq_s3_s2': "s3=s2'"  
+      by (auto dest!: error_free_field_access)
+    with eval_init eval_e fvar check accC'
+    show "G\<turnstile>Norm s0 \<midarrow>{accC',statDeclC,stat}e..fn=\<succ>v\<rightarrow> s2'"
+      by (auto intro: eval.FVar)
+  next
+    case AVar
+    with wf show ?case
+      apply -
+      apply (erule wt_elim_cases)
+      apply (blast intro!: eval.AVar dest: eval_type_sound)
+      done
+  next
+    case NewC
+    with wf show ?case
+      apply - 
+      apply (erule wt_elim_cases)
+      apply (blast intro!: eval.NewC dest: eval_type_sound is_acc_classD)
+      done
+  next
+    case (NewA T a e i n s0 s1 s2 s3 L accC Ta) 
+    have hyp_init: "PROP ?EqEval (Norm s0) s1 (In1r (init_comp_ty T)) \<diamondsuit>" .
+    have hyp_size: "PROP ?EqEval s1 s2 (In1l e) (In1 i)" .
+    have "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (New T[e])\<Colon>Ta" .
+    then obtain
+       wt_init: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and
+       wt_size: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-PrimT Integer"
+      by (rule wt_elim_cases) (auto intro: wt_init_comp_ty dest: is_acc_typeD)
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    from this wt_init 
+    have eval_init: "G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1"
+      by (rule hyp_init)
+    moreover
+    from eval_init wt_init wf conf_s0
+    have "s1\<Colon>\<preceq>(G, L)"
+      by (auto dest: eval_type_sound)
+    from this wt_size 
+    have "G\<turnstile>s1 \<midarrow>e-\<succ>i\<rightarrow> s2"
+      by (rule hyp_size)
+    moreover note NewA
+    ultimately show ?case
+      by (blast intro!: eval.NewA)
+  next
+    case Cast
+    with wf show ?case
+      by - (erule wt_elim_cases, rule eval.Cast,auto dest: eval_type_sound)
+  next
+    case Inst
+    with wf show ?case
+      by - (erule wt_elim_cases, rule eval.Inst,auto dest: eval_type_sound)
+  next
+    case Lit
+    show ?case by (rule eval.Lit)
+  next
+    case Super
+    show ?case by (rule eval.Super)
+  next
+    case Acc
+    then show ?case
+      by - (erule wt_elim_cases, rule eval.Acc,auto dest: eval_type_sound)
+  next
+    case Ass
+    with wf show ?case
+      by - (erule wt_elim_cases, blast intro!: eval.Ass dest: eval_type_sound) 
+  next
+    case (Cond b e0 e1 e2 n s0 s1 s2 v L accC T)
+    have hyp_e0: "PROP ?EqEval (Norm s0) s1 (In1l e0) (In1 b)" .
+    have hyp_if: "PROP ?EqEval s1 s2 
+                              (In1l (if the_Bool b then e1 else e2)) (In1 v)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (e0 ? e1 : e2)\<Colon>T" .
+    then obtain T1 T2 statT where
+       wt_e0: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
+       wt_e1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e1\<Colon>-T1" and
+       wt_e2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e2\<Colon>-T2" and 
+       statT: "G\<turnstile>T1\<preceq>T2 \<and> statT = T2  \<or>  G\<turnstile>T2\<preceq>T1 \<and> statT =  T1" and
+       T    : "T=Inl statT"
+      by (rule wt_elim_cases) auto
+    from conf_s0 wt_e0
+    have eval_e0: "G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1"
+      by (rule hyp_e0)
+    moreover
+    from eval_e0 conf_s0 wf wt_e0
+    have "s1\<Colon>\<preceq>(G, L)"
+      by (blast dest: eval_type_sound)
+    with wt_e1 wt_e2 statT hyp_if
+    have "G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2"
+      by (cases "the_Bool b") auto
+    ultimately
+    show ?case
+      by (rule eval.Cond)
+  next
+    case (Call invDeclC a' accC' args e mn mode n pTs' s0 s1 s2 s4 statT 
+           v vs L accC T)
+    (* Repeats large parts of the type soundness proof. One should factor
+       out some lemmata about the relations and conformance of s2, s3 and s3'*)
+    have evaln_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s1" .
+    have evaln_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" .
+    have invDeclC: "invDeclC 
+                      = invocation_declclass G mode (store s2) a' statT 
+                           \<lparr>name = mn, parTs = pTs'\<rparr>" .
+    let ?InitLvars 
+         = "init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> mode a' vs s2"
+    obtain s3 s3' where 
+      init_lvars: "s3 = 
+             init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> mode a' vs s2" and
+      check: "s3' =
+         check_method_access G accC' statT mode \<lparr>name = mn, parTs = pTs'\<rparr> a' s3"
+      by simp
+    have evaln_methd: 
+           "G\<turnstile>?InitLvars \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<midarrow>n\<rightarrow> s4" .
+    have     hyp_e: "PROP ?EqEval (Norm s0) s1 (In1l e) (In1 a')" .
+    have  hyp_args: "PROP ?EqEval s1 s2 (In3 args) (In3 vs)" .
+    have hyp_methd: "PROP ?EqEval ?InitLvars s4 
+                     (In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)) (In1 v)".
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>
+                    \<turnstile>In1l ({accC',statT,mode}e\<cdot>mn( {pTs'}args))\<Colon>T" .
+    from wt obtain pTs statDeclT statM where
+                 wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT" and
+              wt_args: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>args\<Colon>\<doteq>pTs" and
+                statM: "max_spec G accC statT \<lparr>name=mn,parTs=pTs\<rparr> 
+                         = {((statDeclT,statM),pTs')}" and
+                 mode: "mode = invmode statM e" and
+                    T: "T =Inl (resTy statM)" and
+        eq_accC_accC': "accC=accC'"
+      by (rule wt_elim_cases) auto
+    from conf_s0 wt_e hyp_e
+    have eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1"
+      by blast
+    with wf conf_s0 wt_e
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and
+           conf_a': "normal s1 \<Longrightarrow> G, store s1\<turnstile>a'\<Colon>\<preceq>RefT statT" 
+      by (auto dest!: eval_type_sound)
+    from conf_s1 wt_args hyp_args
+    have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
+      by blast
+    with wt_args conf_s1 wf 
+    obtain    conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+            conf_args: "normal s2 
+                         \<Longrightarrow>  list_all2 (conf G (store s2)) vs pTs" 
+      by (auto dest!: eval_type_sound)
+    from statM 
+    obtain
+       statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>" and
+       pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
+      by (blast dest: max_spec2mheads)
+    from check
+    have eq_store_s3'_s3: "store s3'=store s3"
+      by (cases s3) (simp add: check_method_access_def Let_def)
+    obtain invC
+      where invC: "invC = invocation_class mode (store s2) a' statT"
+      by simp
+    with init_lvars
+    have invC': "invC = (invocation_class mode (store s3) a' statT)"
+      by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
+    show "G\<turnstile>Norm s0 \<midarrow>{accC',statT,mode}e\<cdot>mn( {pTs'}args)
+             -\<succ>v\<rightarrow> (set_lvars (locals (store s2))) s4"
+    proof (cases "normal s2")
+      case False
+      with init_lvars 
+      obtain keep_abrupt: "abrupt s3 = abrupt s2" and
+             "store s3 = store (init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> 
+                                            mode a' vs s2)" 
+	by (auto simp add: init_lvars_def2)
+      moreover
+      from keep_abrupt False check
+      have eq_s3'_s3: "s3'=s3" 
+	by (auto simp add: check_method_access_def Let_def)
+      moreover
+      from eq_s3'_s3 False keep_abrupt evaln_methd init_lvars
+      obtain "s4=s3'"
+	 "In1 v=arbitrary3 (In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>))"
+	by auto
+      moreover note False
+      ultimately have
+	"G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+	by (auto)
+      from eval_e eval_args invDeclC init_lvars check this
+      show ?thesis
+	by (rule eval.Call)
+    next
+      case True
+      note normal_s2 = True
+      with eval_args
+      have normal_s1: "normal s1"
+	by (cases "normal s1") auto
+      with conf_a' eval_args 
+      have conf_a'_s2: "G, store s2\<turnstile>a'\<Colon>\<preceq>RefT statT"
+	by (auto dest: eval_gext intro: conf_gext)
+      show ?thesis
+      proof (cases "a'=Null \<longrightarrow> is_static statM")
+	case False
+	then obtain not_static: "\<not> is_static statM" and Null: "a'=Null" 
+	  by blast
+	with normal_s2 init_lvars mode
+	obtain np: "abrupt s3 = Some (Xcpt (Std NullPointer))" and
+                   "store s3 = store (init_lvars G invDeclC 
+                                       \<lparr>name = mn, parTs = pTs'\<rparr> mode a' vs s2)"
+	  by (auto simp add: init_lvars_def2)
+	moreover
+	from np check
+	have eq_s3'_s3: "s3'=s3" 
+	  by (auto simp add: check_method_access_def Let_def)
+	moreover
+	from eq_s3'_s3 np evaln_methd init_lvars
+	obtain "s4=s3'"
+	  "In1 v=arbitrary3 (In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>))"
+	  by auto
+	moreover note np 
+	ultimately have
+	  "G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+	  by (auto)
+	from eval_e eval_args invDeclC init_lvars check this
+	show ?thesis
+	  by (rule eval.Call)
+      next
+	case True
+	with mode have notNull: "mode = IntVir \<longrightarrow> a' \<noteq> Null"
+	  by (auto dest!: Null_staticD)
+	with conf_s2 conf_a'_s2 wf invC 
+	have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
+	  by (cases s2) (auto intro: DynT_propI)
+	with wt_e statM' invC mode wf 
+	obtain dynM where 
+           dynM: "dynlookup G statT invC  \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+           acc_dynM: "G \<turnstile>Methd  \<lparr>name=mn,parTs=pTs'\<rparr> dynM 
+                          in invC dyn_accessible_from accC"
+	  by (force dest!: call_access_ok)
+	with invC' check eq_accC_accC'
+	have eq_s3'_s3: "s3'=s3"
+	  by (auto simp add: check_method_access_def Let_def)
+	from dynT_prop wf wt_e statM' mode invC invDeclC dynM 
+	obtain 
+	   wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
+	     dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+           iscls_invDeclC: "is_class G invDeclC" and
+	        invDeclC': "invDeclC = declclass dynM" and
+	     invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
+	   is_static_eq: "is_static dynM = is_static statM" and
+	   involved_classes_prop:
+             "(if invmode statM 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 invDeclC) \<or>
+                     (\<forall>statC. statT \<noteq> ClassT statC \<and> invDeclC = Object)) \<and>
+                      statDeclT = ClassT invDeclC)"
+	  by (auto dest: DynT_mheadsD)
+	obtain L' where 
+	   L':"L'=(\<lambda> k. 
+                 (case k of
+                    EName e
+                    \<Rightarrow> (case e of 
+                          VNam v 
+                          \<Rightarrow>(table_of (lcls (mbody (mthd dynM)))
+                             (pars (mthd dynM)[\<mapsto>]pTs')) v
+                        | Res \<Rightarrow> Some (resTy dynM))
+                  | This \<Rightarrow> if is_static statM 
+                            then None else Some (Class invDeclC)))"
+	  by simp
+	from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
+              wf eval_args conf_a' mode notNull wf_dynM involved_classes_prop
+	have conf_s3: "s3\<Colon>\<preceq>(G,L')"
+	   apply - 
+          (*FIXME confomrs_init_lvars should be 
+                adjusted to be more directy applicable *)
+	   apply (drule conforms_init_lvars [of G invDeclC 
+                  "\<lparr>name=mn,parTs=pTs'\<rparr>" dynM "store s2" vs pTs "abrupt s2" 
+                  L statT invC a' "(statDeclT,statM)" e])
+	     apply (rule wf)
+	     apply (rule conf_args,assumption)
+	     apply (simp add: pTs_widen)
+	     apply (cases s2,simp)
+	     apply (rule dynM')
+	     apply (force dest: ty_expr_is_type)
+	     apply (rule invC_widen)
+	     apply (force intro: conf_gext dest: eval_gext)
+	     apply simp
+	     apply simp
+	     apply (simp add: invC)
+	     apply (simp add: invDeclC)
+	     apply (force dest: wf_mdeclD1 is_acc_typeD)
+	     apply (cases s2, simp add: L' init_lvars
+	                      cong add: lname.case_cong ename.case_cong)
+	   done
+	from is_static_eq wf_dynM L'
+	obtain mthdT where
+	   "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+            \<turnstile>Body invDeclC (stmt (mbody (mthd dynM)))\<Colon>-mthdT" and
+	   mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
+	  by - (drule wf_mdecl_bodyD,
+                simp cong add: lname.case_cong ename.case_cong)
+	with dynM' iscls_invDeclC invDeclC'
+	have
+	   "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+            \<turnstile>(Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<Colon>-mthdT"
+	  by (auto intro: wt.Methd)
+	with conf_s3 hyp_methd init_lvars eq_s3'_s3
+	have "G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+	  by auto
+	from eval_e eval_args invDeclC init_lvars check this
+	show ?thesis
+	  by (rule eval.Call)
+      qed
+    qed
+  next
+    case Methd
+    with wf show ?case
+      by - (erule wt_elim_cases, rule eval.Methd, 
+            auto dest: eval_type_sound simp add: body_def2)
+  next
+    case Body
+    with wf show ?case
+       by - (erule wt_elim_cases, blast intro!: eval.Body dest: eval_type_sound)
+  next
+    case Nil
+    show ?case by (rule eval.Nil)
+  next
+    case Cons
+    with wf show ?case
+      by - (erule wt_elim_cases, blast intro!: eval.Cons dest: eval_type_sound)
+  next
+    case Skip
+    show ?case by (rule eval.Skip)
+  next
+    case Expr
+    with wf show ?case
+      by - (erule wt_elim_cases, rule eval.Expr,auto dest: eval_type_sound)
+  next
+    case Lab
+    with wf show ?case
+      by - (erule wt_elim_cases, rule eval.Lab,auto dest: eval_type_sound)
+  next
+    case Comp
+    with wf show ?case
+      by - (erule wt_elim_cases, blast intro!: eval.Comp dest: eval_type_sound)
+  next
+    case (If b c1 c2 e n s0 s1 s2 L accC T)
+    have hyp_e: "PROP ?EqEval (Norm s0) s1 (In1l e) (In1 b)" .
+    have hyp_then_else: 
+      "PROP ?EqEval s1 s2 (In1r (if the_Bool b then c1 else c2)) \<diamondsuit>" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (If(e) c1 Else c2)\<Colon>T" .
+    then obtain 
+              wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
+      wt_then_else: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
+      by (rule wt_elim_cases) (auto split add: split_if)
+    from conf_s0 wt_e
+    have eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
+      by (rule hyp_e)
+    moreover
+    from eval_e wt_e conf_s0 wf
+    have "s1\<Colon>\<preceq>(G, L)"
+      by (blast dest: eval_type_sound)
+    from this wt_then_else
+    have "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2"
+      by (rule hyp_then_else)
+    ultimately
+    show ?case
+      by (rule eval.If)
+  next
+    case (Loop b c e l n s0 s1 s2 s3 L accC T)
+    have hyp_e: "PROP ?EqEval (Norm s0) s1 (In1l e) (In1 b)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T" .
+    then obtain wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
+                wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
+      by (rule wt_elim_cases) (blast)
+    from conf_s0 wt_e 
+    have eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
+      by (rule hyp_e)
+    moreover
+    from eval_e wt_e conf_s0 wf
+    have conf_s1: "s1\<Colon>\<preceq>(G, L)"
+      by (blast dest: eval_type_sound)
+    have "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"
+    proof (cases "normal s1 \<and> the_Bool b")
+      case True 
+      from Loop True have hyp_c: "PROP ?EqEval s1 s2 (In1r c) \<diamondsuit>"
+	by (auto)
+      from Loop True have hyp_w: "PROP ?EqEval (abupd (absorb (Cont l)) s2)
+                                        s3 (In1r (l\<bullet> While(e) c)) \<diamondsuit>"
+	by (auto)
+      from conf_s1 wt_c
+      have eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2"
+	by (rule hyp_c)
+      moreover
+      from eval_c conf_s1 wt_c wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (blast dest: eval_type_sound)
+      then
+      have "abupd (absorb (Cont l)) s2 \<Colon>\<preceq>(G, L)"
+	by (cases s2) (auto intro: conforms_absorb)
+      from this and wt
+      have "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
+	by (rule hyp_w)
+      moreover note True
+      ultimately
+      show ?thesis
+	by simp
+    next
+      case False
+      with Loop have "s3 = s1" by simp
+      with False
+      show ?thesis 
+	by auto
+    qed
+    ultimately
+    show ?case
+      by (rule eval.Loop)
+  next
+    case Do
+    show ?case by (rule eval.Do)
+  next
+    case Throw
+    with wf show ?case
+      by - (erule wt_elim_cases, rule eval.Throw,auto dest: eval_type_sound)
+  next
+    case (Try c1 c2 n s0 s1 s2 s3 catchC vn L accC T)
+    have  hyp_c1: "PROP ?EqEval (Norm s0) s1 (In1r c1) \<diamondsuit>" .
+    have conf_s0:"Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>In1r (Try c1 Catch(catchC vn) c2)\<Colon>T" .
+    then obtain 
+      wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+      wt_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<lparr>lcl := L(VName vn\<mapsto>Class catchC)\<rparr>\<turnstile>c2\<Colon>\<surd>"
+      by (rule wt_elim_cases) (auto)
+    from conf_s0 wt_c1
+    have eval_c1: "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1"
+      by (rule hyp_c1)
+    moreover
+    have sxalloc: "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" .
+    moreover
+    from eval_c1 wt_c1 conf_s0 wf
+    have "s1\<Colon>\<preceq>(G, L)"
+      by (blast dest: eval_type_sound)
+    with sxalloc wf
+    have conf_s2: "s2\<Colon>\<preceq>(G, L)" 
+      by (auto dest: sxalloc_type_sound split: option.splits)
+    have "if G,s2\<turnstile>catch catchC then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2"
+    proof (cases "G,s2\<turnstile>catch catchC")
+      case True
+      note Catch = this
+      with Try have hyp_c2: "PROP ?EqEval (new_xcpt_var vn s2) s3 (In1r c2) \<diamondsuit>"
+	by auto
+      show ?thesis
+      proof (cases "normal s1")
+	case True
+	with sxalloc wf 
+	have eq_s2_s1: "s2=s1"
+	  by (auto dest: sxalloc_type_sound split: option.splits)
+	with True 
+	have "\<not>  G,s2\<turnstile>catch catchC"
+	  by (simp add: catch_def)
+	with Catch show ?thesis 
+	  by (contradiction)
+      next 
+	case False
+	with sxalloc wf
+	obtain a 
+	  where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
+	  by (auto dest!: sxalloc_type_sound split: option.splits)
+	with Catch
+	have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class catchC"
+	  by (cases s2) simp
+	with xcpt_s2 conf_s2 wf 
+	have "new_xcpt_var vn s2\<Colon>\<preceq>(G, L(VName vn\<mapsto>Class catchC))"
+	  by (auto dest: Try_lemma)
+	from this wt_c2
+	have "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3"
+	  by (auto intro: hyp_c2)
+	with Catch 
+	show ?thesis
+	  by simp
+      qed
+    next
+      case False
+      with Try
+      have "s3=s2"
+	by simp
+      with False
+      show ?thesis
+	by simp
+    qed
+    ultimately
+    show ?case
+      by (rule eval.Try)
+  next
+    case Fin
+    with wf show ?case
+      by -(erule wt_elim_cases, blast intro!: eval.Fin
+           dest: eval_type_sound intro: conforms_NormI)
+  next
+    case (Init C c n s0 s1 s2 s3 L accC T)
+    have     cls: "the (class G C) = c" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Init C)\<Colon>T" .
+    with cls
+    have cls_C: "class G C = Some c"
+      by - (erule wt_elim_cases,auto)
+    have "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)"
+    proof (cases "inited C (globs s0)")
+      case True
+      with Init have "s3 = Norm s0"
+	by simp
+      with True show ?thesis 
+	by simp
+    next
+      case False
+      with Init
+      obtain 
+	hyp_init_super: 
+        "PROP ?EqEval (Norm ((init_class_obj G C) s0)) s1
+	               (In1r (if C = Object then Skip else Init (super c))) \<diamondsuit>"
+	and 
+        hyp_init_c:
+	   "PROP ?EqEval ((set_lvars empty) s1) s2 (In1r (init c)) \<diamondsuit>" and
+	s3: "s3 = (set_lvars (locals (store s1))) s2"
+	by (simp only: if_False)
+      from conf_s0 wf cls_C False
+      have conf_s0': "(Norm ((init_class_obj G C) s0))\<Colon>\<preceq>(G, L)"
+	by (auto dest: conforms_init_class_obj)
+      moreover
+      from wf cls_C 
+      have wt_init_super:
+           "\<lparr>prg = G, cls = accC, lcl = L\<rparr>
+                  \<turnstile>(if C = Object then Skip else Init (super c))\<Colon>\<surd>"
+	by (cases "C=Object")
+           (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
+      ultimately
+      have eval_init_super: 
+	   "G\<turnstile>Norm ((init_class_obj G C) s0) 
+            \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1"
+	by (rule hyp_init_super)
+      with conf_s0' wt_init_super wf
+      have "s1\<Colon>\<preceq>(G, L)"
+	by (blast dest: eval_type_sound)
+      then
+      have "(set_lvars empty) s1\<Colon>\<preceq>(G, empty)"
+	by (cases s1) (auto dest: conforms_set_locals )
+      with wf cls_C 
+      have eval_init_c: "G\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<rightarrow> s2"
+	by (auto intro!: hyp_init_c dest: wf_prog_cdecl wf_cdecl_wt_init)
+      from False eval_init_super eval_init_c s3
+      show ?thesis
+	by simp
+    qed
+    from cls this
+    show ?case
+      by (rule eval.Init)
+  qed 
+qed
+
+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 le_max3I1: "(n2::nat) \<le> max n1 (max n2 n3)"
+proof -
+  have "n2 \<le> max n2 n3"
+    by (rule le_maxI1)
+  also
+  have "max n2 n3 \<le> max n1 (max n2 n3)"
+    by (rule le_maxI2)
+  finally
+  show ?thesis .
+qed
+
+lemma le_max3I2: "(n3::nat) \<le> max n1 (max n2 n3)"
+proof -
+  have "n3 \<le> max n2 n3"
+    by (rule le_maxI2)
+  also
+  have "max n2 n3 \<le> max n1 (max n2 n3)"
+    by (rule le_maxI2)
+  finally
+  show ?thesis .
+qed
+
+
+lemma eval_evaln: 
+ (assumes eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" and
+          wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T" and  
+     conf_s0: "s0\<Colon>\<preceq>(G, L)" and
+          wf: "wf_prog G"  
+ )  "\<exists>n. G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)"
+proof -
+  from eval 
+  show "\<And> L accC T. \<lbrakk>s0\<Colon>\<preceq>(G, L);\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T\<rbrakk>
+                     \<Longrightarrow> \<exists> n. G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)"
+       (is "PROP ?EqEval s0 s1 t v")
+  proof (induct)
+    case (Abrupt s t xc L accC T)
+    obtain n where
+      "G\<turnstile>(Some xc, s) \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (arbitrary3 t, Some xc, s)"
+      by (rules intro: evaln.Abrupt)
+    then show ?case ..
+  next
+    case Skip
+    show ?case by (blast intro: evaln.Skip)
+  next
+    case (Expr e s0 s1 v L accC T)
+    then obtain n where
+      "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+      by (rules elim!: wt_elim_cases)
+    then have "G\<turnstile>Norm s0 \<midarrow>Expr e\<midarrow>n\<rightarrow> s1"
+      by (rule evaln.Expr) 
+    then show ?case ..
+  next
+    case (Lab c l s0 s1 L accC T)
+    then obtain n where
+      "G\<turnstile>Norm s0 \<midarrow>c\<midarrow>n\<rightarrow> s1"
+      by (rules elim!: wt_elim_cases)
+    then have "G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<midarrow>n\<rightarrow> abupd (absorb (Break l)) s1"
+      by (rule evaln.Lab)
+    then show ?case ..
+  next
+    case (Comp c1 c2 s0 s1 s2 L accC T)
+    with wf obtain n1 n2 where
+      "G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n1\<rightarrow> s1"
+      "G\<turnstile>s1 \<midarrow>c2\<midarrow>n2\<rightarrow> s2"
+      by (blast elim!: wt_elim_cases dest: eval_type_sound)
+    then have "G\<turnstile>Norm s0 \<midarrow>c1;; c2\<midarrow>max n1 n2\<rightarrow> s2"
+      by (blast intro: evaln.Comp dest: evaln_max2 )
+    then show ?case ..
+  next
+    case (If b c1 c2 e s0 s1 s2 L accC T)
+    with wf obtain
+      "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean"
+      "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
+      by (cases "the_Bool b") (auto elim!: wt_elim_cases)
+    with If wf obtain n1 n2 where
+      "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n1\<rightarrow> s1"
+      "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n2\<rightarrow> s2"
+      by (blast dest: eval_type_sound)
+    then have "G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2\<midarrow>max n1 n2\<rightarrow> s2"
+      by (blast intro: evaln.If dest: evaln_max2)
+    then show ?case ..
+  next
+    case (Loop b c e l s0 s1 s2 s3 L accC T)
+    have eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1" .
+    have hyp_e: "PROP ?EqEval (Norm s0) s1 (In1l e) (In1 b)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T" .
+    then obtain wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
+                wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
+      by (rule wt_elim_cases) (blast)
+    from conf_s0 wt_e 
+    obtain n1 where
+      "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n1\<rightarrow> s1"
+      by (rules dest: hyp_e)
+    moreover
+    from eval_e wt_e conf_s0 wf
+    have conf_s1: "s1\<Colon>\<preceq>(G, L)"
+      by (rules dest: eval_type_sound)
+    obtain n2 where
+      "if normal s1 \<and> the_Bool b 
+             then (G\<turnstile>s1 \<midarrow>c\<midarrow>n2\<rightarrow> s2 \<and> 
+                   G\<turnstile>(abupd (absorb (Cont l)) s2)\<midarrow>l\<bullet> While(e) c\<midarrow>n2\<rightarrow> s3)
+	     else s3 = s1"
+    proof (cases "normal s1 \<and> the_Bool b")
+      case True
+      from Loop True have hyp_c: "PROP ?EqEval s1 s2 (In1r c) \<diamondsuit>"
+	by (auto)
+      from Loop True have hyp_w: "PROP ?EqEval (abupd (absorb (Cont l)) s2)
+                                        s3 (In1r (l\<bullet> While(e) c)) \<diamondsuit>"
+	by (auto)
+      from Loop True have eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2"
+	by simp
+      from conf_s1 wt_c
+      obtain m1 where 
+	evaln_c: "G\<turnstile>s1 \<midarrow>c\<midarrow>m1\<rightarrow> s2"
+	by (rules dest: hyp_c)
+      moreover
+      from eval_c conf_s1 wt_c wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (rules dest: eval_type_sound)
+      then
+      have "abupd (absorb (Cont l)) s2 \<Colon>\<preceq>(G, L)"
+	by (cases s2) (auto intro: conforms_absorb)
+      from this and wt
+      obtain m2 where 
+	"G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<midarrow>m2\<rightarrow> s3"
+	by (blast dest: hyp_w)
+      moreover note True and that
+      ultimately show ?thesis
+	by simp (rules intro: evaln_nonstrict le_maxI1 le_maxI2)
+    next
+      case False
+      with Loop have "s3 = s1"
+	by simp
+      with False that
+      show ?thesis
+	by auto 
+    qed
+    ultimately
+    have "G\<turnstile>Norm s0 \<midarrow>l\<bullet> While(e) c\<midarrow>max n1 n2\<rightarrow> s3"
+      apply -
+      apply (rule evaln.Loop)
+      apply   (rules intro: evaln_nonstrict intro: le_maxI1)
+
+      apply   (auto intro: evaln_nonstrict intro: le_maxI2)
+      done
+    then show ?case ..
+  next
+    case (Do j s L accC T)
+    have "G\<turnstile>Norm s \<midarrow>Do j\<midarrow>n\<rightarrow> (Some (Jump j), s)"
+      by (rule evaln.Do)
+    then show ?case ..
+  next
+    case (Throw a e s0 s1 L accC T)
+    then obtain n where
+      "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1"
+      by (rules elim!: wt_elim_cases)
+    then have "G\<turnstile>Norm s0 \<midarrow>Throw e\<midarrow>n\<rightarrow> abupd (throw a) s1"
+      by (rule evaln.Throw)
+    then show ?case ..
+  next 
+    case (Try catchC c1 c2 s0 s1 s2 s3 vn L accC T)
+    have  hyp_c1: "PROP ?EqEval (Norm s0) s1 (In1r c1) \<diamondsuit>" .
+    have eval_c1: "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>In1r (Try c1 Catch(catchC vn) c2)\<Colon>T" .
+    then obtain 
+      wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+      wt_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<lparr>lcl := L(VName vn\<mapsto>Class catchC)\<rparr>\<turnstile>c2\<Colon>\<surd>"
+      by (rule wt_elim_cases) (auto)
+    from conf_s0 wt_c1
+    obtain n1 where
+      "G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n1\<rightarrow> s1"
+      by (blast dest: hyp_c1)
+    moreover 
+    have sxalloc: "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" .
+    moreover
+    from eval_c1 wt_c1 conf_s0 wf
+    have "s1\<Colon>\<preceq>(G, L)"
+      by (blast dest: eval_type_sound)
+    with sxalloc wf
+    have conf_s2: "s2\<Colon>\<preceq>(G, L)" 
+      by (auto dest: sxalloc_type_sound split: option.splits)
+    obtain n2 where
+      "if G,s2\<turnstile>catch catchC then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n2\<rightarrow> s3 else s3 = s2"
+    proof (cases "G,s2\<turnstile>catch catchC")
+      case True
+      note Catch = this
+      with Try have hyp_c2: "PROP ?EqEval (new_xcpt_var vn s2) s3 (In1r c2) \<diamondsuit>"
+	by auto
+      show ?thesis
+      proof (cases "normal s1")
+	case True
+	with sxalloc wf 
+	have eq_s2_s1: "s2=s1"
+	  by (auto dest: sxalloc_type_sound split: option.splits)
+	with True 
+	have "\<not>  G,s2\<turnstile>catch catchC"
+	  by (simp add: catch_def)
+	with Catch show ?thesis 
+	  by (contradiction)
+      next 
+	case False
+	with sxalloc wf
+	obtain a 
+	  where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
+	  by (auto dest!: sxalloc_type_sound split: option.splits)
+	with Catch
+	have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class catchC"
+	  by (cases s2) simp
+	with xcpt_s2 conf_s2 wf 
+	have "new_xcpt_var vn s2\<Colon>\<preceq>(G, L(VName vn\<mapsto>Class catchC))"
+	  by (auto dest: Try_lemma)
+	(* FIXME extract lemma for this conformance, also usefull for
+               eval_type_sound and evaln_eval *)
+	from this wt_c2
+	obtain m where "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>m\<rightarrow> s3"
+	  by (auto dest: hyp_c2)
+	with True that
+	show ?thesis
+	  by simp
+      qed
+    next
+      case False
+      with Try
+      have "s3=s2"
+	by simp
+      with False and that
+      show ?thesis
+	by simp
+    qed
+    ultimately
+    have "G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(catchC vn) c2\<midarrow>max n1 n2\<rightarrow> s3"
+      by (auto intro!: evaln.Try le_maxI1 le_maxI2)
+    then show ?case ..
+  next
+    case (Fin c1 c2 s0 s1 s2 x1 L accC T)
+    with wf obtain n1 n2 where 
+      "G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n1\<rightarrow> (x1, s1)"
+      "G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n2\<rightarrow> s2"
+      by (blast elim!: wt_elim_cases 
+	         dest: eval_type_sound intro: conforms_NormI)
+    then have 
+     "G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<midarrow>max n1 n2\<rightarrow> abupd (abrupt_if (x1 \<noteq> None) x1) s2"
+      by (blast intro: evaln.Fin dest: evaln_max2)
+    then show ?case ..
+  next
+    case (Init C c s0 s1 s2 s3 L accC T)
+    have     cls: "the (class G C) = c" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Init C)\<Colon>T" .
+    with cls
+    have cls_C: "class G C = Some c"
+      by - (erule wt_elim_cases,auto)
+    obtain n where
+      "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)"
+    proof (cases "inited C (globs s0)")
+      case True
+      with Init have "s3 = Norm s0"
+	by simp
+      with True that show ?thesis 
+	by simp
+    next
+      case False
+      with Init
+      obtain 
+	hyp_init_super: 
+        "PROP ?EqEval (Norm ((init_class_obj G C) s0)) s1
+	               (In1r (if C = Object then Skip else Init (super c))) \<diamondsuit>"
+	and 
+        hyp_init_c:
+	   "PROP ?EqEval ((set_lvars empty) s1) s2 (In1r (init c)) \<diamondsuit>" and
+	s3: "s3 = (set_lvars (locals (store s1))) s2" and
+	eval_init_super: 
+	"G\<turnstile>Norm ((init_class_obj G C) s0) 
+           \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1"
+	by (simp only: if_False)
+      from conf_s0 wf cls_C False
+      have conf_s0': "(Norm ((init_class_obj G C) s0))\<Colon>\<preceq>(G, L)"
+	by (auto dest: conforms_init_class_obj)
+      moreover
+      from wf cls_C 
+      have wt_init_super:
+           "\<lparr>prg = G, cls = accC, lcl = L\<rparr>
+                  \<turnstile>(if C = Object then Skip else Init (super c))\<Colon>\<surd>"
+	by (cases "C=Object")
+           (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
+      ultimately
+      obtain m1 where  
+	   "G\<turnstile>Norm ((init_class_obj G C) s0) 
+            \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>m1\<rightarrow> s1"
+	by (rules dest: hyp_init_super)
+      moreover
+      from eval_init_super conf_s0' wt_init_super wf
+      have "s1\<Colon>\<preceq>(G, L)"
+	by (rules dest: eval_type_sound)
+      then
+      have "(set_lvars empty) s1\<Colon>\<preceq>(G, empty)"
+	by (cases s1) (auto dest: conforms_set_locals )
+      with wf cls_C 
+      obtain m2 where
+	"G\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<midarrow>m2\<rightarrow> s2"
+	by (blast dest!: hyp_init_c 
+                   dest: wf_prog_cdecl intro!: wf_cdecl_wt_init)
+      moreover note s3 and False and that
+      ultimately show ?thesis
+	by simp (rules intro: evaln_nonstrict le_maxI1 le_maxI2)
+    qed
+    from cls this have "G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s3"
+      by (rule evaln.Init)
+    then show ?case ..
+  next
+    case (NewC C a s0 s1 s2 L accC T)
+    with wf obtain n where 
+     "G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1"
+      by (blast elim!: wt_elim_cases dest: is_acc_classD)
+    with NewC 
+    have "G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<midarrow>n\<rightarrow> s2"
+      by (rules intro: evaln.NewC)
+    then show ?case ..
+  next
+    case (NewA T a e i s0 s1 s2 s3 L accC Ta)
+    hence "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" 
+      by (auto elim!: wt_elim_cases 
+              intro!: wt_init_comp_ty dest: is_acc_typeD)
+    with NewA wf obtain n1 n2 where 
+      "G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<midarrow>n1\<rightarrow> s1"
+      "G\<turnstile>s1 \<midarrow>e-\<succ>i\<midarrow>n2\<rightarrow> s2"      
+      by (blast elim!: wt_elim_cases dest: eval_type_sound)
+    moreover
+    have "G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3" .
+    ultimately
+    have "G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<midarrow>max n1 n2\<rightarrow> s3"
+      by (blast intro: evaln.NewA dest: evaln_max2)
+    then show ?case ..
+  next
+    case (Cast castT e s0 s1 s2 v L accC T)
+    with wf obtain n where
+      "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+      by (rules elim!: wt_elim_cases)
+    moreover 
+    have "s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits castT) ClassCast) s1" .
+    ultimately
+    have "G\<turnstile>Norm s0 \<midarrow>Cast castT e-\<succ>v\<midarrow>n\<rightarrow> s2"
+      by (rule evaln.Cast)
+    then show ?case ..
+  next
+    case (Inst T b e s0 s1 v L accC T')
+    with wf obtain n where
+      "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+      by (rules elim!: wt_elim_cases)
+    moreover 
+    have "b = (v \<noteq> Null \<and> G,snd s1\<turnstile>v fits RefT T)" .
+    ultimately
+    have "G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<midarrow>n\<rightarrow> s1"
+      by (rule evaln.Inst)
+    then show ?case ..
+  next
+    case (Lit s v L accC T)
+    have "G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<midarrow>n\<rightarrow> Norm s"
+      by (rule evaln.Lit)
+    then show ?case ..
+  next
+    case (Super s L accC T)
+    have "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<midarrow>n\<rightarrow> Norm s"
+      by (rule evaln.Super)
+    then show ?case ..
+  next
+    case (Acc f s0 s1 v va L accC T)
+    with wf obtain n where
+      "G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v, f)\<midarrow>n\<rightarrow> s1"
+      by (rules elim!: wt_elim_cases)
+    then
+    have "G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<midarrow>n\<rightarrow> s1"
+      by (rule evaln.Acc)
+    then show ?case ..
+  next
+    case (Ass e f s0 s1 s2 v var w L accC T)
+    with wf obtain n1 n2 where 
+      "G\<turnstile>Norm s0 \<midarrow>var=\<succ>(w, f)\<midarrow>n1\<rightarrow> s1"
+      "G\<turnstile>s1 \<midarrow>e-\<succ>v\<midarrow>n2\<rightarrow> s2"      
+      by (blast elim!: wt_elim_cases dest: eval_type_sound)
+    then
+    have "G\<turnstile>Norm s0 \<midarrow>var:=e-\<succ>v\<midarrow>max n1 n2\<rightarrow> assign f v s2"
+      by (blast intro: evaln.Ass dest: evaln_max2)
+    then show ?case ..
+  next
+    case (Cond b e0 e1 e2 s0 s1 s2 v L accC T)
+    have hyp_e0: "PROP ?EqEval (Norm s0) s1 (In1l e0) (In1 b)" .
+    have hyp_if: "PROP ?EqEval s1 s2 
+                              (In1l (if the_Bool b then e1 else e2)) (In1 v)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (e0 ? e1 : e2)\<Colon>T" .
+    then obtain T1 T2 statT where
+       wt_e0: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
+       wt_e1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e1\<Colon>-T1" and
+       wt_e2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e2\<Colon>-T2" and 
+       statT: "G\<turnstile>T1\<preceq>T2 \<and> statT = T2  \<or>  G\<turnstile>T2\<preceq>T1 \<and> statT =  T1" and
+       T    : "T=Inl statT"
+      by (rule wt_elim_cases) auto
+    have eval_e0: "G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1" .
+    from conf_s0 wt_e0
+    obtain n1 where 
+      "G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<midarrow>n1\<rightarrow> s1"
+      by (rules dest: hyp_e0)
+    moreover
+    from eval_e0 conf_s0 wf wt_e0
+    have "s1\<Colon>\<preceq>(G, L)"
+      by (blast dest: eval_type_sound)
+    with wt_e1 wt_e2 statT hyp_if obtain n2 where
+      "G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<midarrow>n2\<rightarrow> s2"
+      by  (cases "the_Bool b") force+
+    ultimately
+    have "G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<midarrow>max n1 n2\<rightarrow> s2"
+      by (blast intro: evaln.Cond dest: evaln_max2)
+    then show ?case ..
+  next
+    case (Call invDeclC a' accC' args e mn mode pTs' s0 s1 s2 s3 s3' s4 statT 
+      v vs L accC T)
+    (* Repeats large parts of the type soundness proof. One should factor
+       out some lemmata about the relations and conformance of s2, s3 and s3'*)
+    have eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1" .
+    have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2" .
+    have invDeclC: "invDeclC 
+                      = invocation_declclass G mode (store s2) a' statT 
+                           \<lparr>name = mn, parTs = pTs'\<rparr>" .
+    have
+      init_lvars: "s3 = 
+             init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> mode a' vs s2" .
+    have
+      check: "s3' =
+       check_method_access G accC' statT mode \<lparr>name = mn, parTs = pTs'\<rparr> a' s3" .
+    have eval_methd: 
+           "G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<rightarrow> s4" .
+    have     hyp_e: "PROP ?EqEval (Norm s0) s1 (In1l e) (In1 a')" .
+    have  hyp_args: "PROP ?EqEval s1 s2 (In3 args) (In3 vs)" .
+    have hyp_methd: "PROP ?EqEval s3' s4 
+                     (In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)) (In1 v)".
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>
+                    \<turnstile>In1l ({accC',statT,mode}e\<cdot>mn( {pTs'}args))\<Colon>T" .
+    from wt obtain pTs statDeclT statM where
+                 wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT" and
+              wt_args: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>args\<Colon>\<doteq>pTs" and
+                statM: "max_spec G accC statT \<lparr>name=mn,parTs=pTs\<rparr> 
+                         = {((statDeclT,statM),pTs')}" and
+                 mode: "mode = invmode statM e" and
+                    T: "T =Inl (resTy statM)" and
+        eq_accC_accC': "accC=accC'"
+      by (rule wt_elim_cases) auto
+    from conf_s0 wt_e
+    obtain n1 where
+      evaln_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<midarrow>n1\<rightarrow> s1"
+      by (rules dest: hyp_e)
+    from wf eval_e conf_s0 wt_e
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and
+           conf_a': "normal s1 \<Longrightarrow> G, store s1\<turnstile>a'\<Colon>\<preceq>RefT statT"  
+      by (auto dest!: eval_type_sound)
+    from conf_s1 wt_args
+    obtain n2 where
+      evaln_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n2\<rightarrow> s2"
+      by (blast dest: hyp_args)
+    from wt_args conf_s1 eval_args wf 
+    obtain    conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+            conf_args: "normal s2 
+                         \<Longrightarrow>  list_all2 (conf G (store s2)) vs pTs"  
+      by (auto dest!: eval_type_sound)
+    from statM 
+    obtain
+       statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>" and
+       pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
+      by (blast dest: max_spec2mheads)
+    from check
+    have eq_store_s3'_s3: "store s3'=store s3"
+      by (cases s3) (simp add: check_method_access_def Let_def)
+    obtain invC
+      where invC: "invC = invocation_class mode (store s2) a' statT"
+      by simp
+    with init_lvars
+    have invC': "invC = (invocation_class mode (store s3) a' statT)"
+      by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
+    obtain n3 where
+     "G\<turnstile>Norm s0 \<midarrow>{accC',statT,mode}e\<cdot>mn( {pTs'}args)-\<succ>v\<midarrow>n3\<rightarrow> 
+          (set_lvars (locals (store s2))) s4"
+    proof (cases "normal s2")
+      case False
+      with init_lvars 
+      obtain keep_abrupt: "abrupt s3 = abrupt s2" and
+             "store s3 = store (init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> 
+                                            mode a' vs s2)" 
+	by (auto simp add: init_lvars_def2)
+      moreover
+      from keep_abrupt False check
+      have eq_s3'_s3: "s3'=s3" 
+	by (auto simp add: check_method_access_def Let_def)
+      moreover
+      from eq_s3'_s3 False keep_abrupt eval_methd init_lvars
+      obtain "s4=s3'"
+	 "In1 v=arbitrary3 (In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>))"
+	by auto
+      moreover note False evaln.Abrupt
+      ultimately obtain m where 
+	"G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<midarrow>m\<rightarrow> s4"
+	by force
+      from evaln_e evaln_args invDeclC init_lvars eq_s3'_s3 this
+      have 
+       "G\<turnstile>Norm s0 \<midarrow>{accC',statT,mode}e\<cdot>mn( {pTs'}args)-\<succ>v\<midarrow>max n1 (max n2 m)\<rightarrow> 
+            (set_lvars (locals (store s2))) s4"
+	by (auto intro!: evaln.Call le_maxI1 le_max3I1 le_max3I2)
+      with that show ?thesis 
+	by rules
+    next
+      case True
+      note normal_s2 = True
+      with eval_args
+      have normal_s1: "normal s1"
+	by (cases "normal s1") auto
+      with conf_a' eval_args 
+      have conf_a'_s2: "G, store s2\<turnstile>a'\<Colon>\<preceq>RefT statT"
+	by (auto dest: eval_gext intro: conf_gext)
+      show ?thesis
+      proof (cases "a'=Null \<longrightarrow> is_static statM")
+	case False
+	then obtain not_static: "\<not> is_static statM" and Null: "a'=Null" 
+	  by blast
+	with normal_s2 init_lvars mode
+	obtain np: "abrupt s3 = Some (Xcpt (Std NullPointer))" and
+                   "store s3 = store (init_lvars G invDeclC 
+                                       \<lparr>name = mn, parTs = pTs'\<rparr> mode a' vs s2)"
+	  by (auto simp add: init_lvars_def2)
+	moreover
+	from np check
+	have eq_s3'_s3: "s3'=s3" 
+	  by (auto simp add: check_method_access_def Let_def)
+	moreover
+	from eq_s3'_s3 np eval_methd init_lvars
+	obtain "s4=s3'"
+	  "In1 v=arbitrary3 (In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>))"
+	  by auto
+	moreover note np
+	ultimately obtain m where 
+	  "G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<midarrow>m\<rightarrow> s4"
+	  by force
+	from evaln_e evaln_args invDeclC init_lvars eq_s3'_s3 this
+	have 
+        "G\<turnstile>Norm s0 \<midarrow>{accC',statT,mode}e\<cdot>mn( {pTs'}args)-\<succ>v\<midarrow>max n1 (max n2 m)\<rightarrow> 
+            (set_lvars (locals (store s2))) s4"
+	  by (auto intro!: evaln.Call le_maxI1 le_max3I1 le_max3I2)
+	with that show ?thesis 
+	  by rules
+      next
+	case True
+	with mode have notNull: "mode = IntVir \<longrightarrow> a' \<noteq> Null"
+	  by (auto dest!: Null_staticD)
+	with conf_s2 conf_a'_s2 wf invC 
+	have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
+	  by (cases s2) (auto intro: DynT_propI)
+	with wt_e statM' invC mode wf 
+	obtain dynM where 
+           dynM: "dynlookup G statT invC  \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+           acc_dynM: "G \<turnstile>Methd  \<lparr>name=mn,parTs=pTs'\<rparr> dynM 
+                          in invC dyn_accessible_from accC"
+	  by (force dest!: call_access_ok)
+	with invC' check eq_accC_accC'
+	have eq_s3'_s3: "s3'=s3"
+	  by (auto simp add: check_method_access_def Let_def)
+	from dynT_prop wf wt_e statM' mode invC invDeclC dynM 
+	obtain 
+	   wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
+	     dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+           iscls_invDeclC: "is_class G invDeclC" and
+	        invDeclC': "invDeclC = declclass dynM" and
+	     invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
+	   is_static_eq: "is_static dynM = is_static statM" and
+	   involved_classes_prop:
+             "(if invmode statM 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 invDeclC) \<or>
+                     (\<forall>statC. statT \<noteq> ClassT statC \<and> invDeclC = Object)) \<and>
+                      statDeclT = ClassT invDeclC)"
+	  by (auto dest: DynT_mheadsD)
+	obtain L' where 
+	   L':"L'=(\<lambda> k. 
+                 (case k of
+                    EName e
+                    \<Rightarrow> (case e of 
+                          VNam v 
+                          \<Rightarrow>(table_of (lcls (mbody (mthd dynM)))
+                             (pars (mthd dynM)[\<mapsto>]pTs')) v
+                        | Res \<Rightarrow> Some (resTy dynM))
+                  | This \<Rightarrow> if is_static statM 
+                            then None else Some (Class invDeclC)))"
+	  by simp
+	from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
+              wf eval_args conf_a' mode notNull wf_dynM involved_classes_prop
+	have conf_s3: "s3\<Colon>\<preceq>(G,L')"
+	   apply - 
+          (*FIXME confomrs_init_lvars should be 
+                adjusted to be more directy applicable *)
+	   apply (drule conforms_init_lvars [of G invDeclC 
+                  "\<lparr>name=mn,parTs=pTs'\<rparr>" dynM "store s2" vs pTs "abrupt s2" 
+                  L statT invC a' "(statDeclT,statM)" e])
+	     apply (rule wf)
+	     apply (rule conf_args,assumption)
+	     apply (simp add: pTs_widen)
+	     apply (cases s2,simp)
+	     apply (rule dynM')
+	     apply (force dest: ty_expr_is_type)
+	     apply (rule invC_widen)
+	     apply (force intro: conf_gext dest: eval_gext)
+	     apply simp
+	     apply simp
+	     apply (simp add: invC)
+	     apply (simp add: invDeclC)
+	     apply (force dest: wf_mdeclD1 is_acc_typeD)
+	     apply (cases s2, simp add: L' init_lvars
+	                      cong add: lname.case_cong ename.case_cong)
+	   done
+	with is_static_eq wf_dynM L'
+	obtain mthdT where
+	   "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+            \<turnstile>Body invDeclC (stmt (mbody (mthd dynM)))\<Colon>-mthdT" 
+	  by - (drule wf_mdecl_bodyD,
+                simp cong add: lname.case_cong ename.case_cong)
+	with dynM' iscls_invDeclC invDeclC'
+	have
+	   "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+            \<turnstile>(Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<Colon>-mthdT"
+	  by (auto intro: wt.Methd)
+	with conf_s3 eq_s3'_s3 hyp_methd
+	obtain m where
+	   "G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<midarrow>m\<rightarrow> s4"
+	  by (blast)
+	from evaln_e evaln_args invDeclC init_lvars  eq_s3'_s3 this
+	have 
+        "G\<turnstile>Norm s0 \<midarrow>{accC',statT,mode}e\<cdot>mn( {pTs'}args)-\<succ>v\<midarrow>max n1 (max n2 m)\<rightarrow> 
+            (set_lvars (locals (store s2))) s4"
+	  by (auto intro!: evaln.Call le_maxI1 le_max3I1 le_max3I2)
+	with that show ?thesis 
+	  by rules
+      qed
+    qed
+    then show ?case ..
+  next
+    case (Methd D s0 s1 sig v L accC T)
+    then obtain n where
+      "G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<midarrow>n\<rightarrow> s1"
+      by - (erule wt_elim_cases, force simp add: body_def2)
+    then have "G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<midarrow>Suc n\<rightarrow> s1"
+      by (rule evaln.Methd)
+    then show ?case ..
+  next
+    case (Body D c s0 s1 s2 L accC T)
+    with wf obtain n1 n2 where 
+      "G\<turnstile>Norm s0 \<midarrow>Init D\<midarrow>n1\<rightarrow> s1"
+      "G\<turnstile>s1 \<midarrow>c\<midarrow>n2\<rightarrow> s2"
+      by (blast elim!: wt_elim_cases dest: eval_type_sound)
+    then have 
+     "G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)\<midarrow>max n1 n2
+       \<rightarrow> abupd (absorb Ret) s2"
+      by (blast intro: evaln.Body dest: evaln_max2)
+    then show ?case ..
+  next
+    case (LVar s vn L accC T)
+    obtain n where
+      "G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<midarrow>n\<rightarrow> Norm s"
+      by (rules intro: evaln.LVar)
+    then show ?case ..
+  next
+    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v L accC' T)
+    have eval_init: "G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1" .
+    have eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2" .
+    have check: "s3 = check_field_access G accC statDeclC fn stat a s2'" .
+    have hyp_init: "PROP ?EqEval (Norm s0) s1 (In1r (Init statDeclC)) \<diamondsuit>" .
+    have hyp_e: "PROP ?EqEval s1 s2 (In1l e) (In1 a)" .
+    have fvar: "(v, s2') = fvar statDeclC stat fn a s2" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have wt: "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile>In2 ({accC,statDeclC,stat}e..fn)\<Colon>T" .
+    then obtain statC f where
+                wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-Class statC" and
+            accfield: "accfield G accC statC fn = Some (statDeclC,f)" and
+                stat: "stat=is_static f" and
+               accC': "accC'=accC" and
+	           T: "T=(Inl (type f))"
+       by (rule wt_elim_cases) (auto simp add: member_is_static_simp)
+    from wf wt_e 
+    have iscls_statC: "is_class G statC"
+      by (auto dest: ty_expr_is_type type_is_class)
+    with wf accfield 
+    have iscls_statDeclC: "is_class G statDeclC"
+      by (auto dest!: accfield_fields dest: fields_declC)
+    then 
+    have wt_init: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
+      by simp
+    from conf_s0 wt_init
+    obtain n1 where
+      evaln_init: "G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<midarrow>n1\<rightarrow> s1"
+      by (rules dest: hyp_init)
+    from eval_init wt_init conf_s0 wf 
+    have conf_s1: "s1\<Colon>\<preceq>(G, L)"
+      by (blast dest: eval_type_sound)
+    with wt_e
+    obtain n2 where
+      evaln_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<midarrow>n2\<rightarrow> s2"
+      by (blast dest: hyp_e)
+    from eval_e wf conf_s1 wt_e
+    obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+            conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+      by (auto dest!: eval_type_sound)
+    from accfield wt_e eval_init eval_e conf_s2 conf_a fvar stat check  wf
+    have eq_s3_s2': "s3=s2'"  
+      by (auto dest!: error_free_field_access)
+    with evaln_init evaln_e fvar accC'
+    have "G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v\<midarrow>max n1 n2\<rightarrow> s3"
+      by (auto intro: evaln.FVar dest: evaln_max2)
+    then show ?case ..
+  next
+    case (AVar a e1 e2 i s0 s1 s2 s2' v L accC T)
+    with wf obtain n1 n2 where 
+      "G\<turnstile>Norm s0 \<midarrow>e1-\<succ>a\<midarrow>n1\<rightarrow> s1"
+      "G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n2\<rightarrow> s2"      
+      by (blast elim!: wt_elim_cases dest: eval_type_sound)
+    moreover 
+    have "(v, s2') = avar G i a s2" .
+    ultimately 
+    have "G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<midarrow>max n1 n2\<rightarrow> s2'"
+      by (blast intro!: evaln.AVar dest: evaln_max2)
+    then show ?case ..
+  next
+    case (Nil s0 L accC T)
+    show ?case by (rules intro: evaln.Nil)
+  next
+    case (Cons e es s0 s1 s2 v vs L accC T)
+    with wf obtain n1 n2 where 
+      "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n1\<rightarrow> s1"
+      "G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n2\<rightarrow> s2"      
+      by (blast elim!: wt_elim_cases dest: eval_type_sound)
+    then
+    have "G\<turnstile>Norm s0 \<midarrow>e # es\<doteq>\<succ>v # vs\<midarrow>max n1 n2\<rightarrow> s2"
+      by (blast intro!: evaln.Cons dest: evaln_max2)
+    then show ?case ..
+  qed
+qed
+
 end
--- a/src/HOL/Bali/Example.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/Example.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -7,7 +7,6 @@
 
 theory Example = Eval + WellForm:
 
-
 text {*
 The following example Bali program includes:
 \begin{itemize}
@@ -43,7 +42,7 @@
   }
 }
 
-public class Example {
+public class Main {
   public static void main(String args[]) throws Throwable {
     Base e = new Ext();
     try {e.foo(null); }
@@ -54,7 +53,6 @@
 }
 \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))"
@@ -68,7 +66,7 @@
 section "type and expression names"
 
 (** unfortunately cannot simply instantiate tnam **)
-datatype tnam_  = HasFoo_ | Base_ | Ext_
+datatype tnam_  = HasFoo_ | Base_ | Ext_ | Main_
 datatype vnam_  = arr_ | vee_ | z_ | e_
 datatype label_ = lab1_
 
@@ -94,6 +92,7 @@
   HasFoo :: qtname
   Base   :: qtname
   Ext    :: qtname
+  Main   :: qtname
   arr :: ename
   vee :: ename
   z   :: ename
@@ -104,6 +103,7 @@
   "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>"
+  "Main"   == "\<lparr>pid=java_lang,tid=TName (tnam_ Main_)\<rparr>"
   "arr"    ==        "(vnam_ arr_)"
   "vee"    ==        "(vnam_ vee_)"
   "z"      ==        "(vnam_ z_)"
@@ -117,12 +117,18 @@
 lemma neq_Ext_Object [simp]: "Ext\<noteq>Object"
 by (simp add: Object_def)
 
+lemma neq_Main_Object [simp]: "Main\<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)
 
+lemma neq_Main_Object [simp]: "Main\<noteq>SXcpt xn"
+by (simp add: SXcpt_def)
+
 section "classes and interfaces"
 
 defs
@@ -147,26 +153,28 @@
   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>)"
-  
+
+constdefs
   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 := 
+                     ,stmt=Expr({Ext,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"
+arr_viewed_from :: "qtname \<Rightarrow> qtname \<Rightarrow> var"
+"arr_viewed_from accC C \<equiv> {accC,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)]),
+           init=Expr(arr_viewed_from Base Base 
+                     :=New (PrimT Boolean)[Lit (Intg 2)]),
            super=Object,
            superIfs=[HasFoo]\<rparr>"
   
@@ -178,6 +186,15 @@
            super=Base,
            superIfs=[]\<rparr>"
 
+MainCl :: class
+"MainCl \<equiv> \<lparr>access=Public,
+           cfields=[], 
+           methods=[], 
+           init=Skip,
+           super=Object,
+           superIfs=[]\<rparr>"
+(* The "main" method is modeled seperately (see tprg) *)
+
 constdefs
   
   HasFooInt :: iface
@@ -187,7 +204,7 @@
  "Ifaces \<equiv> [(HasFoo,HasFooInt)]"
 
  "Classes" ::"cdecl list"
- "Classes \<equiv> [(Base,BaseCl),(Ext,ExtCl)]@standard_classes"
+ "Classes \<equiv> [(Base,BaseCl),(Ext,ExtCl),(Main,MainCl)]@standard_classes"
 
 lemmas table_classes_defs = 
      Classes_def standard_classes_def ObjectC_def SXcptC_def
@@ -231,6 +248,10 @@
 apply (simp (no_asm) add: Object_def SXcpt_def)
 done
 
+lemma table_classes_Main [simp]: "table_of Classes Main = Some MainCl"
+apply (unfold table_classes_defs )
+apply (simp (no_asm) add: Object_def SXcpt_def)
+done
 
 section "program"
 
@@ -243,9 +264,10 @@
 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> Try Expr({Main,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)"
+           (lab1\<bullet> While(Acc 
+                        (Acc (arr_viewed_from Main Ext).[Lit (Intg 2)])) Skip)"
 
 
 section "well-structuredness"
@@ -278,7 +300,7 @@
 apply (erule ssubst)
 apply (rule tnam_.induct)
 apply  safe
-apply (auto dest!: tranclD subcls1D simp add: BaseCl_def ExtCl_def)
+apply (auto dest!: tranclD subcls1D simp add: BaseCl_def ExtCl_def MainCl_def)
 apply (drule rtranclD)
 apply auto
 done
@@ -314,8 +336,13 @@
 apply auto
 done
 
+lemma ws_cdecl_Main: "ws_cdecl tprg Main Object"
+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
+                   ws_cdecl_Base ws_cdecl_Ext ws_cdecl_Main
 
 declare not_Object_subcls_any [rule del]
           not_Throwable_subcls_SXcpt [rule del] 
@@ -329,7 +356,7 @@
 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 (simp (no_asm) add: Classes_def BaseCl_def ExtCl_def MainCl_def)
 apply (auto intro!: ws_cdecls simp add: standard_classes_def ObjectC_def 
         SXcptC_def)
 done
@@ -438,12 +465,6 @@
 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)
@@ -504,7 +525,7 @@
 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
+                      BaseCl_def ExtCl_def MainCl_def
                       SXcptC_def ObjectC_def) 
 done
 
@@ -522,6 +543,13 @@
     by (auto simp add: superclasses_rec  ExtCl_def BaseCl_def)
 qed
 
+lemma superclassesMain [simp]: "superclasses tprg Main={Object}"
+proof -
+  have ws: "ws_prog tprg" by (rule ws_tprg)
+  then show ?thesis
+    by (auto simp add: superclasses_rec  MainCl_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)
 
@@ -564,12 +592,6 @@
   "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 
@@ -617,16 +639,6 @@
 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))"
@@ -667,25 +679,42 @@
            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 Base_foo_member_in_Base: 
+  "tprg\<turnstile>(Base,mdecl Base_foo) member_in Base"
+by (rule member_of_to_member_in [OF Base_foo_member_of_Base])
+
+lemma Base_foo_member_of_Base: 
+  "tprg\<turnstile>(Base,mdecl Base_foo) member_of Base"
+by (auto intro!: members.Immediate Base_declares_foo)
+
 lemma Ext_foo_member_of_Ext: 
   "tprg\<turnstile>(Ext,mdecl Ext_foo) member_of Ext"
 by (auto intro!: members.Immediate Ext_declares_foo)
 
+lemma Ext_foo_member_in_Ext: 
+  "tprg\<turnstile>(Ext,mdecl Ext_foo) member_in Ext"
+by (rule member_of_to_member_in [OF Ext_foo_member_of_Ext])
+
 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 
+by (auto intro: accessible_fromR.Immediate 
                 Base_foo_member_of_Base Base_foo_permits_acc)
 
+lemma Base_foo_dyn_accessible [simp]:
+ "tprg\<turnstile>(Base,mdecl Base_foo) in Base dyn_accessible_from S"
+apply (rule dyn_accessible_fromR.Immediate)
+apply   (rule Base_foo_member_in_Base)
+apply   (rule Base_foo_permits_acc)
+done
+
 lemma accmethd_Base [simp]: 
   "accmethd tprg S Base = methd tprg Base"
 apply (simp add: accmethd_def)
@@ -699,17 +728,15 @@
 
 lemma Ext_foo_accessible [simp]:
  "tprg\<turnstile>(Ext,mdecl Ext_foo) of Ext accessible_from S"
-by (auto intro: accessible_fromR.immediate 
+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)
+lemma Ext_foo_dyn_accessible [simp]:
+ "tprg\<turnstile>(Ext,mdecl Ext_foo) in Ext dyn_accessible_from S"
+apply (rule dyn_accessible_fromR.Immediate) 
+apply   (rule Ext_foo_member_in_Ext)
+apply   (rule Ext_foo_permits_acc)
 done
-*)
 
 lemma Ext_foo_overrides_Base_foo:
  "tprg\<turnstile>(Ext,Ext_foo) overrides (Base,Base_foo)"
@@ -732,29 +759,6 @@
     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)
@@ -762,7 +766,6 @@
 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:
@@ -790,7 +793,7 @@
                       declared_in_def 
                       cdeclaredfield_def
                intro!: filter_tab_all_True_Some filter_tab_None
-                       accessible_fromR.immediate
+                       accessible_fromR.Immediate
                intro: members.Immediate)
 done
 
@@ -802,6 +805,12 @@
 by (auto intro: members.Immediate 
        simp add: declared_in_def cdeclaredfield_def BaseCl_def)
  
+lemma arr_member_in_Base:
+  "tprg\<turnstile>(Base, fdecl (arr, 
+                 \<lparr>access = Public, static = True, type = PrimT Boolean.[]\<rparr>))
+          member_in Base"
+by (rule member_of_to_member_in [OF arr_member_of_Base])
+
 lemma arr_member_of_Ext:
   "tprg\<turnstile>(Base, fdecl (arr, 
                     \<lparr>access = Public, static = True, type = PrimT Boolean.[]\<rparr>))
@@ -812,6 +821,12 @@
 apply   (auto intro: arr_member_of_Base simp add: subcls1_def ExtCl_def)
 done
 
+lemma arr_member_in_Ext:
+  "tprg\<turnstile>(Base, fdecl (arr, 
+                 \<lparr>access = Public, static = True, type = PrimT Boolean.[]\<rparr>))
+          member_in Ext"
+by (rule member_of_to_member_in [OF arr_member_of_Ext])
+
 lemma Ext_fields_accessible[simp]:
 "accfield tprg S Ext 
   = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (DeclConcepts.fields tprg Ext))"
@@ -822,11 +837,27 @@
                       ExtCl_def
                       permits_acc_def
                intro!: filter_tab_all_True_Some filter_tab_None
-                       accessible_fromR.immediate)
+                       accessible_fromR.Immediate)
 apply (auto intro: members.Immediate arr_member_of_Ext
             simp add: declared_in_def cdeclaredfield_def ExtCl_def)
 done
 
+lemma arr_Base_dyn_accessible [simp]:
+"tprg\<turnstile>(Base, fdecl (arr, \<lparr>access=Public,static=True ,type=PrimT Boolean.[]\<rparr>)) 
+       in Base dyn_accessible_from S"
+apply (rule dyn_accessible_fromR.Immediate)
+apply   (rule arr_member_in_Base)
+apply   (simp add: permits_acc_def)
+done
+
+lemma arr_Ext_dyn_accessible[simp]:
+"tprg\<turnstile>(Base, fdecl (arr, \<lparr>access=Public,static=True ,type=PrimT Boolean.[]\<rparr>)) 
+       in Ext dyn_accessible_from S"
+apply (rule dyn_accessible_fromR.Immediate)
+apply   (rule arr_member_in_Ext)
+apply   (simp add: permits_acc_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)
@@ -853,6 +884,8 @@
                       member_is_static_simp )
 done
 
+
+declare member_is_static_simp [simp]
 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
@@ -903,6 +936,11 @@
 apply  blast+
 done
 
+lemma wf_MainC: "wf_cdecl tprg (Main,MainCl)"
+apply (unfold wf_cdecl_def MainCl_def)
+apply (auto intro: ws_cdecl_Main)
+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))
@@ -922,8 +960,9 @@
 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_BaseC [THEN conjI])
+apply   (rule wf_ExtC [THEN conjI])
+apply  (rule wf_MainC [THEN conjI])
 apply (rule wf_cdecl_all_standard_classes)
 done
 
@@ -966,7 +1005,7 @@
 
 section "well-typedness"
 
-lemma wt_test: "\<lparr>prg=tprg,cls=S,lcl=empty(VName e\<mapsto>Class Base)\<rparr>\<turnstile>test ?pTs\<Colon>\<surd>"
+lemma wt_test: "\<lparr>prg=tprg,cls=Main,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 (* ;; *))
@@ -999,6 +1038,7 @@
 apply   (simp)
 apply  (simp)
 apply  (simp)
+apply  (simp)
 apply (rule wtIs (* While *))
 apply  (rule wtIs (* Acc *))
 apply   (rule wtIs (* AVar *))
@@ -1009,6 +1049,7 @@
 apply   (simp)
 apply   (simp )
 apply   (simp)
+apply   (simp)
 apply  (rule wtIs (* LVar *))
 apply  (simp)
 apply (rule wtIs (* Skip *))
@@ -1165,9 +1206,10 @@
 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 init_done, simp)
+apply        (rule eval_Is (* StatRef *))
+apply       (simp)
+apply     (simp add: check_field_access_def Let_def) 
 apply   (rule eval_Is (* NewA *))
 apply     (simp)
 apply    (rule eval_Is (* Lit *))
@@ -1201,7 +1243,12 @@
 apply      (rule eval_Is (* Lit *))
 apply     (rule eval_Is (* Nil *))
 apply    (simp)
-apply   (simp)
+apply    (simp)
+apply    (subgoal_tac
+             "tprg\<turnstile>(Ext,mdecl Ext_foo) in Ext dyn_accessible_from Main")
+apply      (simp add: check_method_access_def Let_def
+                      invocation_declclass_def dynlookup_def dynmethd_Ext_foo)
+apply      (rule Ext_foo_dyn_accessible)
 apply   (rule eval_Is (* Methd *))
 apply   (simp add: body_def Let_def)
 apply   (rule eval_Is (* Body *))
@@ -1216,7 +1263,8 @@
 apply       (rule eval_Is (* Acc *))
 apply       (rule eval_Is (* LVar *))
 apply      (simp)
-apply     (simp split del: split_if)
+apply      (simp split del: split_if)
+apply      (simp add: check_field_access_def Let_def)
 apply    (rule eval_Is (* XcptE *))
 apply   (simp)
       (* end method call *)
@@ -1239,6 +1287,7 @@
 apply      (rule init_done, simp)
 apply     (rule eval_Is (* StatRef *))
 apply    (simp)
+apply    (simp add: check_field_access_def Let_def)
 apply   (rule eval_Is (* Lit *))
 apply  (simp (no_asm_simp))
 apply (auto simp add: in_bounds_def)
--- a/src/HOL/Bali/State.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/State.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -22,8 +22,10 @@
 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 *)
+     (* | CStat qtname     the tag is irrelevant for a class object,
+			   i.e. the static fields of a class,
+                           since its type is given already by the reference to 
+                           it (see below) *)
 
 types	vn   = "fspec + int"                    (* variable name      *)
 record	obj  = 
@@ -489,9 +491,14 @@
 	= Loc loc    (* location of allocated execption object *)
 	| Std xname  (* intermediate standard exception, see Eval.thy *)
 
+datatype error
+       = AccessViolation (* Access to a member that isn't permitted *)
+
 datatype abrupt      (* abrupt completion *) 
         = Xcpt xcpt  (* exception *)
         | Jump jump  (* break, continue, return *)
+        | Error error (* runtime errors, we wan't to detect and proof absent
+                         in welltyped programms *)
 consts
 
   the_Xcpt :: "abrupt \<Rightarrow> xcpt"
@@ -542,12 +549,14 @@
   raise_if :: "bool \<Rightarrow> xname \<Rightarrow> abopt \<Rightarrow> abopt"
   np       :: "val  \<spacespace>        \<Rightarrow> abopt \<Rightarrow> abopt"
   check_neg:: "val  \<spacespace>        \<Rightarrow> abopt \<Rightarrow> abopt"
+  error_if :: "bool \<Rightarrow> error \<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"
+ "error_if c e"  == "abrupt_if c (Some (Error e))"
 
 lemma raise_if_None [simp]: "(raise_if c x y = None) = (\<not>c \<and> y = None)"
 apply (simp add: abrupt_if_def)
@@ -569,6 +578,26 @@
 apply auto
 done
 
+lemma error_if_None [simp]: "(error_if c e y = None) = (\<not>c \<and> y = None)"
+apply (simp add: abrupt_if_def)
+by auto
+declare error_if_None [THEN iffD1, dest!]
+
+lemma if_error_if_None [simp]: 
+  "((if b then y else error_if c e y) = None) = ((c \<longrightarrow> b) \<and> y = None)"
+apply (simp add: abrupt_if_def)
+apply auto
+done
+
+lemma raise_if_SomeD [dest!]:
+  "error_if c e y = Some z \<Longrightarrow> c \<and> z=(Error e) \<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"
@@ -733,7 +762,97 @@
 apply (simp (no_asm))
 done
 
+section {* @{text error_free} *}
+constdefs error_free:: "state \<Rightarrow> bool"
+"error_free s \<equiv> \<not> (\<exists> err. abrupt s = Some (Error err))"
 
+lemma error_free_Norm [simp,intro]: "error_free (Norm s)"
+by (simp add: error_free_def)
+
+lemma error_free_normal [simp,intro]: "normal s \<Longrightarrow> error_free s"
+by (simp add: error_free_def)
+
+lemma error_free_Xcpt [simp]: "error_free (Some (Xcpt x),s)"
+by (simp add: error_free_def)
+
+lemma error_free_Jump [simp,intro]: "error_free (Some (Jump j),s)"
+by (simp add: error_free_def)
+
+lemma error_free_Error [simp]: "error_free (Some (Error e),s) = False"
+by (simp add: error_free_def)  
+
+lemma error_free_Some [simp,intro]: 
+ "\<not> (\<exists> err. x=Error err) \<Longrightarrow> error_free ((Some x),s)"
+by (auto simp add: error_free_def)
+
+lemma error_free_absorb [simp,intro]: 
+ "error_free s \<Longrightarrow> error_free (abupd (absorb j) s)"
+by (cases s) 
+   (auto simp add: error_free_def absorb_def
+         split: split_if_asm)
+
+lemma error_free_absorb [simp,intro]: 
+ "error_free (a,s) \<Longrightarrow> error_free (absorb j a, s)"
+by (auto simp add: error_free_def absorb_def
+            split: split_if_asm)
+
+lemma error_free_abrupt_if [simp,intro]:
+"\<lbrakk>error_free s; \<not> (\<exists> err. x=Error err)\<rbrakk>
+ \<Longrightarrow> error_free (abupd (abrupt_if p (Some x)) s)"
+by (cases s)
+   (auto simp add: abrupt_if_def
+            split: split_if)
+
+lemma error_free_abrupt_if1 [simp,intro]:
+"\<lbrakk>error_free (a,s); \<not> (\<exists> err. x=Error err)\<rbrakk>
+ \<Longrightarrow> error_free (abrupt_if p (Some x) a, s)"
+by  (auto simp add: abrupt_if_def
+            split: split_if)
+
+lemma error_free_abrupt_if_Xcpt [simp,intro]:
+ "error_free s 
+  \<Longrightarrow> error_free (abupd (abrupt_if p (Some (Xcpt x))) s)"
+by simp 
+
+lemma error_free_abrupt_if_Xcpt1 [simp,intro]:
+ "error_free (a,s) 
+  \<Longrightarrow> error_free (abrupt_if p (Some (Xcpt x)) a, s)" 
+by simp 
+
+lemma error_free_abrupt_if_Jump [simp,intro]:
+ "error_free s 
+  \<Longrightarrow> error_free (abupd (abrupt_if p (Some (Jump j))) s)" 
+by simp
+
+lemma error_free_abrupt_if_Jump1 [simp,intro]:
+ "error_free (a,s) 
+  \<Longrightarrow> error_free (abrupt_if p (Some (Jump j)) a, s)" 
+by simp
+
+lemma error_free_raise_if [simp,intro]:
+ "error_free s \<Longrightarrow> error_free (abupd (raise_if p x) s)"
+by simp 
+
+lemma error_free_raise_if1 [simp,intro]:
+ "error_free (a,s) \<Longrightarrow> error_free ((raise_if p x a), s)"
+by simp 
+
+lemma error_free_supd [simp,intro]:
+ "error_free s \<Longrightarrow> error_free (supd f s)"
+by (cases s) (simp add: error_free_def)
+
+lemma error_free_supd1 [simp,intro]:
+ "error_free (a,s) \<Longrightarrow> error_free (a,f s)"
+by (simp add: error_free_def)
+
+lemma error_free_set_lvars [simp,intro]:
+"error_free s \<Longrightarrow> error_free ((set_lvars l) s)"
+by (cases s) simp
+
+lemma error_free_set_locals [simp,intro]: 
+"error_free (x, s)
+       \<Longrightarrow> error_free (x, set_locals l s')"
+by (simp add: error_free_def)
 
 end
 
--- a/src/HOL/Bali/Table.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/Table.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -31,7 +31,6 @@
 \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 *)
@@ -148,6 +147,10 @@
  "\<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_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 filter_tab_None: "t k = None \<Longrightarrow> filter_tab p t k = None"
 apply (simp add: filter_tab_def expand_fun_eq)
 done
@@ -180,6 +183,7 @@
     = 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"
--- a/src/HOL/Bali/Term.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/Term.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -78,7 +78,8 @@
 
 datatype var
 	= LVar                  lname(* local variable (incl. parameters) *)
-        | FVar qtname bool expr vname(*class field*)("{_,_}_.._"[10,10,85,99]90)
+        | FVar qtname qtname bool expr vname
+                                (*class field*)("{_,_,_}_.._"[10,10,10,85,99]90)
 	| AVar        expr expr      (* array component *) ("_.[_]"[90,10   ]90)
 
 and expr
@@ -91,8 +92,8 @@
 	| 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)
+        | Call qtname ref_ty inv_mode expr mname "(ty list)" (* method call *)
+                  "(expr list)" ("{_,_,_}_\<cdot>_'( {_}_')"[10,10,10,85,99,10,10]85)
         | Methd qtname sig          (*   (folded) method (see below) *)
         | Body qtname stmt          (* (unfolded) method body *)
 and  stmt
@@ -157,75 +158,4 @@
 *}
 
 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
--- a/src/HOL/Bali/Trans.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/Trans.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -8,118 +8,153 @@
 
 PRELIMINARY!!!!!!!!
 
+improvements over Java Specification 1.0 (cf. 15.11.4.4):
+* dynamic method lookup does not need to check the return type
+* throw raises a NullPointer exception if a null reference is given, and each
+  throw of a system exception yield a fresh exception object (was not specified)
+* if there is not enough memory even to allocate an OutOfMemory exception,
+  evaluation/execution fails, i.e. simply stops (was not specified)
+
+design issues:
+* Lit expressions and Skip statements are considered completely evaluated.
+* the expr entry in rules is redundant in case of exceptions, but its full
+  inclusion helps to make the rule structure independent of exception occurence.
+* the rule format is such that the start state may contain an exception.
+  ++ faciliates exception handling (to be added later)
+  +  symmetry
+* the rules are defined carefully in order to be applicable even in not
+  type-correct situations (yielding undefined values),
+  e.g. the_Adr (Val (Bool b)) = arbitrary.
+  ++ fewer rules 
+  -  less readable because of auxiliary functions like the_Adr
+  Alternative: "defensive" evaluation throwing some InternalError exception
+               in case of (impossible, for correct programs) type mismatches
+
+simplifications:
+* just simple handling (i.e. propagation) of exceptions so far
+* dynamic method lookup does not check return type (should not be necessary)
 *)
 
 Trans = Eval +
 
 consts
-  texpr_tstmt	:: "prog \<Rightarrow> (((expr \<times> state) \<times> (expr \<times> state)) +
-		            ((stmt \<times> state) \<times> (stmt \<times> state))) set"
+  texpr_tstmt	:: "prog ë (((expr ò state) ò (expr ò state)) +
+		            ((stmt ò state) ò (stmt ò 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"
+  texpr :: "[prog, expr ò state, expr ò state] ë bool "("_É_ è1 _"[61,82,82] 81)
+  tstmt :: "[prog, stmt ò state, stmt ò state] ë bool "("_É_ í1 _"[61,82,82] 81)
+  Ref   :: "loc ë expr"
 
 translations
 
-  "G\<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"
+  "GÉe_s è1 ex_s'" == "Inl (e_s, ex_s') Î texpr_tstmt G"
+  "GÉs_s í1 s'_s'" == "Inr (s_s, s'_s') Î texpr_tstmt G"
   "Ref a" == "Lit (Addr a)"
 
+constdefs
+  
+  sub_expr_expr :: "(expr ë expr) ë prop"
+  "sub_expr_expr ef Ú (ÄG e s e' s'. GÉ(   e,s) è1 (   e',s') êë
+				     GÉ(ef e,s) è1 (ef e',s'))"
+
 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)"
+  XcptE	"ËÂv. e Û Lit vÌ êë
+				  GÉ(e,Some xc,s) è1 (Lit arbitrary,Some xc,s)"
 
+ CastXX "PROP sub_expr_expr (Cast T)"
+
+(*
   (* 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')"
+  NewC	"Ënew_Addr (heap s) = Some (a,x);
+	  s' = assign (hupd[aíinit_Obj G C]s) (x,s)Ì êë
+				GÉ(NewC C,None,s) è1 (Ref a,s')"
 
   (* cf. 15.9.1 *)
-  NewA1	"\<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')"
+(*NewA1	"sub_expr_expr (NewA T)"*)
+  NewA1	"ËGÉ(e,None,s) è1 (e',s')Ì êë
+			      GÉ(New T[e],None,s) è1 (New T[e'],s')"
+  NewA	"Ëi = the_Intg i'; new_Addr (heap s) = Some (a, x);
+	  s' = assign (hupd[aíinit_Arr T i]s)(raise_if (i<#0) NegArrSize x,s)Ì êë
+			GÉ(New T[Lit i'],None,s) è1 (Ref a,s')"
   (* cf. 15.15 *)
-  Cast1	"\<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)"
+  Cast1	"ËGÉ(e,None,s) è1 (e',s')Ì êë
+			      GÉ(Cast T e,None,s) è1 (Cast T e',s')"
+  Cast	"Ëx'= raise_if (\<questiondown>G,sÉv fits T) ClassCast NoneÌ êë
+		        GÉ(Cast T (Lit v),None,s) è1 (Lit v,x',s)"
 
   (* cf. 15.7.1 *)
-(*Lit				"G\<turnstile>(Lit v,None,s) \<rightarrow>1 (Lit v,None,s)"*)
+(*Lit				"GÉ(Lit v,None,s) è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)"
+  LAcc	"Ëv = the (locals s vn)Ì êë
+			       GÉ(LAcc vn,None,s) è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)"
+  LAss1	"ËGÉ(e,None,s) è1 (e',s')Ì êë
+				 GÉ(f vn:=e,None,s) è1 (vn:=e',s')"
+  LAss			    "GÉ(f vn:=Lit v,None,s) è1 (Lit v,None,lupd[vnív]s)"
 
   (* cf. 15.10.1, 15.2 *)
-  FAcc1	"\<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)"
+  FAcc1	"ËGÉ(e,None,s) è1 (e',s')Ì êë
+			       GÉ({T}e..fn,None,s) è1 ({T}e'..fn,s')"
+  FAcc	"Ëv = the (snd (the_Obj (heap s (the_Addr a'))) (fn,T))Ì êë
+			  GÉ({T}Lit a'..fn,None,s) è1 (Lit v,np a' None,s)"
 
   (* cf. 15.25.1 *)
-  FAss1	"\<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')"
+  FAss1	"ËGÉ(e1,None,s) è1 (e1',s')Ì êë
+			  GÉ(f ({T}e1..fn):=e2,None,s) è1 (f({T}e1'..fn):=e2,s')"
+  FAss2	"ËGÉ(e2,np a' None,s) è1 (e2',s')Ì êë
+		      GÉ(f({T}Lit a'..fn):=e2,None,s) è1 (f({T}Lit a'..fn):=e2',s')"
+  FAss	"Ëa = the_Addr a'; (c,fs) = the_Obj (heap s a);
+	  s'= assign (hupd[aíObj c (fs[(fn,T)ív])]s) (None,s)Ì êë
+		   GÉ(f({T}Lit a'..fn):=Lit v,None,s) è1 (Lit v,s')"
 
 
 
 
 
   (* cf. 15.12.1 *)
-  AAcc1	"\<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)"
+  AAcc1	"ËGÉ(e1,None,s) è1 (e1',s')Ì êë
+				GÉ(e1[e2],None,s) è1 (e1'[e2],s')"
+  AAcc2	"ËGÉ(e2,None,s) è1 (e2',s')Ì êë
+			    GÉ(Lit a'[e2],None,s) è1 (Lit a'[e2'],s')"
+  AAcc	"Ëvo = snd (the_Arr (heap s (the_Addr a'))) (the_Intg i');
+	  x' = raise_if (vo = None) IndOutBound (np a' None)Ì êë
+			GÉ(Lit a'[Lit i'],None,s) è1 (Lit (the vo),x',s)"
 
 
   (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *)
-  Call1	"\<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))"
+  Call1	"ËGÉ(e,None,s) è1 (e',s')Ì êë
+			  GÉ(e..mn({pT}p),None,s) è1 (e'..mn({pT}p),s')"
+  Call2	"ËGÉ(p,None,s) è1 (p',s')Ì êë
+		     GÉ(Lit a'..mn({pT}p),None,s) è1 (Lit a'..mn({pT}p'),s')"
+  Call	"Ëa = the_Addr a'; (md,(pn,rT),lvars,blk,res) = 
+ 			   the (cmethd G (fst (the_Obj (h a))) (mn,pT))Ì êë
+	    GÉ(Lit a'..mn({pT}Lit pv),None,(h,l)) è1 
+      (Body blk res l,np a' x,(h,init_vals lvars[Thisía'][Supería'][pnípv]))"
+  Body1	"ËGÉ(s0,None,s) í1 (s0',s')Ì êë
+		   GÉ(Body s0    e      l,None,s) è1 (Body s0'  e  l,s')"
+  Body2	"ËGÉ(e ,None,s) è1 (e',s')Ì êë
+		   GÉ(Body Skip  e      l,None,s) è1 (Body Skip e' l,s')"
+  Body		  "GÉ(Body Skip (Lit v) l,None,s) è1 (Lit v,None,(heap s,l))"
 
 (* execution of statements *)
 
   (* cf. 14.1 *)
-  XcptS	"\<lbrakk>s0 \<noteq> Skip\<rbrakk> \\<Longrightarrow>
-				 G\<turnstile>(s0,Some xc,s) \<mapsto>1 (Skip,Some xc,s)"
+  XcptS	"Ës0 Û SkipÌ êë
+				 GÉ(s0,Some xc,s) í1 (Skip,Some xc,s)"
 
   (* cf. 14.5 *)
-(*Skip	 			 "G\<turnstile>(Skip,None,s) \<mapsto>1 (Skip,None,s)"*)
+(*Skip	 			 "GÉ(Skip,None,s) í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)"
+  Comp1	"ËGÉ(s1,None,s) í1 (s1',s')Ì êë
+			       GÉ(s1;; s2,None,s) í1 (s1';; s2,s')"
+  Comp			    "GÉ(Skip;; s2,None,s) í1 (s2,None,s)"
 
 
 
@@ -127,18 +162,20 @@
 
 
   (* 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)"
+  Expr1	"ËGÉ(e ,None,s) è1 (e',s')Ì êë
+				GÉ(Expr e,None,s) í1 (Expr e',s')"
+  Expr			 "GÉ(Expr (Lit v),None,s) í1 (Skip,None,s)"
 
   (* cf. 14.8.2 *)
-  If1	"\<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 
+  If1	"ËGÉ(e ,None,s) è1 (e',s')Ì êë
+		      GÉ(If(e) s1 Else s2,None,s) í1 (If(e') s1 Else s2,s')"
+  If		 "GÉ(If(Lit v) s1 Else s2,None,s) í1 
 		    (if the_Bool v then s1 else s2,None,s)"
 
   (* cf. 14.10, 14.10.1 *)
-  Loop			  "G\<turnstile>(While(e) s0,None,s) \<mapsto>1 
+  Loop			  "GÉ(While(e) s0,None,s) í1 
 			     (If(e) (s0;; While(e) s0) Else Skip,None,s)"
+*)
+  con_defs "[sub_expr_expr_def]"
 
 end
--- a/src/HOL/Bali/TypeSafe.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/TypeSafe.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -1,20 +1,105 @@
 (*  Title:      HOL/Bali/TypeSafe.thy
     ID:         $Id$
-    Author:     David von Oheimb
+    Author:     David von Oheimb and Norbert Schirmer
     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 *)
 header {* The type soundness proof for Java *}
 
+theory TypeSafe = Eval + WellForm + Conform:
 
-theory TypeSafe = Eval + WellForm + Conform:
+section "error free"
+ 
+lemma error_free_halloc:
+ (assumes halloc: "G\<turnstile>s0 \<midarrow>halloc oi\<succ>a\<rightarrow> s1" and
+          error_free_s0: "error_free s0"
+ ) "error_free s1"
+proof -
+  from halloc error_free_s0
+  obtain abrupt0 store0 abrupt1 store1
+    where eqs: "s0=(abrupt0,store0)" "s1=(abrupt1,store1)" and
+          halloc': "G\<turnstile>(abrupt0,store0) \<midarrow>halloc oi\<succ>a\<rightarrow> (abrupt1,store1)" and
+          error_free_s0': "error_free (abrupt0,store0)"
+    by (cases s0,cases s1) auto
+  from halloc' error_free_s0'
+  have "error_free (abrupt1,store1)"
+  proof (induct)
+    case Abrupt 
+    then show ?case
+      .
+  next
+    case New
+    then show ?case
+      by (auto split: split_if_asm)
+  qed
+  with eqs 
+  show ?thesis
+    by simp
+qed
+
+lemma error_free_sxalloc:
+ (assumes sxalloc: "G\<turnstile>s0 \<midarrow>sxalloc\<rightarrow> s1" and error_free_s0: "error_free s0") 
+ "error_free s1"
+proof -
+  from sxalloc error_free_s0
+  obtain abrupt0 store0 abrupt1 store1
+    where eqs: "s0=(abrupt0,store0)" "s1=(abrupt1,store1)" and
+          sxalloc': "G\<turnstile>(abrupt0,store0) \<midarrow>sxalloc\<rightarrow> (abrupt1,store1)" and
+          error_free_s0': "error_free (abrupt0,store0)"
+    by (cases s0,cases s1) auto
+  from sxalloc' error_free_s0'
+  have "error_free (abrupt1,store1)"
+  proof (induct)
+  qed (auto)
+  with eqs 
+  show ?thesis 
+    by simp
+qed
+
+lemma error_free_check_field_access_eq:
+ "error_free (check_field_access G accC statDeclC fn stat a s)
+ \<Longrightarrow> (check_field_access G accC statDeclC fn stat a s) = s"
+apply (cases s)
+apply (auto simp add: check_field_access_def Let_def error_free_def 
+                      abrupt_if_def 
+            split: split_if_asm)
+done
+
+lemma error_free_check_method_access_eq:
+"error_free (check_method_access G accC statT mode sig a' s)
+ \<Longrightarrow> (check_method_access G accC statT mode sig a' s) = s"
+apply (cases s)
+apply (auto simp add: check_method_access_def Let_def error_free_def 
+                      abrupt_if_def 
+            split: split_if_asm)
+done
+
+lemma error_free_FVar_lemma: 
+     "error_free s 
+       \<Longrightarrow> error_free (abupd (if stat then id else np a) s)"
+  by (case_tac s) (auto split: split_if) 
+
+lemma error_free_init_lvars [simp,intro]:
+"error_free s \<Longrightarrow> 
+  error_free (init_lvars G C sig mode a pvs s)"
+by (cases s) (auto simp add: init_lvars_def Let_def split: split_if)
+
+lemma error_free_LVar_lemma:   
+"error_free s \<Longrightarrow> error_free (assign (\<lambda>v. supd lupd(vn\<mapsto>v)) w s)"
+by (cases s) simp
+
+lemma error_free_throw [simp,intro]:
+  "error_free s \<Longrightarrow> error_free (abupd (throw x) s)"
+by (cases s) (simp add: throw_def)
+
 
 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"
+"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) \<and>
+ (\<forall>s' w. error_free s' \<longrightarrow> (error_free (assign f w s')))"      
 
   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)
@@ -91,7 +176,8 @@
 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
+ simp  add: lvar_def fvar_def2 avar_def2 init_lvars_def2 
+            check_field_access_def check_method_access_def Let_def
  split del: split_if_asm split add: sum3.split)
 (* 6 subgoals *)
 apply force+
@@ -162,7 +248,7 @@
 
 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)"
+  (if is_static m then x else (np a') x)"
 apply (simp (no_asm) add: init_lvars_def2)
 done
 
@@ -175,7 +261,8 @@
        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);
+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)
@@ -348,38 +435,34 @@
     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; 
+"\<lbrakk>G\<turnstile>invmode 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
+  (statDeclT,sm) \<in> mheads G C statT sig; 
+  invC = invocation_class (invmode sm e) s a' statT;
+  declC =invocation_declclass G (invmode 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> 
+  methd G declC sig = Some dm \<and> dynlookup G statT invC sig = Some dm  \<and> 
+  G\<turnstile>resTy (mthd dm)\<preceq>resTy 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
+  (if invmode 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))"
+            statDeclT = ClassT (declclass dm))"
 proof -
-  assume invC_prop: "G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT" 
+  assume invC_prop: "G\<turnstile>invmode 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        sm: "(statDeclT,sm) \<in> mheads G C statT sig" 
+     and      invC: "invC = invocation_class (invmode sm e) s a' statT"
      and     declC: "declC = 
-                    invocation_declclass G (invmode (mhd sm) e) s a' statT sig"
+                    invocation_declclass G (invmode 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
@@ -388,13 +471,13 @@
     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)"
+                                        invmode 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)"
+  have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode sm e \<noteq> SuperM)"
     by (auto dest: invocationTypeExpr_noClassD)
   show ?thesis
-  proof (cases "invmode (mhd sm) e = IntVir")
+  proof (cases "invmode sm e = IntVir")
     case True
     with invC_prop not_Null
     have invC_prop': " is_class G invC \<and> 
@@ -403,15 +486,15 @@
       by (auto simp add: DynT_prop_def) 
     from True 
     have "\<not> is_static sm"
-      by (simp add: invmode_IntVir_eq)
+      by (simp add: invmode_IntVir_eq member_is_static_simp)
     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
+      resT_dm: "G\<turnstile>resTy (mthd dm)\<preceq>resTy sm"      and
        static: "is_static dm = is_static sm"
-      by - (drule dynamic_mheadsD,auto)
+      by  - (drule dynamic_mheadsD,force+)
     with declC invC not_Null 
     have declC': "declC = (declclass dm)" 
       by (auto simp add: invocation_declclass_def)
@@ -428,13 +511,14 @@
     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
+         dm
     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)) "
+                      (\<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
@@ -443,16 +527,19 @@
                                             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)
+	eq_declC_sm_dm:"statDeclT = ClassT (declclass dm)"  and
+	     eq_mheads:"sm=mhead (mthd dm) "
+      by - (drule static_mheadsD, (force dest: accmethd_SomeD)+)
+    then have static: "is_static dm = is_static sm" by - (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 dynlookup_static dm 
+    have dm'': "dynlookup G statT invC sig = Some dm"
+      by simp
     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 )
@@ -464,126 +551,11 @@
     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' 
+    from False dm' 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"
-*}
+qed   
 
 lemma DynT_conf: "\<lbrakk>G\<turnstile>invocation_class mode s a' statT \<preceq>\<^sub>C declC; wf_prog G;
  isrtype G (statT);
@@ -604,12 +576,11 @@
 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))"
+lemma Ass_lemma:
+"\<lbrakk> G\<turnstile>Norm s0 \<midarrow>var=\<succ>(w, f)\<rightarrow> Norm s1; G\<turnstile>Norm s1 \<midarrow>e-\<succ>v\<rightarrow> Norm s2;
+   G,s2\<turnstile>v\<Colon>\<preceq>eT;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>
+      (normal (assign f v (Norm s2)) \<longrightarrow> G,store (assign f v (Norm s2))\<turnstile>v\<Colon>\<preceq>eT)"
 apply (drule_tac x = "None" and s = "s2" and v = "v" in evar_gext_f)
 apply (drule eval_gext', clarsimp)
 apply (erule conf_gext)
@@ -639,14 +610,18 @@
 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); 
+lemma FVar_lemma1: 
+"\<lbrakk>table_of (DeclConcepts.fields G statC) (fn, statDeclC) = Some f ; 
+  x2 = None \<longrightarrow> G,s2\<turnstile>a\<Colon>\<preceq> Class statC; wf_prog G; G\<turnstile>statC\<preceq>\<^sub>C statDeclC; 
+  statDeclC \<noteq> Object; 
+  class G statDeclC = Some y; (x2,s2)\<Colon>\<preceq>(G, L); s1\<le>|s2; 
+  inited statDeclC (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)"
+  \<exists>obj. globs s2 (if static f then Inr statDeclC else Inl (the_Addr a)) 
+                  = Some obj \<and> 
+  var_tys G (tag obj)  (if static f then Inr statDeclC else Inl(the_Addr a)) 
+          (Inl(fn,statDeclC)) = Some (type f)"
 apply (drule initedD)
 apply (frule subcls_is_class2, simp (no_asm_simp))
 apply (case_tac "static f")
@@ -665,11 +640,39 @@
 apply (auto elim!: fields_mono subcls_is_class2)
 done
 
+lemma FVar_lemma2: "error_free state
+       \<Longrightarrow> error_free
+           (assign
+             (\<lambda>v. supd
+                   (upd_gobj
+                     (if static field then Inr statDeclC
+                      else Inl (the_Addr a))
+                     (Inl (fn, statDeclC)) v))
+             w state)"
+proof -
+  assume error_free: "error_free state"
+  obtain a s where "state=(a,s)"
+    by (cases state) simp
+  with error_free
+  show ?thesis
+    by (cases a) auto
+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 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>  
+"\<lbrakk>((v, f), Norm s2') = fvar statDeclC (static field) fn a (x2, s2); 
+  G\<turnstile>statC\<preceq>\<^sub>C statDeclC;  
+  table_of (DeclConcepts.fields G statC) (fn, statDeclC) = Some field; 
+  wf_prog G;   
+  x2 = None \<longrightarrow> G,s2\<turnstile>a\<Colon>\<preceq>Class statC; 
+  statDeclC \<noteq> Object; class G statDeclC = Some y;   
+  (x2, s2)\<Colon>\<preceq>(G, L); s1\<le>|s2; inited statDeclC (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)
@@ -678,9 +681,20 @@
 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)
+apply (rule conjI)
+apply   clarsimp
+apply   (drule (1) rev_gext_objD)
+apply   (force elim!: conforms_upd_gobj)
+
+apply   (blast intro: FVar_lemma2)
 done
+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 {*
+claset_ref()  := claset() addSbefore ("split_all_tac", split_all_tac);
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
+*}
 
 
 lemma AVar_lemma1: "\<lbrakk>globs s (Inl a) = Some obj;tag obj=Arr ty i; 
@@ -700,6 +714,22 @@
     by auto
 qed
  
+lemma AVar_lemma2: "error_free state 
+       \<Longrightarrow> error_free
+           (assign
+             (\<lambda>v (x, s').
+                 ((raise_if (\<not> G,s'\<turnstile>v fits T) ArrStore) x,
+                  upd_gobj (Inl a) (Inr (the_Intg i)) v s'))
+             w state)"
+proof -
+  assume error_free: "error_free state"
+  obtain a s where "state=(a,s)"
+    by (cases state) simp
+  with error_free
+  show ?thesis
+    by (cases a) 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)"
@@ -714,14 +744,14 @@
 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)
+apply   (force dest: AVar_lemma1)
+
+apply   (force elim!: fits_Array dest: gext_objD 
+         intro: var_tys_Some_eq [THEN iffD2] conforms_upd_gobj)
 done
 
+section "Call"
 
-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)); 
@@ -763,7 +793,13 @@
 apply (auto intro: prim_ty.induct)
 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"
+*}
 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);  
@@ -789,7 +825,7 @@
                                   \<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)) 
+                 | This \<Rightarrow> if (is_static (mthd sm)) 
                               then None else Some (Class declC)))"
 apply (simp add: init_lvars_def2)
 apply (rule conforms_set_locals)
@@ -806,103 +842,58 @@
 apply (case_tac "is_static sm")
 apply simp_all
 done
-
+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 {*
+claset_ref()  := claset() addSbefore ("split_all_tac", split_all_tac);
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
+*}
 
-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"
 
+
+(* #### stat raus und gleich is_static f schreiben *) 
 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
+     not_Null: "\<not> stat \<longrightarrow> a\<noteq>Null" and
+    conform_a: "G,(store s)\<turnstile>a\<Colon>\<preceq> Class statC" and
+    conform_s: "s\<Colon>\<preceq>(G, L)" and 
+     normal_s: "normal s" and
+         wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<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)"
+         dynC: "if stat then dynC=declclass f  
+                        else dynC=obj_class (lookup_obj (store s) a)" and
+         stat: "if stat then (is_static f) else (\<not> is_static f)"
   ) "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 stat have static: "(is_static f)" by simp
+  from True dynC 
+  have dynC': "dynC=declclass f" by simp
   with f
-  have "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)"
+  have "table_of (DeclConcepts.fields G statC) (fn,declclass f) = Some (fld f)"
     by (auto simp add: accfield_def Let_def intro!: table_of_remap_SomeD)
-  with dynC' f
+  moreover
+  from wt_e wf have "is_class G statC"
+    by (auto dest!: ty_expr_is_type)
+  moreover note wf dynC'
+  ultimately have
+     "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)"
+    by (auto dest: fields_declC)
+  with dynC' f static wf
   show ?thesis
-    by (auto intro!: static_to_dynamic_accessible_from
-         dest: accfield_accessibleD accessible_from_commonD)
+    by (auto dest: static_to_dynamic_accessible_from_static
+            dest!: accfield_accessibleD )
 next
   case False
-  with wf conform_a not_Null conform_s2 dynC
+  with wf conform_a not_Null conform_s dynC
   obtain subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
     "is_class G dynC"
-    by (auto dest!: conforms_RefTD [of _ _ _ _ "(fst s2)" L]
+    by (auto dest!: conforms_RefTD [of _ _ _ _ "(fst s)" L]
               dest: obj_ty_obj_class1
           simp add: obj_ty_obj_class )
   with wf f
@@ -919,12 +910,167 @@
     by blast
 qed
 
-lemma call_access_ok: 
-(assumes invC_prop: "G\<turnstile>invmode (mhd statM) e\<rightarrow>invC\<preceq>statT" 
+(*
+theorem dynamic_field_access_ok:
+  (assumes wf: "wf_prog G" and
+     not_Null: "\<not> is_static f \<longrightarrow> a\<noteq>Null" and
+    conform_a: "G,(store s)\<turnstile>a\<Colon>\<preceq> Class statC" and
+    conform_s: "s\<Colon>\<preceq>(G, L)" and 
+     normal_s: "normal s" and
+         wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-Class statC" and
+            f: "accfield G accC statC fn = Some f" and
+         dynC: "if is_static f 
+                   then dynC=declclass f  
+                   else dynC=obj_class (lookup_obj (store s) 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 "is_static f")
+  case True
+  from True dynC 
+  have dynC': "dynC=declclass f" by simp
+  with f
+  have "table_of (DeclConcepts.fields G statC) (fn,declclass f) = Some (fld f)"
+    by (auto simp add: accfield_def Let_def intro!: table_of_remap_SomeD)
+  moreover
+  from wt_e wf have "is_class G statC"
+    by (auto dest!: ty_expr_is_type)
+  moreover note wf dynC'
+  ultimately have
+     "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)"
+    by (auto dest: fields_declC)
+  with dynC' f True wf
+  show ?thesis
+    by (auto dest: static_to_dynamic_accessible_from_static
+            dest!: accfield_accessibleD )
+next
+  case False
+  with wf conform_a not_Null conform_s dynC
+  obtain subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
+    "is_class G dynC"
+    by (auto dest!: conforms_RefTD [of _ _ _ _ "(fst s)" 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
+*)
+
+
+(* ### Einsetzen in case FVar des TypeSoundness Proofs *)
+(*
+lemma FVar_check_error_free:
+(assumes fvar: "(v, s2') = fvar statDeclC stat fn a s2" and 
+        check: "s3 = check_field_access G accC statDeclC fn stat a s2'" and
+       conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC" and
+      conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+    initd_statDeclC_s2: "initd statDeclC s2" and
+    wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-Class statC" and
+    accfield: "accfield G accC statC fn = Some (statDeclC,f)" and
+    stat: "stat=is_static f" and
+      wf: "wf_prog G"
+)  "s3=s2'"
+proof -
+  from fvar 
+  have store_s2': "store s2'=store s2"
+    by (cases s2) (simp add: fvar_def2)
+  with fvar conf_s2 
+  have conf_s2': "s2'\<Colon>\<preceq>(G, L)"
+    by (cases s2,cases stat) (auto simp add: fvar_def2)
+  from initd_statDeclC_s2 store_s2' 
+  have initd_statDeclC_s2': "initd statDeclC s2"
+    by simp
+  show ?thesis
+  proof (cases "normal s2'")
+    case False
+    with check show ?thesis 
+      by (auto simp add: check_field_access_def Let_def)
+  next
+    case True
+    with fvar store_s2' 
+    have not_Null: "\<not> stat \<longrightarrow> a\<noteq>Null" 
+      by (cases s2) (auto simp add: fvar_def2)
+    from True fvar store_s2'
+    have "normal s2"
+      by (cases s2,cases stat) (auto simp add: fvar_def2)
+    with conf_a store_s2'
+    have conf_a': "G,store s2'\<turnstile>a\<Colon>\<preceq>Class statC"
+      by simp 
+    from conf_a' conf_s2'  check True initd_statDeclC_s2' 
+      dynamic_field_access_ok [OF wf not_Null conf_a' conf_s2' 
+                                   True wt_e accfield ] stat
+    show ?thesis
+      by (cases stat)
+         (auto dest!: initedD
+           simp add: check_field_access_def Let_def)
+  qed
+qed
+*)
+
+lemma error_free_field_access:
+ (assumes accfield: "accfield G accC statC fn = Some (statDeclC, f)" and
+              wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-Class statC" and
+         eval_init: "G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1" and
+            eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2" and
+           conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+            conf_a: "normal s2 \<Longrightarrow> G, store s2\<turnstile>a\<Colon>\<preceq>Class statC" and
+              fvar: "(v,s2')=fvar statDeclC (is_static f) fn a s2" and
+                wf: "wf_prog G"   
+ ) "check_field_access G accC statDeclC fn (is_static f) a s2' = s2'"
+proof -
+  from fvar
+  have store_s2': "store s2'=store s2"
+    by (cases s2) (simp add: fvar_def2)
+  with fvar conf_s2 
+  have conf_s2': "s2'\<Colon>\<preceq>(G, L)"
+    by (cases s2,cases "is_static f") (auto simp add: fvar_def2)
+  from eval_init 
+  have initd_statDeclC_s1: "initd statDeclC s1"
+    by (rule init_yields_initd)
+  with eval_e store_s2'
+  have initd_statDeclC_s2': "initd statDeclC s2'"
+    by (auto dest: eval_gext intro: inited_gext)
+  show ?thesis
+  proof (cases "normal s2'")
+    case False
+    then show ?thesis 
+      by (auto simp add: check_field_access_def Let_def)
+  next
+    case True
+    with fvar store_s2' 
+    have not_Null: "\<not> (is_static f) \<longrightarrow> a\<noteq>Null" 
+      by (cases s2) (auto simp add: fvar_def2)
+    from True fvar store_s2'
+    have "normal s2"
+      by (cases s2,cases "is_static f") (auto simp add: fvar_def2)
+    with conf_a store_s2'
+    have conf_a': "G,store s2'\<turnstile>a\<Colon>\<preceq>Class statC"
+      by simp
+    from conf_a' conf_s2' True initd_statDeclC_s2' 
+      dynamic_field_access_ok [OF wf not_Null conf_a' conf_s2' 
+                                   True wt_e accfield ] 
+    show ?thesis
+      by  (cases "is_static f")
+          (auto dest!: initedD
+           simp add: check_field_access_def Let_def)
+  qed
+qed
+
+lemma call_access_ok:
+(assumes invC_prop: "G\<turnstile>invmode 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"
+     and     statM: "(statDeclT,statM) \<in> mheads G accC statT sig" 
+     and      invC: "invC = invocation_class (invmode 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 -
@@ -933,13 +1079,13 @@
   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)"
+                                        invmode 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)"
+  have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode statM e \<noteq> SuperM)"
     by (auto dest: invocationTypeExpr_noClassD)
   show ?thesis
-  proof (cases "invmode (mhd statM) e = IntVir")
+  proof (cases "invmode statM e = IntVir")
     case True
     with invC_prop not_Null
     have invC_prop': "is_class G invC \<and>  
@@ -948,8 +1094,7 @@
       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 *)
+     by (cases statT) (auto simp add: invmode_def) 
     with statM type_statT wf 
     show ?thesis
       by - (rule dynlookup_access,auto)
@@ -970,242 +1115,1227 @@
      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 *)
+     by (cases statT) (auto simp add: invmode_def) 
    with statM type_statT wf 
     show ?thesis
       by - (rule dynlookup_access,auto)
   qed
 qed
 
+lemma error_free_call_access:
+ (assumes     
+   eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2" and
+        wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-(RefT statT)" and  
+       statM: "max_spec G accC statT \<lparr>name = mn, parTs = pTs\<rparr> 
+               = {((statDeclT, statM), pTs')}" and
+     conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+      conf_a: "normal s1 \<Longrightarrow> G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT" and
+     invProp: "normal s3 \<Longrightarrow>
+                G\<turnstile>invmode statM e\<rightarrow>invC\<preceq>statT" and
+          s3: "s3=init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> 
+                        (invmode statM e) a vs s2" and
+        invC: "invC = invocation_class (invmode statM e) (store s2) a statT"and
+    invDeclC: "invDeclC = invocation_declclass G (invmode statM e) (store s2) 
+                             a statT \<lparr>name = mn, parTs = pTs'\<rparr>" and
+          wf: "wf_prog G"
+ )"check_method_access G accC statT (invmode statM e) \<lparr>name=mn,parTs=pTs'\<rparr> a s3
+   = s3"
+proof (cases "normal s2")
+  case False
+  with s3 
+  have "abrupt s3 = abrupt s2"  
+    by (auto simp add: init_lvars_def2)
+  with False
+  show ?thesis
+    by (auto simp add: check_method_access_def Let_def)
+next
+  case True
+  note normal_s2 = True
+  with eval_args
+  have normal_s1: "normal s1"
+    by (cases "normal s1") auto
+  with conf_a eval_args 
+  have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
+    by (auto dest: eval_gext intro: conf_gext)
+  show ?thesis
+  proof (cases "a=Null \<longrightarrow> (is_static statM)")
+    case False
+    then obtain "\<not> is_static statM" "a=Null" 
+      by blast
+    with normal_s2 s3
+    have "abrupt s3 = Some (Xcpt (Std NullPointer))" 
+      by (auto simp add: init_lvars_def2)
+    then show ?thesis
+      by (auto simp add: check_method_access_def Let_def)
+  next
+    case True
+    from statM 
+    obtain
+      statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>" 
+      by (blast dest: max_spec2mheads)
+    from True normal_s2 s3
+    have "normal s3"
+      by (auto simp add: init_lvars_def2)
+    then have "G\<turnstile>invmode statM e\<rightarrow>invC\<preceq>statT"
+      by (rule invProp)
+    with wt_e statM' wf invC
+    obtain dynM where 
+      dynM: "dynlookup G statT invC  \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+      acc_dynM: "G \<turnstile>Methd  \<lparr>name=mn,parTs=pTs'\<rparr> dynM 
+                          in invC dyn_accessible_from accC"
+      by (force dest!: call_access_ok)
+    moreover
+    from s3 invC
+    have invC': "invC=(invocation_class (invmode statM e) (store s3) a statT)"
+      by (cases s2,cases "invmode statM e") 
+         (simp add: init_lvars_def2 del: invmode_Static_eq)+
+    ultimately
+    show ?thesis
+      by (auto simp add: check_method_access_def Let_def)
+  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)"
+lemma eval_type_sound:
+      (assumes eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" and
+                 wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T" and
+                 wf: "wf_prog G" and 
+            conf_s0: "s0\<Colon>\<preceq>(G,L)"           
+      ) "s1\<Colon>\<preceq>(G,L) \<and>  (normal s1 \<longrightarrow> G,L,store s1\<turnstile>t\<succ>v\<Colon>\<preceq>T) \<and> 
+         (error_free s0 = error_free s1)"
+proof -
+  from eval 
+  have "\<And> L accC T. \<lbrakk>s0\<Colon>\<preceq>(G,L);\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T\<rbrakk>  
+        \<Longrightarrow> s1\<Colon>\<preceq>(G,L) \<and> (normal s1 \<longrightarrow> G,L,store s1\<turnstile>t\<succ>v\<Colon>\<preceq>T)
+            \<and> (error_free s0 = error_free s1)"
+   (is "PROP ?TypeSafe s0 s1 t v"
+    is "\<And> L accC T. ?Conform L s0 \<Longrightarrow> ?WellTyped L accC T t  
+                 \<Longrightarrow> ?Conform L s1 \<and> ?ValueTyped L T s1 t v \<and>
+                     ?ErrorFree s0 s1")
+  proof (induct)
+    case (Abrupt s t xc L accC T) 
+    have "(Some xc, s)\<Colon>\<preceq>(G,L)" .
+    then show "(Some xc, s)\<Colon>\<preceq>(G,L) \<and> 
+      (normal (Some xc, s) 
+      \<longrightarrow> G,L,store (Some xc,s)\<turnstile>t\<succ>arbitrary3 t\<Colon>\<preceq>T) \<and> 
+      (error_free (Some xc, s) = error_free (Some xc, s))"
+      by (simp)
+  next
+    case (Skip s L accC T)
+    have "Norm s\<Colon>\<preceq>(G, L)" and  
+           "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r Skip\<Colon>T" .
+    then show "Norm s\<Colon>\<preceq>(G, L) \<and>
+              (normal (Norm s) \<longrightarrow> G,L,store (Norm s)\<turnstile>In1r Skip\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and> 
+              (error_free (Norm s) = error_free (Norm s))"
+      by (simp)
+  next
+    case (Expr e s0 s1 v L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1" .
+    have     hyp: "PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Expr e)\<Colon>T" .
+    then obtain eT 
+      where "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l e\<Colon>eT"
+      by (rule wt_elim_cases) (blast)
+    with conf_s0 hyp 
+    obtain "s1\<Colon>\<preceq>(G, L)" and "error_free s1"
+      by (blast)
+    with wt
+    show "s1\<Colon>\<preceq>(G, L) \<and>
+          (normal s1 \<longrightarrow> G,L,store s1\<turnstile>In1r (Expr e)\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and> 
+          (error_free (Norm s0) = error_free s1)"
+      by (simp)
+  next
+    case (Lab c l s0 s1 L accC T)
+    have     hyp: "PROP ?TypeSafe (Norm s0) s1 (In1r c) \<diamondsuit>" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (l\<bullet> c)\<Colon>T" .
+    then have "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
+      by (rule wt_elim_cases) (blast)
+    with conf_s0 hyp
+    obtain       conf_s1: "s1\<Colon>\<preceq>(G, L)" and 
+           error_free_s1: "error_free s1" 
+      by (blast)
+    from conf_s1 have "abupd (absorb (Break l)) s1\<Colon>\<preceq>(G, L)"
+      by (cases s1) (auto intro: conforms_absorb)
+    with wt error_free_s1
+    show "abupd (absorb (Break l)) s1\<Colon>\<preceq>(G, L) \<and>
+          (normal (abupd (absorb (Break l)) s1)
+           \<longrightarrow> G,L,store (abupd (absorb (Break l)) s1)\<turnstile>In1r (l\<bullet> c)\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and>
+          (error_free (Norm s0) = error_free (abupd (absorb (Break l)) s1))"
+      by (simp)
+  next
+    case (Comp c1 c2 s0 s1 s2 L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1" .
+    have "G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2" .
+    have  hyp_c1: "PROP ?TypeSafe (Norm s0) s1 (In1r c1) \<diamondsuit>" .
+    have  hyp_c2: "PROP ?TypeSafe s1        s2 (In1r c2) \<diamondsuit>" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (c1;; c2)\<Colon>T" .
+    then obtain wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+                wt_c2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c2\<Colon>\<surd>"
+      by (rule wt_elim_cases) (blast)
+    with conf_s0 hyp_c1 hyp_c2
+    obtain "s2\<Colon>\<preceq>(G, L)" and "error_free s2"
+      by (blast)
+    with wt
+    show "s2\<Colon>\<preceq>(G, L) \<and>
+          (normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1r (c1;; c2)\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and>
+          (error_free (Norm s0) = error_free s2)"
+      by (simp)
+  next
+    case (If b c1 c2 e s0 s1 s2 L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1" .
+    have "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2" .
+    have hyp_e: "PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 b)" .
+    have hyp_then_else: 
+            "PROP ?TypeSafe s1 s2 (In1r (if the_Bool b then c1 else c2)) \<diamondsuit>" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (If(e) c1 Else c2)\<Colon>T" .
+    then obtain "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean"
+                "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
+      by (rule wt_elim_cases) (auto split add: split_if)
+    with conf_s0 hyp_e hyp_then_else
+    obtain "s2\<Colon>\<preceq>(G, L)" and "error_free s2"
+      by (blast)
+    with wt
+    show "s2\<Colon>\<preceq>(G, L) \<and>
+           (normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1r (If(e) c1 Else c2)\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and>
+           (error_free (Norm s0) = error_free s2)"
+      by (simp)
+  next
+    case (Loop b c e l s0 s1 s2 s3 L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1" .
+    have hyp_e: "PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 b)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T" .
+    then obtain wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
+                wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
+      by (rule wt_elim_cases) (blast)
+    from conf_s0 wt_e hyp_e
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
+      by blast
+    show "s3\<Colon>\<preceq>(G, L) \<and>
+          (normal s3 \<longrightarrow> G,L,store s3\<turnstile>In1r (l\<bullet> While(e) c)\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and>
+          (error_free (Norm s0) = error_free s3)"
+    proof (cases "normal s1 \<and> the_Bool b")
+      case True
+      from Loop True have "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" by auto
+      from Loop True have "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
+	by auto
+      from Loop True have hyp_c: "PROP ?TypeSafe s1 s2 (In1r c) \<diamondsuit>"
+	by (auto)
+      from Loop True have hyp_w: "PROP ?TypeSafe (abupd (absorb (Cont l)) s2)
+                                       s3 (In1r (l\<bullet> While(e) c)) \<diamondsuit>"
+	by (auto)
+      from conf_s1 error_free_s1 wt_c hyp_c
+      obtain conf_s2:  "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
+	by blast
+      from conf_s2 have "abupd (absorb (Cont l)) s2 \<Colon>\<preceq>(G, L)"
+	by (cases s2) (auto intro: conforms_absorb)
+      moreover
+      from error_free_s2 have "error_free (abupd (absorb (Cont l)) s2)"
+	by simp
+      moreover note wt hyp_w
+      ultimately obtain "s3\<Colon>\<preceq>(G, L)" and "error_free s3"
+	by blast
+      with wt 
+      show ?thesis
+	by (simp)
+    next
+      case False
+      with Loop have "s3=s1" by simp
+      with conf_s1 error_free_s1 wt
+      show ?thesis
+	by (simp)
+    qed
+  next
+    case (Do j s L accC T)
+    have "Norm s\<Colon>\<preceq>(G, L)" .
+    then 
+    show "(Some (Jump j), s)\<Colon>\<preceq>(G, L) \<and>
+           (normal (Some (Jump j), s) 
+           \<longrightarrow> G,L,store (Some (Jump j), s)\<turnstile>In1r (Do j)\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and>
+           (error_free (Norm s) = error_free (Some (Jump j), s))"
+      by simp
+  next
+    case (Throw a e s0 s1 L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1" .
+    have hyp: "PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 a)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Throw e)\<Colon>T" .
+    then obtain tn 
+      where      wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-Class tn" and
+            throwable: "G\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable"
+      by (rule wt_elim_cases) (auto)
+    from conf_s0 wt_e hyp obtain
+      "s1\<Colon>\<preceq>(G, L)" and
+      "(normal s1 \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>Class tn)" and
+      error_free_s1: "error_free s1"
+      by force
+    with wf throwable
+    have "abupd (throw a) s1\<Colon>\<preceq>(G, L)" 
+      by (cases s1) (auto dest: Throw_lemma)
+    with wt error_free_s1
+    show "abupd (throw a) s1\<Colon>\<preceq>(G, L) \<and>
+            (normal (abupd (throw a) s1) \<longrightarrow>
+            G,L,store (abupd (throw a) s1)\<turnstile>In1r (Throw e)\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and>
+            (error_free (Norm s0) = error_free (abupd (throw a) s1))"
+      by simp
+  next
+    case (Try catchC c1 c2 s0 s1 s2 s3 vn L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1" .
+    have sx_alloc: "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" .
+    have hyp_c1: "PROP ?TypeSafe (Norm s0) s1 (In1r c1) \<diamondsuit>" .
+    have conf_s0:"Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>In1r (Try c1 Catch(catchC vn) c2)\<Colon>T" .
+    then obtain 
+      wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+      wt_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<lparr>lcl := L(VName vn\<mapsto>Class catchC)\<rparr>\<turnstile>c2\<Colon>\<surd>" and
+      fresh_vn: "L(VName vn)=None"
+      by (rule wt_elim_cases) (auto)
+    with conf_s0 hyp_c1
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
+      by blast
+    from conf_s1 sx_alloc wf 
+    have conf_s2: "s2\<Colon>\<preceq>(G, L)" 
+      by (auto dest: sxalloc_type_sound split: option.splits)
+    from sx_alloc error_free_s1 
+    have error_free_s2: "error_free s2"
+      by (rule error_free_sxalloc)
+    show "s3\<Colon>\<preceq>(G, L) \<and>
+          (normal s3 \<longrightarrow> G,L,store s3\<turnstile>In1r (Try c1 Catch(catchC vn) c2)\<succ>\<diamondsuit>\<Colon>\<preceq>T)\<and>
+          (error_free (Norm s0) = error_free s3)"
+    proof (cases "normal s1")  
+      case True
+      with sx_alloc wf 
+      have eq_s2_s1: "s2=s1"
+	by (auto dest: sxalloc_type_sound split: option.splits)
+      with True 
+      have "\<not>  G,s2\<turnstile>catch catchC"
+	by (simp add: catch_def)
+      with Try
+      have "s3=s2"
+	by simp
+      with wt conf_s1 error_free_s1 eq_s2_s1
+      show ?thesis
+	by simp
+    next
+      case False
+      note exception_s1 = this
+      show ?thesis
+      proof (cases "G,s2\<turnstile>catch catchC") 
+	case False
+	with Try
+	have "s3=s2"
+	  by simp
+	with wt conf_s2 error_free_s2 
+	show ?thesis
+	  by simp
+      next
+	case True
+	with Try have "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3" by simp
+	from True Try 
+	have hyp_c2: "PROP ?TypeSafe (new_xcpt_var vn s2) s3 (In1r c2) \<diamondsuit>"
+	  by auto
+	from exception_s1 sx_alloc wf
+	obtain a 
+	  where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
+	  by (auto dest!: sxalloc_type_sound split: option.splits)
+	with True
+	have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class catchC"
+	  by (cases s2) simp
+	with xcpt_s2 conf_s2 wf
+	have "Norm (lupd(VName vn\<mapsto>Addr a) (store s2))
+              \<Colon>\<preceq>(G, L(VName vn\<mapsto>Class catchC))"
+	  by (auto dest: Try_lemma)
+	with hyp_c2 wt_c2 xcpt_s2 error_free_s2
+	obtain       conf_s3: "s3\<Colon>\<preceq>(G, L(VName vn\<mapsto>Class catchC))" and
+               error_free_s3: "error_free s3"
+	  by (cases s2) auto
+	from conf_s3 fresh_vn 
+	have "s3\<Colon>\<preceq>(G,L)"
+	  by (blast intro: conforms_deallocL)
+	with wt error_free_s3
+	show ?thesis
+	  by simp
+      qed
+    qed
+  next
+    case (Fin c1 c2 s0 s1 s2 x1 L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1, s1)" .
+    have c2: "G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2" .
+    have  hyp_c1: "PROP ?TypeSafe (Norm s0) (x1,s1) (In1r c1) \<diamondsuit>" .
+    have  hyp_c2: "PROP ?TypeSafe (Norm s1) s2      (In1r c2) \<diamondsuit>" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (c1 Finally c2)\<Colon>T" .
+    then obtain
+      wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+      wt_c2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c2\<Colon>\<surd>"
+      by (rule wt_elim_cases) blast
+    from conf_s0 wt_c1 hyp_c1  
+    obtain conf_s1: "(x1,s1)\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free (x1,s1)"
+      by blast
+    from conf_s1 have "Norm s1\<Colon>\<preceq>(G, L)"
+      by (rule conforms_NormI)
+    with wt_c2 hyp_c2
+    obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
+      by blast
+    show "abupd (abrupt_if (x1 \<noteq> None) x1) s2\<Colon>\<preceq>(G, L) \<and>
+          (normal (abupd (abrupt_if (x1 \<noteq> None) x1) s2) 
+           \<longrightarrow> G,L,store (abupd (abrupt_if (x1 \<noteq> None) x1) s2)
+               \<turnstile>In1r (c1 Finally c2)\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and> 
+          (error_free (Norm s0) =
+              error_free (abupd (abrupt_if (x1 \<noteq> None) x1) s2))"
+    proof (cases x1)
+      case None with conf_s2 wt show ?thesis by auto
+    next
+      case (Some x) 
+      with c2 wf conf_s1 conf_s2
+      have conf: "(abrupt_if True (Some x) (abrupt s2), store s2)\<Colon>\<preceq>(G, L)"
+	by (cases s2) (auto dest: Fin_lemma)
+      from Some error_free_s1
+      have "\<not> (\<exists> err. x=Error err)"
+	by (simp add: error_free_def)
+      with error_free_s2
+      have "error_free (abrupt_if True (Some x) (abrupt s2), store s2)"
+	by (cases s2) simp
+      with Some wt conf show ?thesis
+	by (cases s2) auto
+    qed
+  next
+    case (Init C c s0 s1 s2 s3 L accC T)
+    have     cls: "the (class G C) = c" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Init C)\<Colon>T" .
+    with cls
+    have cls_C: "class G C = Some c"
+      by - (erule wt_elim_cases,auto)
+    show "s3\<Colon>\<preceq>(G, L) \<and> (normal s3 \<longrightarrow> G,L,store s3\<turnstile>In1r (Init C)\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and>
+          (error_free (Norm s0) = error_free s3)"
+    proof (cases "inited C (globs s0)")
+      case True
+      with Init have "s3 = Norm s0"
+	by simp
+      with conf_s0 wt show ?thesis 
+	by simp
+    next
+      case False
+      with Init 
+      have "G\<turnstile>Norm ((init_class_obj G C) s0) 
+              \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1" and
+        eval_init: "G\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<rightarrow> s2" and
+	s3: "s3 = (set_lvars (locals (store s1))) s2" 
+	by auto
+      from False Init 
+      have hyp_init_super: 
+             "PROP ?TypeSafe (Norm ((init_class_obj G C) s0)) s1
+	              (In1r (if C = Object then Skip else Init (super c))) \<diamondsuit>"
+	by auto
+      with False Init (* without chaining hyp_init_super, the simplifier will
+                          loop! *)
+      have hyp_init_c:
+	"PROP ?TypeSafe ((set_lvars empty) s1) s2 (In1r (init c)) \<diamondsuit>"
+	by auto
+      from conf_s0 wf cls_C False
+      have conf_s0': "(Norm ((init_class_obj G C) s0))\<Colon>\<preceq>(G, L)"
+	by (auto dest: conforms_init_class_obj)
+      from wf cls_C have
+	wt_super:"\<lparr>prg = G, cls = accC, lcl = L\<rparr>
+                   \<turnstile>(if C = Object then Skip else Init (super c))\<Colon>\<surd>"
+	by (cases "C=Object")
+           (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
+      with conf_s0' hyp_init_super
+      obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
+	by blast 
+      then
+      have "(set_lvars empty) s1\<Colon>\<preceq>(G, empty)"
+	by (cases s1) (auto dest: conforms_set_locals )
+      moreover from error_free_s1
+      have "error_free ((set_lvars empty) s1)"
+	by simp
+      moreover note hyp_init_c wf cls_C 
+      ultimately
+      obtain conf_s2: "s2\<Colon>\<preceq>(G, empty)" and error_free_s2: "error_free s2"
+	by (auto dest!: wf_prog_cdecl wf_cdecl_wt_init)
+      with s3 conf_s1 eval_init
+      have "s3\<Colon>\<preceq>(G, L)"
+	by (cases s2,cases s1) (force dest: conforms_return eval_gext')
+      moreover from error_free_s2 s3
+      have "error_free s3"
+	by simp
+      moreover note wt
+      ultimately show ?thesis
+	by simp
+    qed
+  next
+    case (NewC C a s0 s1 s2 L accC T)
+    have         "G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1" .
+    have halloc: "G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2" .
+    have hyp: "PROP ?TypeSafe (Norm s0) s1 (In1r (Init C)) \<diamondsuit>" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (NewC C)\<Colon>T" .
+    then obtain is_cls_C: "is_class G C" and
+                       T: "T=Inl (Class C)"
+      by (rule wt_elim_cases) (auto dest: is_acc_classD)
+    with conf_s0 hyp
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
+      by auto
+    from conf_s1 halloc wf is_cls_C
+    obtain halloc_type_safe: "s2\<Colon>\<preceq>(G, L)" 
+                             "(normal s2 \<longrightarrow> G,store s2\<turnstile>Addr a\<Colon>\<preceq>Class C)"
+      by (cases s2) (auto dest!: halloc_type_sound)
+    from halloc error_free_s1 
+    have "error_free s2"
+      by (rule error_free_halloc)
+    with halloc_type_safe T
+    show "s2\<Colon>\<preceq>(G, L) \<and> 
+          (normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1l (NewC C)\<succ>In1 (Addr a)\<Colon>\<preceq>T)  \<and>
+          (error_free (Norm s0) = error_free s2)"
+      by auto
+  next
+    case (NewA T a e i s0 s1 s2 s3 L accC Ta)
+    have "G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1" .
+    have "G\<turnstile>s1 \<midarrow>e-\<succ>i\<rightarrow> s2" .
+    have halloc: "G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3" .
+    have hyp_init: "PROP ?TypeSafe (Norm s0) s1 (In1r (init_comp_ty T)) \<diamondsuit>" .
+    have hyp_size: "PROP ?TypeSafe s1 s2 (In1l e) (In1 i)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have     wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (New T[e])\<Colon>Ta" .
+    then obtain
+      wt_init: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and
+      wt_size: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-PrimT Integer" and
+            T: "is_type G T" and
+           Ta: "Ta=Inl (T.[])"
+      by (rule wt_elim_cases) (auto intro: wt_init_comp_ty dest: is_acc_typeD)
+    from conf_s0 wt_init hyp_init
+    obtain "s1\<Colon>\<preceq>(G, L)" "error_free s1"
+      by blast
+    with wt_size hyp_size
+    obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
+      by blast
+    from conf_s2 have "abupd (check_neg i) s2\<Colon>\<preceq>(G, L)"
+      by (cases s2) auto
+    with halloc wf T 
+    have halloc_type_safe:
+          "s3\<Colon>\<preceq>(G, L) \<and> (normal s3 \<longrightarrow> G,store s3\<turnstile>Addr a\<Colon>\<preceq>T.[])"
+      by (cases s3) (auto dest!: halloc_type_sound)
+    from halloc error_free_s2
+    have "error_free s3"
+      by (auto dest: error_free_halloc)
+    with halloc_type_safe Ta
+    show "s3\<Colon>\<preceq>(G, L) \<and> 
+          (normal s3 \<longrightarrow> G,L,store s3\<turnstile>In1l (New T[e])\<succ>In1 (Addr a)\<Colon>\<preceq>Ta) \<and>
+          (error_free (Norm s0) = error_free s3) "
+      by simp
+  next
+    case (Cast castT e s0 s1 s2 v L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1" .
+    have s2:"s2 = abupd (raise_if (\<not> G,store s1\<turnstile>v fits castT) ClassCast) s1" .
+    have hyp: "PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (Cast castT e)\<Colon>T" .
+    then obtain eT
+      where wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" and
+              eT: "G\<turnstile>eT\<preceq>? castT" and 
+               T: "T=Inl castT"
+      by (rule wt_elim_cases) auto
+    with conf_s0 hyp
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
+      by blast
+    from conf_s1 s2 
+    have conf_s2: "s2\<Colon>\<preceq>(G, L)"
+      by (cases s1) simp
+    from error_free_s1 s2
+    have error_free_s2: "error_free s2"
+      by simp
+    {
+      assume norm_s2: "normal s2"
+      have "G,L,store s2\<turnstile>In1l (Cast castT e)\<succ>In1 v\<Colon>\<preceq>T"
+      proof -
+	from s2 norm_s2 have "normal s1"
+	  by (cases s1) simp
+	with wt_e conf_s0 hyp 
+	have "G,store s1\<turnstile>v\<Colon>\<preceq>eT"
+	  by force
+	with eT wf s2 T norm_s2
+	show ?thesis
+	  by (cases s1) (auto dest: fits_conf)
+      qed
+    }
+    with conf_s2 error_free_s2
+    show "s2\<Colon>\<preceq>(G, L) \<and> 
+           (normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1l (Cast castT e)\<succ>In1 v\<Colon>\<preceq>T)  \<and>
+           (error_free (Norm s0) = error_free s2)"
+      by blast
+  next
+    case (Inst T b e s0 s1 v L accC T')
+    then show ?case
+      by (auto elim!: wt_elim_cases)
+  next
+    case (Lit s v L accC T)
+    then show ?case
+      by (auto elim!: wt_elim_cases 
+               intro: conf_litval simp add: empty_dt_def)
+  next
+    case (Super s L accC T)
+    have conf_s: "Norm s\<Colon>\<preceq>(G, L)" .
+    have     wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l Super\<Colon>T" .
+    then obtain C c where
+             C: "L This = Some (Class C)" and
+       neq_Obj: "C\<noteq>Object" and
+         cls_C: "class G C = Some c" and
+             T: "T=Inl (Class (super c))"
+      by (rule wt_elim_cases) auto
+    from C conf_s have "G,s\<turnstile>val_this s\<Colon>\<preceq>Class C"
+      by (blast intro: conforms_localD [THEN lconfD])
+    with neq_Obj cls_C wf
+    have "G,s\<turnstile>val_this s\<Colon>\<preceq>Class (super c)"
+      by (auto intro: conf_widen
+                dest: subcls_direct[THEN widen.subcls])
+    with T conf_s
+    show "Norm s\<Colon>\<preceq>(G, L) \<and>
+           (normal (Norm s) \<longrightarrow> 
+              G,L,store (Norm s)\<turnstile>In1l Super\<succ>In1 (val_this s)\<Colon>\<preceq>T) \<and>
+           (error_free (Norm s) = error_free (Norm s))"
+      by simp
+  next
+    case (Acc f s0 s1 v va L accC T)
+    then show ?case
+      by (force elim!: wt_elim_cases)
+  next
+    case (Ass e f s0 s1 s2 v var w L accC T)
+    have eval_var: "G\<turnstile>Norm s0 \<midarrow>var=\<succ>(w, f)\<rightarrow> s1" .
+    have   eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2" .
+    have  hyp_var: "PROP ?TypeSafe (Norm s0) s1 (In2 var) (In2 (w,f))" .
+    have    hyp_e: "PROP ?TypeSafe s1 s2 (In1l e) (In1 v)" .
+    have  conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have       wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (var:=e)\<Colon>T" .
+    then obtain varT eT where
+	 wt_var: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>var\<Colon>=varT" and
+	   wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" and
+	  widen: "G\<turnstile>eT\<preceq>varT" and
+              T: "T=Inl eT"
+      by (rule wt_elim_cases) auto
+    from conf_s0 wt_var hyp_var
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
+      by blast
+    with wt_e hyp_e
+    obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
+      by blast
+    show "assign f v s2\<Colon>\<preceq>(G, L) \<and>
+           (normal (assign f v s2) \<longrightarrow>
+            G,L,store (assign f v s2)\<turnstile>In1l (var:=e)\<succ>In1 v\<Colon>\<preceq>T) \<and>
+            (error_free (Norm s0) = error_free (assign f v s2))"
+    proof (cases "normal s1")
+      case False
+      with eval_e 
+      have "s2=s1"
+	by auto
+      with False conf_s1 error_free_s1
+      show ?thesis
+	by auto
+    next
+      case True
+      note normal_s1=this
+      show ?thesis 
+      proof (cases "normal s2")
+	case False
+	with conf_s2 error_free_s2 
+	show ?thesis
+	  by auto
+      next
+	case True
+	from True normal_s1 conf_s1 wt_e hyp_e
+	have conf_v_eT: "G,store s2\<turnstile>v\<Colon>\<preceq>eT"
+	  by force
+	with widen wf
+	have conf_v_varT: "G,store s2\<turnstile>v\<Colon>\<preceq>varT"
+	  by (auto intro: conf_widen)
+	from conf_s0 normal_s1 wt_var hyp_var
+	have "G,L,store s1\<turnstile>In2 var\<succ>In2 (w, f)\<Colon>\<preceq>Inl varT"
+	  by blast
+	then 
+	have conf_assign:  "store s1\<le>|f\<preceq>varT\<Colon>\<preceq>(G, L)" 
+	  by auto
+	from conf_v_eT conf_v_varT conf_assign normal_s1 True wf eval_var 
+	  eval_e T conf_s2 error_free_s2
+	show ?thesis
+	  by (cases s1, cases s2) 
+	     (auto dest!: Ass_lemma simp add: assign_conforms_def)
+      qed
+    qed
+  next
+    case (Cond b e0 e1 e2 s0 s1 s2 v L accC T)
+    have eval_e0: "G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1" .
+    have "G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2" .
+    have hyp_e0: "PROP ?TypeSafe (Norm s0) s1 (In1l e0) (In1 b)" .
+    have hyp_if: "PROP ?TypeSafe s1 s2 
+                       (In1l (if the_Bool b then e1 else e2)) (In1 v)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (e0 ? e1 : e2)\<Colon>T" .
+    then obtain T1 T2 statT where
+      wt_e0: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
+      wt_e1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e1\<Colon>-T1" and
+      wt_e2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e2\<Colon>-T2" and 
+      statT: "G\<turnstile>T1\<preceq>T2 \<and> statT = T2  \<or>  G\<turnstile>T2\<preceq>T1 \<and> statT =  T1" and
+      T    : "T=Inl statT"
+      by (rule wt_elim_cases) auto
+    with wt_e0 conf_s0 hyp_e0
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1" 
+      by blast
+    with wt_e1 wt_e2 statT hyp_if
+    obtain dynT where
+      conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2" and
+      conf_res: 
+          "(normal s2 \<longrightarrow>
+        G,L,store s2\<turnstile>In1l (if the_Bool b then e1 else e2)\<succ>In1 v\<Colon>\<preceq>Inl dynT)" and
+      dynT: "dynT = (if the_Bool b then T1 else T2)"
+      by (cases "the_Bool b") force+
+    from statT dynT  
+    have "G\<turnstile>dynT\<preceq>statT"
+      by (cases "the_Bool b") auto
+    with conf_s2 conf_res error_free_s2 T wf
+    show "s2\<Colon>\<preceq>(G, L) \<and>
+           (normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1l (e0 ? e1 : e2)\<succ>In1 v\<Colon>\<preceq>T) \<and>
+           (error_free (Norm s0) = error_free s2)"
+      by (auto)
+  next
+    case (Call invDeclC a' accC' args e mn mode pTs' s0 s1 s2 s3 s3' s4 statT 
+           v vs L accC T)
+    have eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1" .
+    have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2" .
+    have invDeclC: "invDeclC 
+                      = invocation_declclass G mode (store s2) a' statT 
+                           \<lparr>name = mn, parTs = pTs'\<rparr>" .
+    have init_lvars: 
+           "s3 = init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> mode a' vs s2".
+    have check: "s3' =
+       check_method_access G accC' statT mode \<lparr>name = mn, parTs = pTs'\<rparr> a' s3" .
+    have eval_methd: 
+           "G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<rightarrow> s4" .
+    have     hyp_e: "PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 a')" .
+    have  hyp_args: "PROP ?TypeSafe s1 s2 (In3 args) (In3 vs)" .
+    have hyp_methd: "PROP ?TypeSafe s3' s4 
+                     (In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)) (In1 v)".
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>
+                    \<turnstile>In1l ({accC',statT,mode}e\<cdot>mn( {pTs'}args))\<Colon>T" .
+    from wt obtain pTs statDeclT statM where
+                 wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT" and
+              wt_args: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>args\<Colon>\<doteq>pTs" and
+                statM: "max_spec G accC statT \<lparr>name=mn,parTs=pTs\<rparr> 
+                         = {((statDeclT,statM),pTs')}" and
+                 mode: "mode = invmode statM e" and
+                    T: "T =Inl (resTy statM)" and
+        eq_accC_accC': "accC=accC'"
+      by (rule wt_elim_cases) auto
+    from conf_s0 wt_e hyp_e 
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and
+           conf_a': "normal s1 \<Longrightarrow> G, store s1\<turnstile>a'\<Colon>\<preceq>RefT statT" and
+           error_free_s1: "error_free s1" 
+      by force
+    with wt_args hyp_args
+    obtain    conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+            conf_args: "normal s2 
+                         \<Longrightarrow>  list_all2 (conf G (store s2)) vs pTs" and
+        error_free_s2: "error_free s2" 
+      by force
+    from error_free_s2 init_lvars
+    have error_free_s3: "error_free s3"
+      by (auto simp add: init_lvars_def2)
+    from statM 
+    obtain
+      statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>" and
+      pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
+      by (blast dest: max_spec2mheads)
+    from check
+    have eq_store_s3'_s3: "store s3'=store s3"
+      by (cases s3) (simp add: check_method_access_def Let_def)
+    obtain invC
+      where invC: "invC = invocation_class mode (store s2) a' statT"
+      by simp
+    with init_lvars
+    have invC': "invC = (invocation_class mode (store s3) a' statT)"
+      by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
+    show "(set_lvars (locals (store s2))) s4\<Colon>\<preceq>(G, L) \<and>
+             (normal ((set_lvars (locals (store s2))) s4) \<longrightarrow>
+               G,L,store ((set_lvars (locals (store s2))) s4)
+               \<turnstile>In1l ({accC',statT,mode}e\<cdot>mn( {pTs'}args))\<succ>In1 v\<Colon>\<preceq>T) \<and>
+             (error_free (Norm s0) =
+                error_free ((set_lvars (locals (store s2))) s4))"
+    proof (cases "normal s2")
+      case False
+      with init_lvars 
+      obtain keep_abrupt: "abrupt s3 = abrupt s2" and
+             "store s3 = store (init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> 
+                                            mode a' vs s2)" 
+	by (auto simp add: init_lvars_def2)
+      moreover
+      from keep_abrupt False check
+      have eq_s3'_s3: "s3'=s3" 
+	by (auto simp add: check_method_access_def Let_def)
+      moreover
+      from eq_s3'_s3 False keep_abrupt eval_methd
+      have "s4=s3'"
+	by auto
+      ultimately have
+	"set_lvars (locals (store s2)) s4 = s2"
+	by (cases s2,cases s4) (simp add: init_lvars_def2)
+      with False conf_s2 error_free_s2
+      show ?thesis
+	by auto
+    next
+      case True
+      note normal_s2 = True
+      with eval_args
+      have normal_s1: "normal s1"
+	by (cases "normal s1") auto
+      with conf_a' eval_args 
+      have conf_a'_s2: "G, store s2\<turnstile>a'\<Colon>\<preceq>RefT statT"
+	by (auto dest: eval_gext intro: conf_gext)
+      show ?thesis
+      proof (cases "a'=Null \<longrightarrow> is_static statM")
+	case False
+	then obtain not_static: "\<not> is_static statM" and Null: "a'=Null" 
+	  by blast
+	with normal_s2 init_lvars mode
+	obtain np: "abrupt s3 = Some (Xcpt (Std NullPointer))" and
+                   "store s3 = store (init_lvars G invDeclC 
+                                       \<lparr>name = mn, parTs = pTs'\<rparr> mode a' vs s2)"
+	  by (auto simp add: init_lvars_def2)
+	moreover
+	from np check
+	have eq_s3'_s3: "s3'=s3" 
+	  by (auto simp add: check_method_access_def Let_def)
+	moreover
+	from eq_s3'_s3 np eval_methd
+	have "s4=s3'"
+	  by auto
+	ultimately have
+	  "set_lvars (locals (store s2)) s4 
+           = (Some (Xcpt (Std NullPointer)),store s2)"
+	  by (cases s2,cases s4) (simp add: init_lvars_def2)
+	with conf_s2 error_free_s2
+	show ?thesis
+	  by (cases s2) (auto dest: conforms_NormI)
+      next
+	case True
+	with mode have notNull: "mode = IntVir \<longrightarrow> a' \<noteq> Null"
+	  by (auto dest!: Null_staticD)
+	with conf_s2 conf_a'_s2 wf invC  
+	have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
+	  by (cases s2) (auto intro: DynT_propI)
+	with wt_e statM' invC mode wf 
+	obtain dynM where 
+          dynM: "dynlookup G statT invC  \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+          acc_dynM: "G \<turnstile>Methd  \<lparr>name=mn,parTs=pTs'\<rparr> dynM 
+                          in invC dyn_accessible_from accC"
+	  by (force dest!: call_access_ok)
+	with invC' check eq_accC_accC'
+	have eq_s3'_s3: "s3'=s3"
+	  by (auto simp add: check_method_access_def Let_def)
+	from dynT_prop wf wt_e statM' mode invC invDeclC dynM 
+	obtain 
+	   wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
+	     dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+           iscls_invDeclC: "is_class G invDeclC" and
+	        invDeclC': "invDeclC = declclass dynM" and
+	     invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
+	    resTy_widen: "G\<turnstile>resTy dynM\<preceq>resTy statM" and
+	   is_static_eq: "is_static dynM = is_static statM" and
+	   involved_classes_prop:
+             "(if invmode statM 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 invDeclC) \<or>
+                     (\<forall>statC. statT \<noteq> ClassT statC \<and> invDeclC = Object)) \<and>
+                      statDeclT = ClassT invDeclC)"
+	  by (auto dest: DynT_mheadsD)
+	obtain L' where 
+	   L':"L'=(\<lambda> k. 
+                 (case k of
+                    EName e
+                    \<Rightarrow> (case e of 
+                          VNam v 
+                          \<Rightarrow>(table_of (lcls (mbody (mthd dynM)))
+                             (pars (mthd dynM)[\<mapsto>]pTs')) v
+                        | Res \<Rightarrow> Some (resTy dynM))
+                  | This \<Rightarrow> if is_static statM 
+                            then None else Some (Class invDeclC)))"
+	  by simp
+	from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
+             wf eval_args conf_a' mode notNull wf_dynM involved_classes_prop
+	have conf_s3: "s3\<Colon>\<preceq>(G,L')"
+	  apply - 
+             (* FIXME confomrs_init_lvars should be 
+                adjusted to be more directy applicable *)
+	  apply (drule conforms_init_lvars [of G invDeclC 
+                  "\<lparr>name=mn,parTs=pTs'\<rparr>" dynM "store s2" vs pTs "abrupt s2" 
+                  L statT invC a' "(statDeclT,statM)" e])
+	  apply (rule wf)
+	  apply (rule conf_args,assumption)
+	  apply (simp add: pTs_widen)
+	  apply (cases s2,simp)
+	  apply (rule dynM')
+	  apply (force dest: ty_expr_is_type)
+	  apply (rule invC_widen)
+	  apply (force intro: conf_gext dest: eval_gext)
+	  apply simp
+	  apply simp
+	  apply (simp add: invC)
+	  apply (simp add: invDeclC)
+	  apply (force dest: wf_mdeclD1 is_acc_typeD)
+	  apply (cases s2, simp add: L' init_lvars
+	                      cong add: lname.case_cong ename.case_cong)
+	  done 
+	from  is_static_eq wf_dynM L'
+	obtain mthdT where
+	   "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+            \<turnstile>Body invDeclC (stmt (mbody (mthd dynM)))\<Colon>-mthdT" and
+	   mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
+	  by - (drule wf_mdecl_bodyD,
+                simp cong add: lname.case_cong ename.case_cong)
+	with dynM' iscls_invDeclC invDeclC'
+	have
+	   "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+            \<turnstile>(Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<Colon>-mthdT"
+	  by (auto intro: wt.Methd)
+	with eq_s3'_s3 conf_s3 error_free_s3 
+             hyp_methd [of L' invDeclC "Inl mthdT"]
+	obtain  conf_s4: "s4\<Colon>\<preceq>(G, L')" and 
+	       conf_Res: "normal s4 \<longrightarrow> G,store s4\<turnstile>v\<Colon>\<preceq>mthdT" and
+	  error_free_s4: "error_free s4"
+	  by auto
+	from init_lvars eval_methd eq_s3'_s3 
+	have "store s2\<le>|store s4"
+	  by (cases s2) (auto dest!: eval_gext simp add: init_lvars_def2 )
+	with conf_s2 conf_s4
+	have "(set_lvars (locals (store s2))) s4\<Colon>\<preceq>(G, L)"
+	  by (cases s2,cases s4) (auto intro: conforms_return)
+	moreover 
+	from conf_Res mthdT_widen resTy_widen wf
+	have "normal s4 
+             \<longrightarrow> G,store s4\<turnstile>v\<Colon>\<preceq>(resTy statM)"
+	  by (auto dest: widen_trans)
+	then
+	have "normal ((set_lvars (locals (store s2))) s4)
+             \<longrightarrow> G,store((set_lvars (locals (store s2))) s4) \<turnstile>v\<Colon>\<preceq>(resTy statM)"
+	  by (cases s4) auto
+	moreover note error_free_s4 T
+	ultimately 
+	show ?thesis
+	  by simp
+      qed
+    qed
+  next
+    case (Methd D s0 s1 sig v L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<rightarrow> s1" .
+    have hyp: "PROP ?TypeSafe (Norm s0) s1 (In1l (body G D sig)) (In1 v)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have      wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (Methd D sig)\<Colon>T" .
+    then obtain m bodyT where
+      D: "is_class G D" and
+      m: "methd G D sig = Some m" and
+      wt_body: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>
+                   \<turnstile>Body (declclass m) (stmt (mbody (mthd m)))\<Colon>-bodyT" and
+      T: "T=Inl bodyT"
+      by (rule wt_elim_cases) auto
+    with hyp [of _ _ "(Inl bodyT)"] conf_s0 
+    show "s1\<Colon>\<preceq>(G, L) \<and> 
+           (normal s1 \<longrightarrow> G,L,snd s1\<turnstile>In1l (Methd D sig)\<succ>In1 v\<Colon>\<preceq>T) \<and>
+           (error_free (Norm s0) = error_free s1)"
+      by (auto simp add: Let_def body_def)
+  next
+    case (Body D c s0 s1 s2 L accC T)
+    have "G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1" .
+    have "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" .
+    have hyp_init: "PROP ?TypeSafe (Norm s0) s1 (In1r (Init D)) \<diamondsuit>" .
+    have hyp_c: "PROP ?TypeSafe s1 s2 (In1r c) \<diamondsuit>" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (Body D c)\<Colon>T" .
+    then obtain bodyT where
+         iscls_D: "is_class G D" and
+            wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>" and
+         resultT: "L Result = Some bodyT" and
+      isty_bodyT: "is_type G bodyT" and (* ### not needed! remove from wt? *)
+               T: "T=Inl bodyT"
+      by (rule wt_elim_cases) auto
+    from conf_s0 iscls_D hyp_init
+    obtain "s1\<Colon>\<preceq>(G, L)" "error_free s1"
+      by auto
+    with wt_c hyp_c
+    obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
+      by blast
+    from conf_s2
+    have "abupd (absorb Ret) s2\<Colon>\<preceq>(G, L)"
+      by (cases s2) (auto intro: conforms_absorb)
+    moreover
+    from error_free_s2
+    have "error_free (abupd (absorb Ret) s2)"
+      by simp
+    moreover note T resultT
+    ultimately
+    show "abupd (absorb Ret) s2\<Colon>\<preceq>(G, L) \<and>
+           (normal (abupd (absorb Ret) s2) \<longrightarrow>
+             G,L,store (abupd (absorb Ret) s2)
+              \<turnstile>In1l (Body D c)\<succ>In1 (the (locals (store s2) Result))\<Colon>\<preceq>T) \<and>
+          (error_free (Norm s0) = error_free (abupd (absorb Ret) s2)) "
+      by (cases s2) (auto intro: conforms_locals)
+  next
+    case (LVar s vn L accC T)
+    have conf_s: "Norm s\<Colon>\<preceq>(G, L)" and 
+             wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In2 (LVar vn)\<Colon>T" .
+    then obtain vnT where
+      vnT: "L vn = Some vnT" and
+        T: "T=Inl vnT"
+      by (auto elim!: wt_elim_cases)
+    from conf_s vnT
+    have conf_fst: "G,s\<turnstile>fst (lvar vn s)\<Colon>\<preceq>vnT"  
+      by (auto elim: conforms_localD [THEN lconfD]  
+               simp add: lvar_def)
+    moreover
+    from conf_s conf_fst vnT 
+    have "s\<le>|snd (lvar vn s)\<preceq>vnT\<Colon>\<preceq>(G, L)"
+      by (auto elim: conforms_lupd simp add: assign_conforms_def lvar_def)
+    moreover note conf_s T
+    ultimately 
+    show "Norm s\<Colon>\<preceq>(G, L) \<and>
+                 (normal (Norm s) \<longrightarrow>
+                    G,L,store (Norm s)\<turnstile>In2 (LVar vn)\<succ>In2 (lvar vn s)\<Colon>\<preceq>T) \<and>
+                 (error_free (Norm s) = error_free (Norm s))"
+      by simp 
+  next
+    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v L accC' T)
+    have eval_init: "G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1" .
+    have eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2" .
+    have fvar: "(v, s2') = fvar statDeclC stat fn a s2" .
+    have check: "s3 = check_field_access G accC statDeclC fn stat a s2'" .
+    have hyp_init: "PROP ?TypeSafe (Norm s0) s1 (In1r (Init statDeclC)) \<diamondsuit>" .
+    have hyp_e: "PROP ?TypeSafe s1 s2 (In1l e) (In1 a)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have wt: "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile>In2 ({accC,statDeclC,stat}e..fn)\<Colon>T" .
+    then obtain statC f where
+                wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-Class statC" and
+            accfield: "accfield G accC statC fn = Some (statDeclC,f)" and
+       eq_accC_accC': "accC=accC'" and
+                stat: "stat=is_static f" and
+	           T: "T=(Inl (type f))"
+      by (rule wt_elim_cases) (auto simp add: member_is_static_simp)
+    from wf wt_e 
+    have iscls_statC: "is_class G statC"
+      by (auto dest: ty_expr_is_type type_is_class)
+    with wf accfield 
+    have iscls_statDeclC: "is_class G statDeclC"
+      by (auto dest!: accfield_fields dest: fields_declC)
+    with conf_s0 hyp_init
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
+      by auto
+    from conf_s1 wt_e hyp_e
+    obtain       conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+                  conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC" 
+      by force
+    from conf_s1 wt_e error_free_s1 hyp_e
+    have error_free_s2: "error_free s2"
+      by auto
+    from fvar 
+    have store_s2': "store s2'=store s2"
+      by (cases s2) (simp add: fvar_def2)
+    with fvar conf_s2 
+    have conf_s2': "s2'\<Colon>\<preceq>(G, L)"
+      by (cases s2,cases stat) (auto simp add: fvar_def2)
+    from eval_init 
+    have initd_statDeclC_s1: "initd statDeclC s1"
+      by (rule init_yields_initd)
+    from accfield wt_e eval_init eval_e conf_s2 conf_a fvar stat check  wf
+    have eq_s3_s2': "s3=s2'"  
+      by (auto dest!: error_free_field_access)
+    have conf_v: "normal s2' \<Longrightarrow> 
+           G,store s2'\<turnstile>fst v\<Colon>\<preceq>type f \<and> store s2'\<le>|snd v\<preceq>type f\<Colon>\<preceq>(G, L)"
+    proof - (*###FVar_lemma should be adjusted to be more directy applicable *)
+      assume normal: "normal s2'"
+      obtain vv vf x2 store2 store2'
+	where  v: "v=(vv,vf)" and
+              s2: "s2=(x2,store2)" and
+         store2': "store s2' = store2'"
+	by (cases v,cases s2,cases s2') blast
+      from iscls_statDeclC obtain c
+	where c: "class G statDeclC = Some c"
+	by auto
+      have "G,store2'\<turnstile>vv\<Colon>\<preceq>type f \<and> store2'\<le>|vf\<preceq>type f\<Colon>\<preceq>(G, L)"
+      proof (rule FVar_lemma [of vv vf store2' statDeclC f fn a x2 store2 
+                               statC G c L "store s1"])
+	from v normal s2 fvar stat store2' 
+	show "((vv, vf), Norm store2') = 
+               fvar statDeclC (static f) fn a (x2, store2)"
+	  by (auto simp add: member_is_static_simp)
+	from accfield iscls_statC wf
+	show "G\<turnstile>statC\<preceq>\<^sub>C statDeclC"
+	  by (auto dest!: accfield_fields dest: fields_declC)
+	from accfield
+	show fld: "table_of (fields G statC) (fn, statDeclC) = Some f"
+	  by (auto dest!: accfield_fields)
+	from wf show "wf_prog G" .
+	from conf_a s2 show "x2 = None \<longrightarrow> G,store2\<turnstile>a\<Colon>\<preceq>Class statC"
+	  by auto
+	from fld wf iscls_statC
+	show "statDeclC \<noteq> Object "
+	  by (cases "statDeclC=Object") (drule fields_declC,simp+)+
+	from c show "class G statDeclC = Some c" .
+	from conf_s2 s2 show "(x2, store2)\<Colon>\<preceq>(G, L)" by simp
+	from eval_e s2 show "snd s1\<le>|store2" by (auto dest: eval_gext)
+	from initd_statDeclC_s1 show "inited statDeclC (globs (snd s1))" 
+	  by simp
+      qed
+      with v s2 store2'  
+      show ?thesis
+	by simp
+    qed
+    from fvar error_free_s2
+    have "error_free s2'"
+      by (cases s2)
+         (auto simp add: fvar_def2 intro!: error_free_FVar_lemma)
+    with conf_v T conf_s2' eq_s3_s2'
+    show "s3\<Colon>\<preceq>(G, L) \<and>
+          (normal s3 
+           \<longrightarrow> G,L,store s3\<turnstile>In2 ({accC,statDeclC,stat}e..fn)\<succ>In2 v\<Colon>\<preceq>T) \<and>
+          (error_free (Norm s0) = error_free s3)"
+      by auto
+  next
+    case (AVar a e1 e2 i s0 s1 s2 s2' v L accC T)
+    have eval_e1: "G\<turnstile>Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1" .
+    have eval_e2: "G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2" .
+    have hyp_e1: "PROP ?TypeSafe (Norm s0) s1 (In1l e1) (In1 a)" .
+    have hyp_e2: "PROP ?TypeSafe s1 s2 (In1l e2) (In1 i)" .
+    have avar: "(v, s2') = avar G i a s2" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have wt:  "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In2 (e1.[e2])\<Colon>T" .
+    then obtain elemT
+       where wt_e1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e1\<Colon>-elemT.[]" and
+             wt_e2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e2\<Colon>-PrimT Integer" and
+                 T: "T= Inl elemT"
+      by (rule wt_elim_cases) auto
+    from  conf_s0 wt_e1 hyp_e1 
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and
+            conf_a: "(normal s1 \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>elemT.[])" and
+            error_free_s1: "error_free s1"
+      by force
+    from conf_s1 error_free_s1 wt_e2 hyp_e2
+    obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
+      by blast
+    from avar 
+    have "store s2'=store s2"
+      by (cases s2) (simp add: avar_def2)
+    with avar conf_s2 
+    have conf_s2': "s2'\<Colon>\<preceq>(G, L)"
+      by (cases s2) (auto simp add: avar_def2)
+    from avar error_free_s2
+    have error_free_s2': "error_free s2'"
+      by (cases s2) (auto simp add: avar_def2 )
+    have "normal s2' \<Longrightarrow> 
+           G,store s2'\<turnstile>fst v\<Colon>\<preceq>elemT \<and> store s2'\<le>|snd v\<preceq>elemT\<Colon>\<preceq>(G, L)"
+    proof -(*###AVar_lemma should be adjusted to be more directy applicable *)
+      assume normal: "normal s2'"
+      show ?thesis
+      proof -
+	obtain vv vf x1 store1 x2 store2 store2'
+	   where  v: "v=(vv,vf)" and
+                 s1: "s1=(x1,store1)" and
+                 s2: "s2=(x2,store2)" and
+	    store2': "store2'=store s2'"
+	  by (cases v,cases s1, cases s2, cases s2') blast 
+	have "G,store2'\<turnstile>vv\<Colon>\<preceq>elemT \<and> store2'\<le>|vf\<preceq>elemT\<Colon>\<preceq>(G, L)"
+	proof (rule AVar_lemma [of G x1 store1 e2 i x2 store2 vv vf store2' a,
+                                 OF wf])
+	  from s1 s2 eval_e2 show "G\<turnstile>(x1, store1) \<midarrow>e2-\<succ>i\<rightarrow> (x2, store2)"
+	    by simp
+	  from v normal s2 store2' avar 
+	  show "((vv, vf), Norm store2') = avar G i a (x2, store2)"
+	    by auto
+	  from s2 conf_s2 show "(x2, store2)\<Colon>\<preceq>(G, L)" by simp
+	  from s1 conf_a show  "x1 = None \<longrightarrow> G,store1\<turnstile>a\<Colon>\<preceq>elemT.[]" by simp 
+	  from eval_e2 s1 s2 show "store1\<le>|store2" by (auto dest: eval_gext)
+	qed
+	with v s1 s2 store2' 
+	show ?thesis
+	  by simp
+      qed
+    qed
+    with conf_s2' error_free_s2' T 
+    show "s2'\<Colon>\<preceq>(G, L) \<and>
+           (normal s2' \<longrightarrow> G,L,store s2'\<turnstile>In2 (e1.[e2])\<succ>In2 v\<Colon>\<preceq>T) \<and>
+           (error_free (Norm s0) = error_free s2') "
+      by auto
+  next
+    case (Nil s0 L accC T)
+    then show ?case
+      by (auto elim!: wt_elim_cases)
+  next
+    case (Cons e es s0 s1 s2 v vs L accC T)
+    have eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1" .
+    have eval_es: "G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s2" .
+    have hyp_e: "PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)" .
+    have hyp_es: "PROP ?TypeSafe s1 s2 (In3 es) (In3 vs)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    have wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In3 (e # es)\<Colon>T" .
+    then obtain eT esT where
+       wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" and
+       wt_es: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>es\<Colon>\<doteq>esT" and
+       T: "T=Inr (eT#esT)"
+      by (rule wt_elim_cases) blast
+    from hyp_e [OF conf_s0 wt_e]
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1" and 
+      conf_v: "normal s1 \<longrightarrow> G,store s1\<turnstile>v\<Colon>\<preceq>eT"
+      by auto
+    from eval_es conf_v 
+    have conf_v': "normal s2 \<longrightarrow> G,store s2\<turnstile>v\<Colon>\<preceq>eT"
+      apply clarify
+      apply (rule conf_gext)
+      apply (auto dest: eval_gext)
+      done
+    from hyp_es [OF conf_s1 wt_es] error_free_s1 
+    obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and 
+           error_free_s2: "error_free s2" and
+           conf_vs: "normal s2 \<longrightarrow> list_all2 (conf G (store s2)) vs esT"
+      by auto
+    with conf_v' T
+    show 
+      "s2\<Colon>\<preceq>(G, L) \<and> 
+      (normal s2 \<longrightarrow> G,L,store s2\<turnstile>In3 (e # es)\<succ>In3 (v # vs)\<Colon>\<preceq>T) \<and>
+      (error_free (Norm s0) = error_free s2) "
+      by auto
+  qed
+  then show ?thesis .
+qed
+ 
+corollary eval_ts: 
+ "\<lbrakk>G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow> s'; wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-T\<rbrakk> 
+\<Longrightarrow>  s'\<Colon>\<preceq>(G,L) \<and> (normal s' \<longrightarrow> G,store s'\<turnstile>v\<Colon>\<preceq>T) \<and> 
+     (error_free s = error_free s')"
 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)"
+corollary evals_ts: 
+"\<lbrakk>G\<turnstile>s \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s'; wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>es\<Colon>\<doteq>Ts\<rbrakk> 
+\<Longrightarrow>  s'\<Colon>\<preceq>(G,L) \<and> (normal s' \<longrightarrow> list_all2 (conf G (store s')) vs Ts) \<and> 
+     (error_free s = error_free s')" 
 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)"
+corollary evar_ts: 
+"\<lbrakk>G\<turnstile>s \<midarrow>v=\<succ>vf\<rightarrow> s'; wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>v\<Colon>=T\<rbrakk> \<Longrightarrow>  
+  s'\<Colon>\<preceq>(G,L) \<and> (normal s' \<longrightarrow> G,L,(store s')\<turnstile>In2 v\<succ>In2 vf\<Colon>\<preceq>Inl T) \<and> 
+  (error_free s = error_free s')"
 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)"
+"\<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) \<and> (error_free s \<longrightarrow> error_free s')"
 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
--- a/src/HOL/Bali/WellForm.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/WellForm.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Bali/WellForm.thy
     ID:         $Id$
-    Author:     David von Oheimb
+    Author:     David von Oheimb and Norbert Schirmer
     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 *)
 
@@ -31,8 +31,8 @@
 *}
 
 section "well-formed field declarations"
-  (* well-formed field declaration (common part for classes and interfaces),
-     cf. 8.3 and (9.3) *)
+text  {* 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"
@@ -98,7 +98,7 @@
                         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))
+	        | This \<Rightarrow> if is_static m then None else Some (Class C))
           \<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>"
 
 lemma wf_mheadI: 
@@ -122,7 +122,7 @@
                 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))
+        | This \<Rightarrow> if is_static m then None else Some (Class C))
   \<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>
   \<rbrakk> \<Longrightarrow>  
   wf_mdecl G C (sig,m)"
@@ -149,7 +149,7 @@
             \<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))
+          | This \<Rightarrow> if is_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
@@ -1291,12 +1291,6 @@
   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 
@@ -1326,8 +1320,9 @@
       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 note wf False cls_C  
+      ultimately have "G\<turnstile>super c \<preceq>\<^sub>C declclass m"  
+	by (auto intro: Hyp [rule_format])
       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
@@ -1337,104 +1332,10 @@
   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)"
@@ -1471,7 +1372,6 @@
   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
@@ -1757,9 +1657,8 @@
  ) "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)
+  by (auto dest: inheritable_instance_methd intro: inheritance overriding)
 qed
 
 lemma inheritable_instance_methd_props: 
@@ -1978,12 +1877,6 @@
 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> 
@@ -2168,14 +2061,14 @@
 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.
@@ -2189,14 +2082,13 @@
 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)
@@ -2237,7 +2129,6 @@
 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);
@@ -2270,7 +2161,7 @@
       "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)
+      by (force dest!: ws_dynmethd accmethd_SomeD)
     with dynlookup eq_mheads 
     show ?thesis 
       by (cases emh type: *) (auto)
@@ -2293,7 +2184,7 @@
       "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)
+      by (force dest: implmt_methd)
     with dynlookup eq_mheads
     show ?thesis 
       by (cases emh type: *) (auto)
@@ -2319,7 +2210,7 @@
 	 "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])
+                  intro: class_Object [OF wf] intro: that)
        with dynlookup eq_mheads
        show ?thesis 
 	 by (cases emh type: *) (auto)
@@ -2364,6 +2255,11 @@
       by (cases emh type: *) (auto dest: accmethd_SomeD)
   qed
 qed
+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)
+*}
 
 (* Tactical version *)
 (*
@@ -2431,7 +2327,7 @@
 done
 *)
 
-(* ### auf ws_class_induct umstellen *)
+(* FIXME occasionally convert to ws_class_induct*) 
 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"
@@ -2465,8 +2361,8 @@
 	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
+	moreover note wf False cls_C 
+	ultimately show ?thesis by (auto intro: hyp [rule_format])
       next
 	case Some
 	from this methd_C m show ?thesis by auto 
@@ -2504,12 +2400,12 @@
                  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 
@@ -2598,7 +2494,7 @@
 by (erule mheads_cases)
    (auto dest: accmethd_SomeD accessible_from_commonD accimethdsD)
 
-lemma static_to_dynamic_accessible_from:
+lemma static_to_dynamic_accessible_from_aux:
 "\<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)
@@ -2615,22 +2511,36 @@
   from stat_acc subclseq 
   show ?thesis (is "?Dyn_accessible m")
   proof (induct rule: accessible_fromR.induct)
-    case (immediate statC m)
+    case (Immediate statC m)
     then show "?Dyn_accessible m"
-      by (blast intro: dyn_accessible_fromR.immediate
+      by (blast intro: dyn_accessible_fromR.Immediate
                        member_inI
                        permits_acc_inheritance)
   next
-    case (overriding _ _ m)
+    case (Overriding _ _ m)
     with wf show "?Dyn_accessible m"
-      by (blast intro: dyn_accessible_fromR.overriding
+      by (blast intro: dyn_accessible_fromR.Overriding
                        member_inI
                        static_to_dynamic_overriding  
                        rtrancl_trancl_trancl 
-                       static_to_dynamic_accessible_from)
+                       static_to_dynamic_accessible_from_aux)
   qed
 qed
 
+lemma static_to_dynamic_accessible_from_static:
+ (assumes stat_acc: "G\<turnstile>m of statC accessible_from accC" and
+            static: "is_static m" and
+                wf: "wf_prog G"
+ ) "G\<turnstile>m in (declclass m) dyn_accessible_from accC"
+proof -
+  from stat_acc wf 
+  have "G\<turnstile>m in statC dyn_accessible_from accC"
+    by (auto intro: static_to_dynamic_accessible_from)
+  from this static
+  show ?thesis
+    by (rule dyn_accessible_from_static_declC)
+qed
+
 lemma dynmethd_member_in:
  (assumes    m: "dynmethd G statC dynC sig = Some m" and
    iscls_statC: "is_class G statC" and
@@ -2723,7 +2633,7 @@
     moreover
     note override eq_dynM_newM
     ultimately show ?thesis
-      by (cases dynM,cases statM) (auto intro: dyn_accessible_fromR.overriding)
+      by (cases dynM,cases statM) (auto intro: dyn_accessible_fromR.Overriding)
   qed
 qed
 
@@ -2749,7 +2659,7 @@
     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 
+    by (auto intro!: dyn_accessible_fromR.Immediate 
               intro: methd_member_of member_of_to_member_in
                      simp add: permits_acc_def)
   with dynM    
@@ -2972,14 +2882,14 @@
             \<Longrightarrow> pid accC = pid (declclass m)"
     (is "?Pack m \<Longrightarrow> ?P m")
   proof (induct rule: dyn_accessible_fromR.induct)
-    case (immediate C m)
+    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)
+    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
@@ -3001,4 +2911,64 @@
   qed
 qed
 
+
+text {* @{text dyn_accessible_instance_field_Protected} only works for fields
+since methods can break the package bounds due to overriding
+*}
+lemma dyn_accessible_instance_field_Protected:
+ (assumes dyn_acc: "G \<turnstile> f in C dyn_accessible_from accC" and
+             prot: "accmodi f = Protected" and
+            field: "is_field f" and
+   instance_field: "\<not> is_static f" and
+          outside: "pid (declclass f) \<noteq> pid accC"
+ ) "G\<turnstile> C \<preceq>\<^sub>C accC"
+proof -
+  from dyn_acc prot field instance_field outside
+  show ?thesis
+  proof (induct)
+    case (Immediate C f)
+    have "G \<turnstile> f in C permits_acc_to accC" .
+    moreover 
+    assume "accmodi f = Protected" and  "is_field f" and "\<not> is_static f" and
+           "pid (declclass f) \<noteq> pid accC"
+    ultimately 
+    show "G\<turnstile> C \<preceq>\<^sub>C accC"
+      by (auto simp add: permits_acc_def)
+  next
+    case Overriding
+    then show ?case by (simp add: is_field_def)
+  qed
+qed
+   
+lemma dyn_accessible_static_field_Protected:
+ (assumes dyn_acc: "G \<turnstile> f in C dyn_accessible_from accC" and
+             prot: "accmodi f = Protected" and
+            field: "is_field f" and
+     static_field: "is_static f" and
+          outside: "pid (declclass f) \<noteq> pid accC"
+ ) "G\<turnstile> accC \<preceq>\<^sub>C declclass f  \<and> G\<turnstile>C \<preceq>\<^sub>C declclass f"
+proof -
+  from dyn_acc prot field static_field outside
+  show ?thesis
+  proof (induct)
+    case (Immediate C f)
+    assume "accmodi f = Protected" and  "is_field f" and "is_static f" and
+           "pid (declclass f) \<noteq> pid accC"
+    moreover 
+    have "G \<turnstile> f in C permits_acc_to accC" .
+    ultimately
+    have "G\<turnstile> accC \<preceq>\<^sub>C declclass f"
+      by (auto simp add: permits_acc_def)
+    moreover
+    have "G \<turnstile> f member_in C" .
+    then have "G\<turnstile>C \<preceq>\<^sub>C declclass f"
+      by (rule member_in_class_relation)
+    ultimately show ?case
+      by blast
+  next
+    case Overriding
+    then show ?case by (simp add: is_field_def)
+  qed
+qed
+
 end
\ No newline at end of file
--- a/src/HOL/Bali/WellType.thy	Thu Feb 21 20:11:32 2002 +0100
+++ b/src/HOL/Bali/WellType.thy	Fri Feb 22 11:26:44 2002 +0100
@@ -153,28 +153,30 @@
  "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"
+"invmode m e \<equiv> if is_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) add: member_is_static_simp)
+done
+
+
+lemma invmode_Static_eq [simp]: "(invmode m e = Static) = is_static m"
+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)"
+lemma invmode_IntVir_eq: "(invmode m e = IntVir) = (\<not>(is_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"
+  "a'=Null \<longrightarrow> (is_static m) \<Longrightarrow> invmode m e = IntVir \<longrightarrow> a' \<noteq> Null"
 apply (clarsimp simp add: invmode_IntVir_eq)
 done
 
@@ -337,7 +339,7 @@
 	  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)"
+		   E,dt\<Turnstile>{cls E,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;
@@ -367,8 +369,8 @@
 					 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)"
+	  accfield (prg E) (cls E) C fn = Some (statDeclC,f)\<rbrakk> \<Longrightarrow>
+			 E,dt\<Turnstile>{cls E,statDeclC,is_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>
@@ -395,7 +397,7 @@
 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  ({accC,statDeclC,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"
@@ -406,7 +408,7 @@
 	"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 ({accC,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"
@@ -463,15 +465,15 @@
  \<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"
+    = {((statDeclC,m),pTs')};rT=(resTy m);accC=cls E;
+ mode = invmode m e\<rbrakk> \<Longrightarrow> E,dt\<Turnstile>{accC,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"
@@ -493,11 +495,12 @@
 \<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 (statDeclC,f);
+  sf=is_static f; fT=(type f); accC=cls E\<rbrakk> 
+\<Longrightarrow> E,dt\<Turnstile>{accC,statDeclC,sf}e..fn\<Colon>=fT"
+by (auto dest: wt.FVar)
 
-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")