"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
authorschirmer
Thu, 31 Oct 2002 18:27:10 +0100
changeset 13688 a0b16d42d489
parent 13687 22dce9134953
child 13689 3d4ad560b2ff
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
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/DefiniteAssignment.thy
src/HOL/Bali/DefiniteAssignmentCorrect.thy
src/HOL/Bali/Eval.thy
src/HOL/Bali/Evaln.thy
src/HOL/Bali/Example.thy
src/HOL/Bali/Name.thy
src/HOL/Bali/State.thy
src/HOL/Bali/Table.thy
src/HOL/Bali/Term.thy
src/HOL/Bali/Trans.thy
src/HOL/Bali/Type.thy
src/HOL/Bali/TypeRel.thy
src/HOL/Bali/TypeSafe.thy
src/HOL/Bali/Value.thy
src/HOL/Bali/WellForm.thy
src/HOL/Bali/WellType.thy
src/HOL/Bali/document/root.tex
--- a/src/HOL/Bali/AxCompl.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/AxCompl.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Bali/AxCompl.thy
     ID:         $Id$
-    Author:     David von Oheimb
+    Author:     David von Oheimb and Norbert Schirmer
     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 *)
 
@@ -19,6 +19,7 @@
 
 
 
+           
 section "set of not yet initialzed classes"
 
 constdefs
@@ -79,8 +80,8 @@
 apply fast
 done
 
-lemma card_Suc_lemma: "\<lbrakk>card (insert a A) \<le> Suc n; a\<notin>A; finite A\<rbrakk> \<Longrightarrow> card A \<le> n"
-apply (rotate_tac 1)
+lemma card_Suc_lemma: 
+  "\<lbrakk>card (insert a A) \<le> Suc n; a\<notin>A; finite A\<rbrakk> \<Longrightarrow> card A \<le> n"
 apply clarsimp
 done
 
@@ -96,7 +97,8 @@
               simp add: nyinitcls_def inited_def split add: split_if_asm)
 done
 
-ML {* bind_thm("inited_gext'",permute_prems 0 1 (thm "inited_gext")) *}
+lemma inited_gext': "\<lbrakk>s\<le>|s';inited C (globs s)\<rbrakk> \<Longrightarrow> inited C (globs s')"
+by (rule inited_gext)
 
 lemma nyinitcls_gext: "snd s\<le>|snd s' \<Longrightarrow> nyinitcls G s' \<subseteq> nyinitcls G s"
 apply (unfold nyinitcls_def)
@@ -124,7 +126,9 @@
 apply auto
 done
 
-lemma All_init_leD: "\<forall>n::nat. G,A\<turnstile>{P \<and>. G\<turnstile>init\<le>n} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
+lemma All_init_leD: 
+ "\<forall>n::nat. G,(A::'a triple set)\<turnstile>{P \<and>. G\<turnstile>init\<le>n} t\<succ> {Q::'a assn} 
+  \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
 apply (drule spec)
 apply (erule conseq1)
 apply clarsimp
@@ -158,7 +162,6 @@
   "{=:n} t\<succ> {G\<rightarrow>} \<equiv> {\<doteq> \<and>. G\<turnstile>init\<le>n} t\<succ> {G\<rightarrow>}"
 
 (* unused *)
-
 lemma MGF_valid: "wf_prog G \<Longrightarrow> G,{}\<Turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
 apply (unfold MGF_def)
 apply (simp add:  ax_valids_def triple_valid_def2)
@@ -178,7 +181,8 @@
 apply fast
 done
 
-lemma MGF_MGFn_iff: "G,A\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} = (\<forall>n. G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>})"
+lemma MGF_MGFn_iff: 
+"G,(A::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} = (\<forall>n. G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>})"
 apply (simp (no_asm_use) add: MGFn_def2 MGF_def)
 apply safe
 apply  (erule_tac [2] All_init_leD)
@@ -187,7 +191,7 @@
 done
 
 lemma MGFnD: 
-"G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>} \<Longrightarrow>  
+"G,(A::state triple set)\<turnstile>{=:n} t\<succ> {G\<rightarrow>} \<Longrightarrow>  
  G,A\<turnstile>{(\<lambda>Y' s' s. s' = s           \<and> P s) \<and>. G\<turnstile>init\<le>n}  
  t\<succ>  {(\<lambda>Y' s' s. G\<turnstile>s\<midarrow>t\<succ>\<rightarrow>(Y',s') \<and> P s) \<and>. G\<turnstile>init\<le>n}"
 apply (unfold init_le_def)
@@ -198,6 +202,10 @@
 done
 lemmas MGFnD' = MGFnD [of _ _ _ _ "\<lambda>x. True"] 
 
+text {* To derive the most general formula, we can always assume a normal
+state in the precondition, since abrupt cases can be handled uniformally by
+the abrupt rule.
+*}
 lemma MGFNormalI: "G,A\<turnstile>{Normal \<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow>  
   G,(A::state triple set)\<turnstile>{\<doteq>::state assn} t\<succ> {G\<rightarrow>}"
 apply (unfold MGF_def)
@@ -208,12 +216,15 @@
 apply (clarsimp simp add: Let_def)
 done
 
-lemma MGFNormalD: "G,A\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow> G,A\<turnstile>{Normal \<doteq>} t\<succ> {G\<rightarrow>}"
+lemma MGFNormalD: 
+"G,(A::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow> G,A\<turnstile>{Normal \<doteq>} t\<succ> {G\<rightarrow>}"
 apply (unfold MGF_def)
 apply (erule conseq1)
 apply clarsimp
 done
 
+text {* Additionally to @{text MGFNormalI}, we also expand the definition of 
+the most general formula here *} 
 lemma MGFn_NormalI: 
 "G,(A::state triple set)\<turnstile>{Normal((\<lambda>Y' s' s. s'=s \<and> normal s) \<and>. G\<turnstile>init\<le>n)}t\<succ> 
  {\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y,s')} \<Longrightarrow> G,A\<turnstile>{=:n}t\<succ>{G\<rightarrow>}"
@@ -225,6 +236,9 @@
 apply (clarsimp simp add: Let_def)
 done
 
+text {* To derive the most general formula, we can restrict ourselves to 
+welltyped terms, since all others can be uniformally handled by the hazard
+rule. *} 
 lemma MGFn_free_wt: 
   "(\<exists>T L C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) 
     \<longrightarrow> G,(A::state triple set)\<turnstile>{=:n} t\<succ> {G\<rightarrow>} 
@@ -234,8 +248,12 @@
 apply (auto elim: conseq12 simp add: MGFn_def MGF_def)
 done
 
+text {* To derive the most general formula, we can restrict ourselves to 
+welltyped terms and assume that the state in the precondition conforms to the
+environment. All type violations can be uniformally handled by the hazard
+rule. *} 
 lemma MGFn_free_wt_NormalConformI: 
-"(\<forall> T L C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T 
+"(\<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> 
@@ -247,389 +265,1129 @@
 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 (erule exE,erule exE, erule exE, erule exE,erule conjE,drule (1) mp,
+       erule conjE)
 apply (drule spec,drule spec, drule spec, drule (1) mp)
 apply (erule conseq12)
 apply blast
 done
 
+text {* To derive the most general formula, we can restrict ourselves to 
+welltyped terms and assume that the state in the precondition conforms to the
+environment and that the term is definetly assigned with respect to this state.
+All type violations can be uniformally handled by the hazard rule. *} 
+lemma MGFn_free_wt_da_NormalConformI: 
+"(\<forall> T L C B. \<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))
+        \<and>. (\<lambda> s. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>B)}
+      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 exE,erule conjE,drule (1) mp,
+       erule conjE)
+apply (drule spec,drule spec, drule spec,drule spec, drule (1) mp)
+apply (erule conseq12)
+apply blast
+done
 
 section "main lemmas"
 
-declare fun_upd_apply [simp del]
-declare splitI2 [rule del] (*prevents ugly renaming of state variables*)
-
-ML_setup {* 
-Delsimprocs [eval_expr_proc, eval_var_proc, eval_exprs_proc, eval_stmt_proc]
-*} (*prevents modifying rhs of MGF*)
-ML {*
-val eval_css = (claset() delrules [thm "eval.Abrupt"] addSIs (thms "eval.intros") 
-                delrules[thm "eval.Expr", thm "eval.Init", thm "eval.Try"] 
-                addIs   [thm "eval.Expr", thm "eval.Init"]
-                addSEs[thm "eval.Try"] delrules[equalityCE],
-                simpset() addsimps [split_paired_all,Let_def]
- addsimprocs [eval_expr_proc,eval_var_proc,eval_exprs_proc,eval_stmt_proc]);
-val eval_Force_tac = force_tac eval_css;
-
-val wt_prepare_tac = EVERY'[
-    rtac (thm "MGFn_free_wt"),
-    clarsimp_tac (claset() addSEs (thms "wt_elim_cases"), simpset())]
-val compl_prepare_tac = EVERY'[rtac (thm "MGFn_NormalI"), Simp_tac]
-val 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]
-*}
-
-lemma MGFn_Init: "\<forall>m. Suc m\<le>n \<longrightarrow> (\<forall>t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>}) \<Longrightarrow>  
-  G,(A::state triple set)\<turnstile>{=:n} In1r (Init C)\<succ> {G\<rightarrow>}"
-apply (tactic "wt_prepare_tac 1")
-(* requires is_class G C two times for nyinitcls *)
-apply (tactic "compl_prepare_tac 1")
-apply (rule_tac C = "initd C" in ax_cases)
-apply  (rule ax_derivs.Done [THEN conseq1])
-apply  (clarsimp intro!: init_done)
-apply (rule_tac y = n in nat.exhaust, clarsimp)
-apply  (rule ax_impossible [THEN conseq1])
-apply  (force dest!: nyinitcls_emptyD)
-apply clarsimp
-apply (drule_tac x = "nat" in spec)
-apply clarsimp
-apply (rule_tac Q = " (\<lambda>Y s' (x,s) . G\<turnstile> (x,init_class_obj G C s) \<midarrow> (if C=Object then Skip else Init (super (the (class G C))))\<rightarrow> s' \<and> x=None \<and> \<not>inited C (globs s)) \<and>. G\<turnstile>init\<le>nat" in ax_derivs.Init)
-apply   simp
-apply  (rule_tac P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s \<and> normal s \<and> \<not> initd C s) \<and>. G\<turnstile>init\<le>nat) " in conseq1)
-prefer 2
-apply   (force elim!: nyinitcls_le_SucD)
-apply  (simp split add: split_if, rule conjI, clarify)
-apply   (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1")
-apply  clarify
-apply  (drule spec)
-apply  (erule MGFnD' [THEN conseq12])
-apply  (tactic "force_tac (claset(), simpset() addsimprocs[eval_stmt_proc]) 1")
-apply (rule allI)
-apply (drule spec)
-apply (erule MGFnD' [THEN conseq12])
-apply clarsimp
-apply (tactic {* pair_tac "sa" 1 *})
-apply (tactic"clarsimp_tac (claset(), simpset() addsimprocs[eval_stmt_proc]) 1")
-apply (rule eval_Init, force+)
-done
+lemma MGFn_Init: 
+ assumes mgf_hyp: "\<forall>m. Suc m\<le>n \<longrightarrow> (\<forall>t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>})"
+ shows "G,(A::state triple set)\<turnstile>{=:n} \<langle>Init C\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+proof (rule MGFn_free_wt [rule_format],elim exE,rule MGFn_NormalI)
+  fix T L accC
+  assume "\<lparr>prg=G, cls=accC, lcl= L\<rparr>\<turnstile>\<langle>Init C\<rangle>\<^sub>s\<Colon>T"
+  hence is_cls: "is_class G C"
+    by cases simp
+  show "G,A\<turnstile>{Normal ((\<lambda>Y' s' s. s' = s \<and> normal s) \<and>. G\<turnstile>init\<le>n)} 
+            .Init C.
+            {\<lambda>Y s' s. G\<turnstile>s \<midarrow>\<langle>Init C\<rangle>\<^sub>s\<succ>\<rightarrow> (Y, s')}"
+       (is "G,A\<turnstile>{Normal ?P} .Init C. {?R}")
+  proof (rule ax_cases [where ?C="initd C"])
+    show "G,A\<turnstile>{Normal ?P  \<and>. initd C} .Init C. {?R}"
+      by (rule ax_derivs.Done [THEN conseq1]) (fastsimp intro: init_done)
+  next
+    have "G,A\<turnstile>{Normal (?P  \<and>. Not \<circ> initd C)} .Init C. {?R}" 
+    proof (cases n)
+      case 0
+      with is_cls
+      show ?thesis
+	by - (rule ax_impossible [THEN conseq1],fastsimp dest: nyinitcls_emptyD)
+    next
+      case (Suc m)
+      with mgf_hyp have mgf_hyp': "\<And> t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>}"
+	by simp
+      from is_cls obtain c where c: "the (class G C) = c"
+	by auto
+      let ?Q= "(\<lambda>Y s' (x,s) . 
+          G\<turnstile> (x,init_class_obj G C s) 
+             \<midarrow> (if C=Object then Skip else Init (super (the (class G C))))\<rightarrow> s'
+          \<and> x=None \<and> \<not>inited C (globs s)) \<and>. G\<turnstile>init\<le>m"
+      from c
+      show ?thesis
+      proof (rule ax_derivs.Init [where ?Q="?Q"])
+	let ?P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s 
+                           \<and> normal s \<and> \<not> initd C s) \<and>. G\<turnstile>init\<le>m)" 
+	show "G,A\<turnstile>{Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
+                  .(if C = Object then Skip else Init (super c)). 
+                  {?Q}"
+	proof (rule conseq1 [where ?P'="?P'"])
+	  show "G,A\<turnstile>{?P'} .(if C = Object then Skip else Init (super c)). {?Q}"
+	  proof (cases "C=Object")
+	    case True
+	    have "G,A\<turnstile>{?P'} .Skip. {?Q}"
+	      by (rule ax_derivs.Skip [THEN conseq1])
+	         (auto simp add: True intro: eval.Skip)
+            with True show ?thesis 
+	      by simp
+	  next
+	    case False
+	    from mgf_hyp'
+	    have "G,A\<turnstile>{?P'} .Init (super c). {?Q}"
+	      by (rule MGFnD' [THEN conseq12]) (fastsimp simp add: False c)
+	    with False show ?thesis
+	      by simp
+	  qed
+	next
+	  from Suc is_cls
+	  show "Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))
+                \<Rightarrow> ?P'"
+	    by (fastsimp elim: nyinitcls_le_SucD)
+	qed
+      next
+	from mgf_hyp'
+	show "\<forall>l. G,A\<turnstile>{?Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty} 
+                      .init c.
+                      {set_lvars l .; ?R}"
+	  apply (rule MGFnD' [THEN conseq12, THEN allI])
+	  apply (clarsimp simp add: split_paired_all)
+	  apply (rule eval.Init [OF c])
+	  apply (insert c)
+	  apply auto
+	  done
+      qed
+    qed
+    thus "G,A\<turnstile>{Normal ?P  \<and>. Not \<circ> initd C} .Init C. {?R}"
+      by clarsimp
+  qed
+qed
 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>};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
+  assumes mgf_methds: 
+           "\<forall>C sig. G,(A::state triple set)\<turnstile>{=:n} \<langle>(Methd C sig)\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+  and mgf_e: "G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+  and mgf_ps: "G,A\<turnstile>{=:n} \<langle>ps\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
+  and wf: "wf_prog G"
+  shows "G,A\<turnstile>{=:n} \<langle>{accC,statT,mode}e\<cdot>mn({pTs'}ps)\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+proof (rule MGFn_free_wt_da_NormalConformI [rule_format],clarsimp) 
+  note inj_term_simps [simp]
+  fix T L accC' E
+  assume wt: "\<lparr>prg=G,cls=accC',lcl = L\<rparr>\<turnstile>\<langle>({accC,statT,mode}e\<cdot>mn( {pTs'}ps))\<rangle>\<^sub>e\<Colon>T"
+  then 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>ps\<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 cases fastsimp+
+  let ?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> 
+              (\<exists> P. normal s1
+                  \<longrightarrow> \<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile>dom (locals (store s1))\<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright>P)) 
+          \<and>. G\<turnstile>init\<le>n \<and>. (\<lambda> s. s\<Colon>\<preceq>(G, L))::state assn"
+  let ?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 = \<lfloor>pvs\<rfloor>\<^sub>l \<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)))::state assn"
 
-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   (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
+  show "G,A\<turnstile>{Normal ((\<lambda>Y' s' s. s' = s \<and> abrupt s = None) \<and>. G\<turnstile>init\<le>n \<and>.
+                     (\<lambda>s. s\<Colon>\<preceq>(G, L)) \<and>.
+                     (\<lambda>s. \<lparr>prg=G, cls=accC',lcl=L\<rparr> \<turnstile> dom (locals (store s)) 
+                           \<guillemotright> \<langle>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)\<rangle>\<^sub>e\<guillemotright> E))}
+             {accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>
+             {\<lambda>Y s' s. \<exists>v. Y = \<lfloor>v\<rfloor>\<^sub>e \<and> 
+                           G\<turnstile>s \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v\<rightarrow> s'}"
+    (is "G,A\<turnstile>{Normal ?P} {accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ> {?S}")
+  proof (rule ax_derivs.Call [where ?Q="?Q" and ?R="?R"])
+    from mgf_e
+    show "G,A\<turnstile>{Normal ?P} e-\<succ> {?Q}"
+    proof (rule MGFnD' [THEN conseq12],clarsimp)
+      fix s0 s1 a
+      assume conf_s0: "Norm s0\<Colon>\<preceq>(G, L)"
+      assume da: "\<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> 
+                     dom (locals s0) \<guillemotright>\<langle>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)\<rangle>\<^sub>e\<guillemotright> E"
+      assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+      show "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT) \<and>
+            (abrupt s1 = None \<longrightarrow>
+              (\<exists>P. \<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P))
+            \<and> s1\<Colon>\<preceq>(G, L)"
+      proof -
+	from da obtain C where
+	  da_e:  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
+                    dom (locals (store ((Norm s0)::state)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
+	  da_ps: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> E" 
+	  by cases (simp add: eq_accC_accC')
+	from eval_e conf_s0 wt_e da_e wf
+	obtain "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT)"
+	  and  "s1\<Colon>\<preceq>(G, L)"
+	  by (rule eval_type_soundE) simp
+	moreover
+	{
+	  assume normal_s1: "normal s1"
+	  have "\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
+	  proof -
+	    from eval_e wt_e da_e wf normal_s1
+	    have "nrm C \<subseteq>  dom (locals (store s1))"
+	      by (cases rule: da_good_approxE') rules
+	    with da_ps show ?thesis
+	      by (rule da_weakenE) rules
+	  qed
+	}
+	ultimately show ?thesis
+	  using eq_accC_accC' by simp
+      qed
+    qed
+  next
+    show "\<forall>a. G,A\<turnstile>{?Q\<leftarrow>In1 a} ps\<doteq>\<succ> {?R a}" (is "\<forall> a. ?PS a")
+    proof 
+      fix a  
+      show "?PS a"
+      proof (rule MGFnD' [OF mgf_ps, THEN conseq12],
+             clarsimp simp add: eq_accC_accC' [symmetric])
+	fix s0 s1 s2 vs
+	assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
+	assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+	assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+	assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
+	assume da_ps: "abrupt s1 = None \<longrightarrow> 
+                       (\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+                               dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P)"
+	show "(\<exists>s1. G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
+                (abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT) \<and>
+                G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2) \<and>
+              s2\<Colon>\<preceq>(G, L)"
+	proof (cases "normal s1")
+	  case True
+	  with da_ps obtain P where
+	   "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
+	    by auto
+	  from eval_ps conf_s1 wt_args this wf
+	  have "s2\<Colon>\<preceq>(G, L)"
+	    by (rule eval_type_soundE)
+	  with eval_e conf_a eval_ps 
+	  show ?thesis 
+	    by auto
+	next
+	  case False
+	  with eval_ps have "s2=s1" by auto
+	  with eval_e conf_a eval_ps conf_s1 
+	  show ?thesis 
+	    by auto
+	qed
+      qed
+    qed
+  next
+    show "\<forall>a vs invC declC l.
+      G,A\<turnstile>{?R a\<leftarrow>\<lfloor>vs\<rfloor>\<^sub>l \<and>.
+             (\<lambda>s. declC =
+                  invocation_declclass G mode (store s) a statT
+                      \<lparr>name=mn, parTs=pTs'\<rparr> \<and>
+                  invC = invocation_class mode (store s) a statT \<and>
+                  l = locals (store s)) ;.
+             init_lvars G declC \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs \<and>.
+             (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)}
+          Methd declC \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ> 
+          {set_lvars l .; ?S}" 
+      (is "\<forall> a vs invC declC l. ?METHD a vs invC declC l")
+    proof (intro allI)
+      fix a vs invC declC l
+      from mgf_methds [rule_format]
+      show "?METHD a vs invC declC l"
+      proof (rule MGFnD' [THEN conseq12],clarsimp)
+	fix s4 s2 s1::state
+	fix s0 v
+	let ?D= "invocation_declclass G mode (store s2) a statT 
+                    \<lparr>name=mn,parTs=pTs'\<rparr>"
+	let ?s3= "init_lvars G ?D \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs s2"
+	assume inv_prop: "abrupt ?s3=None 
+             \<longrightarrow> G\<turnstile>mode\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
+	assume conf_s2: "s2\<Colon>\<preceq>(G, L)"
+	assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+	assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+	assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
+	assume eval_mthd: "G\<turnstile>?s3 \<midarrow>Methd ?D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+	show "G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v
+                        \<rightarrow> (set_lvars (locals (store s2))) s4"
+	proof -
+	  obtain D where D: "D=?D" by simp
+	  obtain s3 where s3: "s3=?s3" by simp
+	  obtain s3' where 
+	    s3': "s3' = check_method_access G accC statT mode 
+                           \<lparr>name=mn,parTs=pTs'\<rparr> a s3"
+	    by simp
+	  have eq_s3'_s3: "s3'=s3"
+	  proof -
+	    from inv_prop s3 mode
+	    have "normal s3 \<Longrightarrow> 
+             G\<turnstile>invmode statM e\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
+	      by auto
+	    with eval_ps wt_e statM conf_s2 conf_a [rule_format] 
+	    have "check_method_access G accC statT (invmode statM e)
+                      \<lparr>name=mn,parTs=pTs'\<rparr> a s3 = s3"
+	      by (rule error_free_call_access) (auto simp add: s3 mode wf)
+	    thus ?thesis 
+	      by (simp add: s3' mode)
+	  qed
+	  with eval_mthd D s3
+	  have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+	    by simp
+	  with eval_e eval_ps D _ s3' 
+	  show ?thesis
+	    by (rule eval_Call) (auto simp add: s3 mode D)
+	qed
+      qed
+    qed
+  qed
+qed
+	  	  
+lemma eval_expression_no_jump':
+  assumes eval: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<rightarrow> s1"
+  and   no_jmp: "abrupt s0 \<noteq> Some (Jump j)"
+  and      wt: "\<lparr>prg=G, cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-T" 
+  and      wf: "wf_prog G"
+shows "abrupt s1 \<noteq> Some (Jump j)"
+using eval no_jmp wt wf
+by - (rule eval_expression_no_jump 
+            [where ?Env="\<lparr>prg=G, cls=C,lcl=L\<rparr>",simplified],auto)
 
-lemma MGF_altern: "G,A\<turnstile>{Normal (\<doteq> \<and>. p)} t\<succ> {G\<rightarrow>} =  
- G,A\<turnstile>{Normal ((\<lambda>Y s Z. \<forall>w s'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<longrightarrow> (w,s') = Z) \<and>. p)} 
-  t\<succ> {\<lambda>Y s Z. (Y,s) = Z}"
-apply (unfold MGF_def)
-apply (auto del: conjI elim!: conseq12)
-apply (case_tac "\<exists>w s. G\<turnstile>Norm sa \<midarrow>t\<succ>\<rightarrow> (w,s) ")
-apply  (fast dest: unique_eval)
-apply clarsimp
-apply (drule split_paired_All [THEN subst])
-apply (clarsimp elim!: state_not_single)
-done
+
+text {* To derive the most general formula for the loop statement, we need to
+come up with a proper loop invariant, which intuitively states that we are 
+currently inside the evaluation of the loop. To define such an invariant, we
+unroll the loop in iterated evaluations of the expression and evaluations of
+the loop body. *}
+
+constdefs
+ unroll:: "prog \<Rightarrow> label \<Rightarrow> expr \<Rightarrow> stmt \<Rightarrow> (state \<times>  state) set"
+
+ "unroll G l e c \<equiv> {(s,t). \<exists> v s1 s2.
+                             G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s1 \<and> the_Bool v \<and> normal s1 \<and>
+                             G\<turnstile>s1 \<midarrow>c\<rightarrow> s2 \<and> t=(abupd (absorb (Cont l)) s2)}"
+
+
+lemma unroll_while:
+  assumes unroll: "(s, t) \<in> (unroll G l e c)\<^sup>*"
+  and     eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'" 
+  and     normal_termination: "normal s'  \<longrightarrow> \<not> the_Bool v"
+  and     wt: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
+  and     wf: "wf_prog G" 
+  shows "G\<turnstile>s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+using unroll (* normal_s *)
+proof (induct rule: converse_rtrancl_induct) 
+  show "G\<turnstile>t \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+  proof (cases "normal t")
+    case False
+    with eval_e have "s'=t" by auto
+    with False show ?thesis by auto
+  next
+    case True
+    note normal_t = this
+    show ?thesis
+    proof (cases "normal s'")
+      case True
+      with normal_t eval_e normal_termination
+      show ?thesis
+	by (auto intro: eval.Loop)
+    next
+      case False
+      note abrupt_s' = this
+      from eval_e _ wt wf
+      have no_cont: "abrupt s' \<noteq> Some (Jump (Cont l))"
+	by (rule eval_expression_no_jump') (insert normal_t,simp)
+      have
+	"if the_Bool v 
+             then (G\<turnstile>s' \<midarrow>c\<rightarrow> s' \<and> 
+                   G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s')
+	     else s' = s'"
+      proof (cases "the_Bool v")
+	case False thus ?thesis by simp
+      next
+	case True
+	with abrupt_s' have "G\<turnstile>s' \<midarrow>c\<rightarrow> s'" by auto
+	moreover from abrupt_s' no_cont 
+	have no_absorb: "(abupd (absorb (Cont l)) s')=s'"
+	  by (cases s') (simp add: absorb_def split: split_if)
+	moreover
+	from no_absorb abrupt_s'
+	have "G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+	  by auto
+	ultimately show ?thesis
+	  using True by simp
+      qed
+      with eval_e 
+      show ?thesis
+	using normal_t by (auto intro: eval.Loop)
+    qed
+  qed
+next
+  fix s s3
+  assume unroll: "(s,s3) \<in> unroll G l e c"
+  assume while: "G\<turnstile>s3 \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+  show "G\<turnstile>s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+  proof -
+    from unroll obtain v s1 s2 where
+      normal_s1: "normal s1" and
+      eval_e: "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s1" and
+      continue: "the_Bool v" and
+      eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" and
+      s3: "s3=(abupd (absorb (Cont l)) s2)"
+      by  (unfold unroll_def) fast 
+    from eval_e normal_s1 have
+      "normal s"
+      by (rule eval_no_abrupt_lemma [rule_format])
+    with while eval_e continue eval_c s3 show ?thesis
+      by (auto intro!: eval.Loop)
+  qed
+qed
 
 
-lemma MGFn_Loop: 
-"\<lbrakk>G,(A::state triple set)\<turnstile>{=:n} In1l expr\<succ> {G\<rightarrow>};G,A\<turnstile>{=:n} In1r stmnt\<succ> {G\<rightarrow>} \<rbrakk> 
-\<Longrightarrow> 
-  G,A\<turnstile>{=:n} In1r (l\<bullet> While(expr) stmnt)\<succ> {G\<rightarrow>}"
-apply (rule MGFn_NormalI, simp)
-apply (rule_tac p2 = "\<lambda>s. card (nyinitcls G s) \<le> n" in 
-          MGF_altern [unfolded MGF_def, THEN iffD2, THEN conseq1])
-prefer 2
-apply  clarsimp
-apply (rule_tac P' = 
-"((\<lambda>Y s Z. \<forall>w s'. G\<turnstile>s \<midarrow>In1r (l\<bullet>  While(expr) stmnt) \<succ>\<rightarrow> (w,s') \<longrightarrow> (w,s') = Z) 
-  \<and>. (\<lambda>s. card (nyinitcls G s) \<le> n))" in conseq12)
-prefer 2
-apply  clarsimp
-apply  (tactic "smp_tac 1 1", erule_tac V = "All ?P" in thin_rl)
-apply  (rule_tac [2] P' = " (\<lambda>b s (Y',s') . (\<exists>s0. G\<turnstile>s0 \<midarrow>In1l expr\<succ>\<rightarrow> (b,s)) \<and> (if normal s \<and> the_Bool (the_In1 b) then (\<forall>s'' w s0. G\<turnstile>s \<midarrow>stmnt\<rightarrow> s'' \<and> G\<turnstile>(abupd (absorb (Cont l)) s'') \<midarrow>In1r (l\<bullet> While(expr) stmnt) \<succ>\<rightarrow> (w,s0) \<longrightarrow> (w,s0) = (Y',s')) else (\<diamondsuit>,s) = (Y',s'))) \<and>. G\<turnstile>init\<le>n" in polymorphic_Loop)
-apply   (force dest!: eval.Loop split add: split_if_asm)
-prefer 2
-apply  (erule MGFnD' [THEN conseq12])
-apply  clarsimp
-apply  (erule_tac V = "card (nyinitcls G s') \<le> n" in thin_rl)
-apply  (tactic "eval_Force_tac 1")
-apply (erule MGFnD' [THEN conseq12] , clarsimp)
-apply (rule conjI, erule exI)
-apply (tactic "clarsimp_tac eval_css 1")
-apply (case_tac "a")
-prefer 2
-apply  (clarsimp)
-apply (clarsimp split add: split_if)
-apply (rule conjI, (tactic {* force_tac (claset() addSDs [thm "eval.Loop"],
-  simpset() addsimps [split_paired_all] addsimprocs [eval_stmt_proc]) 1*})+)
-done
+ML"Addsimprocs [eval_expr_proc, eval_var_proc, eval_exprs_proc, eval_stmt_proc]"
+  
+lemma MGFn_Loop:
+  assumes mfg_e: "G,(A::state triple set)\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+  and     mfg_c: "G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+  and     wf: "wf_prog G" 
+shows "G,A\<turnstile>{=:n} \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+proof (rule MGFn_free_wt [rule_format],elim exE)
+  fix T L C
+  assume wt: "\<lparr>prg = G, cls = C, lcl = L\<rparr>\<turnstile>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<Colon>T"
+  then obtain eT where
+    wt_e: "\<lparr>prg = G, cls = C, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
+    by cases simp
+  show ?thesis
+  proof (rule MGFn_NormalI)
+    show "G,A\<turnstile>{Normal ((\<lambda>Y' s' s. s' = s \<and> normal s) \<and>. G\<turnstile>init\<le>n)} 
+              .l\<bullet> While(e) c.
+              {\<lambda>Y s' s. G\<turnstile>s \<midarrow>In1r (l\<bullet> While(e) c)\<succ>\<rightarrow> (Y, s')}"
+    proof (rule conseq12 
+           [where ?P'="(\<lambda> Y s' s. (s,s') \<in> (unroll G l e c)\<^sup>* ) \<and>. G\<turnstile>init\<le>n"
+             and  ?Q'="((\<lambda> Y s' s. (\<exists> t b. (s,t) \<in> (unroll G l e c)\<^sup>* \<and> 
+                          Y=\<lfloor>b\<rfloor>\<^sub>e \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s')) 
+                        \<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>"])
+      show  "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n} 
+                  .l\<bullet> While(e) c.
+                 {((\<lambda>Y s' s. (\<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and> 
+                                  Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s')) 
+                              \<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>}"
+      proof (rule ax_derivs.Loop)
+	from mfg_e
+	show "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n} 
+                   e-\<succ>
+                  {(\<lambda>Y s' s. (\<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and> 
+                                     Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s')) 
+                   \<and>. G\<turnstile>init\<le>n}"
+	proof (rule MGFnD' [THEN conseq12],clarsimp)
+	  fix s Z s' v
+	  assume "(Z, s) \<in> (unroll G l e c)\<^sup>*"
+	  moreover
+	  assume "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s'"
+	  ultimately
+	  show "\<exists>t. (Z, t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'"
+	    by blast
+	qed
+      next
+	from mfg_c
+	show "G,A\<turnstile>{Normal (((\<lambda>Y s' s. \<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
+                                       Y = \<lfloor>b\<rfloor>\<^sub>e \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s') 
+                          \<and>. G\<turnstile>init\<le>n)\<leftarrow>=True)}
+                  .c.
+                  {abupd (absorb (Cont l)) .;
+                   ((\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n)}"
+	proof (rule MGFnD' [THEN conseq12],clarsimp)
+	  fix Z s' s v t
+	  assume unroll: "(Z, t) \<in> (unroll G l e c)\<^sup>*"
+	  assume eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> Norm s" 
+	  assume true: "the_Bool v"
+	  assume eval_c: "G\<turnstile>Norm s \<midarrow>c\<rightarrow> s'"
+	  show "(Z, abupd (absorb (Cont l)) s') \<in> (unroll G l e c)\<^sup>*"
+	  proof -
+	    note unroll
+	    also
+	    from eval_e true eval_c
+	    have "(t,abupd (absorb (Cont l)) s') \<in> unroll G l e c" 
+	      by (unfold unroll_def) force
+	    ultimately show ?thesis ..
+	  qed
+	qed
+      qed
+    next
+      show 
+	"\<forall>Y s Z.
+         (Normal ((\<lambda>Y' s' s. s' = s \<and> normal s) \<and>. G\<turnstile>init\<le>n)) Y s Z 
+         \<longrightarrow> (\<forall>Y' s'.
+               (\<forall>Y Z'. 
+                 ((\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n) Y s Z' 
+                 \<longrightarrow> (((\<lambda>Y s' s. \<exists>t b. (s,t) \<in> (unroll G l e c)\<^sup>* 
+                                       \<and> Y=\<lfloor>b\<rfloor>\<^sub>e \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s') 
+                     \<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>) Y' s' Z') 
+               \<longrightarrow> G\<turnstile>Z \<midarrow>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ>\<rightarrow> (Y', s'))"
+      proof (clarsimp)
+	fix Y' s' s
+	assume asm:
+	  "\<forall>Z'. (Z', Norm s) \<in> (unroll G l e c)\<^sup>* 
+                 \<longrightarrow> card (nyinitcls G s') \<le> n \<and>
+                     (\<exists>v. (\<exists>t. (Z', t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s') \<and>
+                     (fst s' = None \<longrightarrow> \<not> the_Bool v)) \<and> Y' = \<diamondsuit>"
+	show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+	proof -
+	  from asm obtain v t where 
+	    -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}  
+	    unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
+            eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'" and
+            normal_termination: "normal s' \<longrightarrow> \<not> the_Bool v" and
+	     Y': "Y' = \<diamondsuit>"
+	    by auto
+	  from unroll eval_e normal_termination wt_e wf
+	  have "G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+	    by (rule unroll_while)
+	  with Y' 
+	  show ?thesis
+	    by simp
+	qed
+      qed
+    qed
+  qed
+qed
 
-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
+  fixes A :: "state triple set"
+ assumes mgf_init: "G,A\<turnstile>{=:n} \<langle>Init statDeclC\<rangle>\<^sub>s\<succ> {G\<rightarrow>}" 
+  and    mgf_e: "G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+  and    wf: "wf_prog G"
+  shows "G,A\<turnstile>{=:n} \<langle>{accC,statDeclC,stat}e..fn\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
+proof (rule MGFn_free_wt_da_NormalConformI [rule_format],clarsimp) 
+  note inj_term_simps [simp]
+  fix T L accC' V
+  assume wt: "\<lparr>prg = G, cls = accC', lcl = L\<rparr>\<turnstile>\<langle>{accC,statDeclC,stat}e..fn\<rangle>\<^sub>v\<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'" and
+    stat: "stat=is_static  f"
+    by (cases) (auto simp add: member_is_static_simp)
+  let ?Q="(\<lambda>Y s1 (x,s) . x = None \<and> 
+                (G\<turnstile>Norm s \<midarrow>Init statDeclC\<rightarrow> s1) \<and>
+                (\<exists> E. \<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile>dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E))
+                \<and>. G\<turnstile>init\<le>n \<and>. (\<lambda> s. s\<Colon>\<preceq>(G, L))"
+  show "G,A\<turnstile>{Normal
+             ((\<lambda>Y' s' s. s' = s \<and> abrupt s = None) \<and>. G\<turnstile>init\<le>n \<and>.
+              (\<lambda>s. s\<Colon>\<preceq>(G, L)) \<and>.
+              (\<lambda>s. \<lparr>prg=G,cls=accC',lcl=L\<rparr>
+                 \<turnstile> dom (locals (store s)) \<guillemotright> \<langle>{accC,statDeclC,stat}e..fn\<rangle>\<^sub>v\<guillemotright> V))
+             } {accC,statDeclC,stat}e..fn=\<succ>
+             {\<lambda>Y s' s. \<exists>vf. Y = \<lfloor>vf\<rfloor>\<^sub>v \<and> 
+                        G\<turnstile>s \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>vf\<rightarrow> s'}"
+    (is "G,A\<turnstile>{Normal ?P} {accC,statDeclC,stat}e..fn=\<succ> {?R}")
+  proof (rule ax_derivs.FVar [where ?Q="?Q" ])
+    from mgf_init
+    show "G,A\<turnstile>{Normal ?P} .Init statDeclC. {?Q}"
+    proof (rule MGFnD' [THEN conseq12],clarsimp)
+      fix s s'
+      assume conf_s: "Norm s\<Colon>\<preceq>(G, L)"
+      assume da: "\<lparr>prg=G,cls=accC',lcl=L\<rparr>
+                    \<turnstile> dom (locals s) \<guillemotright>\<langle>{accC,statDeclC,stat}e..fn\<rangle>\<^sub>v\<guillemotright> V"
+      assume eval_init: "G\<turnstile>Norm s \<midarrow>Init statDeclC\<rightarrow> s'"
+      show "(\<exists>E. \<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E) \<and>
+            s'\<Colon>\<preceq>(G, L)"
+      proof -
+	from da 
+	obtain E where
+	  "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals s) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+	  by cases simp
+	moreover
+	from eval_init
+	have "dom (locals s) \<subseteq> dom (locals (store s'))"
+	  by (rule dom_locals_eval_mono [elim_format]) simp
+	ultimately obtain E' where
+	  "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
+	  by (rule da_weakenE)
+	moreover
+	have "s'\<Colon>\<preceq>(G, L)"
+	proof -
+	  have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
+	  proof -
+	    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)
+	    thus ?thesis by simp
+	  qed
+	  obtain I where 
+	    da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+               \<turnstile> dom (locals (store ((Norm s)::state))) \<guillemotright>\<langle>Init statDeclC\<rangle>\<^sub>s\<guillemotright> I"
+	    by (auto intro: da_Init [simplified] assigned.select_convs)
+	  from eval_init conf_s wt_init da_init  wf
+	  show ?thesis
+	    by (rule eval_type_soundE)
+	qed
+	ultimately show ?thesis by rules
+      qed
+    qed
+  next
+    from mgf_e
+    show "G,A\<turnstile>{?Q} e-\<succ> {\<lambda>Val:a:. fvar statDeclC stat fn a ..; ?R}"
+    proof (rule MGFnD' [THEN conseq12],clarsimp)
+      fix s0 s1 s2 E a
+      let ?fvar = "fvar statDeclC stat fn a s2"
+      assume eval_init: "G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1"
+      assume eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2"
+      assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
+      assume da_e: "\<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+      show "G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>fst ?fvar\<rightarrow> snd ?fvar"
+      proof -
+	obtain v s2' where
+	  v: "v=fst ?fvar" and s2': "s2'=snd ?fvar"
+	  by simp
+	obtain s3 where
+	  s3: "s3= check_field_access G accC' statDeclC fn stat a s2'"
+	  by simp
+	have eq_s3_s2': "s3=s2'"
+	proof -
+	  from eval_e conf_s1 wt_e da_e wf obtain
+	    conf_s2: "s2\<Colon>\<preceq>(G, L)"  and
+	    conf_a: "normal s2 \<Longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+	    by (rule eval_type_soundE) simp
+	  from accfield wt_e eval_init eval_e conf_s2 conf_a _ wf
+	  show ?thesis
+	    by (rule  error_free_field_access 
+                      [where ?v=v and ?s2'=s2',elim_format])
+	       (simp add: s3 v s2' stat)+
+        qed
+	from eval_init eval_e 
+	show ?thesis
+	  apply (rule eval.FVar [where ?s2'=s2'])
+	  apply  (simp add: s2')
+	  apply  (simp add: s3 [symmetric]   eq_s3_s2' eq_accC s2' [symmetric])
+	  done
+      qed
+    qed
+  qed
+qed
 
-text {* For @{text MGFn_Fin} we need the wellformedness of the program to
-switch from the evaln-semantics to the eval-semantics *}
-lemma MGFn_Fin: 
-"\<lbrakk>wf_prog G; G,A\<turnstile>{=:n} In1r stmt1\<succ> {G\<rightarrow>}; G,A\<turnstile>{=:n} In1r stmt2\<succ> {G\<rightarrow>}\<rbrakk>
- \<Longrightarrow> G,(A\<Colon>state triple set)\<turnstile>{=:n} In1r (stmt1 Finally stmt2)\<succ> {G\<rightarrow>}"
-apply (tactic "wt_conf_prepare_tac 1")
-apply (rule_tac Q = " (\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>stmt1\<rightarrow> s' \<and> s\<Colon>\<preceq>(G, L)) 
-\<and>. G\<turnstile>init\<le>n" in ax_derivs.Fin)
-apply (tactic "forw_hyp_tac 1")
-apply (tactic "clarsimp_tac eval_css 1")
-apply (rule allI)
-apply (tactic "clarsimp_tac eval_css 1")
-apply (tactic "forw_hyp_tac 1")
-apply (tactic {* pair_tac "sb" 1 *})
-apply (tactic"clarsimp_tac (claset(),simpset() addsimprocs [eval_stmt_proc]) 1")
-apply (rule wf_eval_Fin)
-apply auto
-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)]: 
- "\<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)
-apply (drule_tac psi = "All ?P" in asm_rl)
-apply (subgoal_tac "\<forall>v e c es. G,A\<turnstile>{=:n} In2 v\<succ> {G\<rightarrow>} \<and> G,A\<turnstile>{=:n} In1l e\<succ> {G\<rightarrow>} \<and> G,A\<turnstile>{=:n} In1r c\<succ> {G\<rightarrow>} \<and> G,A\<turnstile>{=:n} In3 es\<succ> {G\<rightarrow>}")
-apply  (tactic "Clarify_tac 2")
-apply  (induct_tac "t")
-apply    (induct_tac "a")
-apply     fast+
-apply (rule var_expr_stmt.induct) 
-(* 34 subgoals *)
-prefer 17 apply fast (* Methd *)
-prefer 16 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 [29] MGFn_Init)
-prefer 23 apply (erule (1) MGFn_Loop)
-prefer 26 apply (erule (2) MGFn_Fin)
-apply (tactic "ALLGOALS compl_prepare_tac")
+lemma MGFn_Fin:
+  assumes wf: "wf_prog G" 
+  and     mgf_c1: "G,A\<turnstile>{=:n} \<langle>c1\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+  and     mgf_c2: "G,A\<turnstile>{=:n} \<langle>c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+  shows "G,(A\<Colon>state triple set)\<turnstile>{=:n} \<langle>c1 Finally c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+proof (rule MGFn_free_wt_da_NormalConformI [rule_format],clarsimp)
+  fix T L accC C 
+  assume 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 cases simp
+  let  ?Q = "(\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>c1\<rightarrow> s' \<and> 
+               (\<exists> C1. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1)
+               \<and> s\<Colon>\<preceq>(G, L)) 
+             \<and>. G\<turnstile>init\<le>n"
+  show "G,A\<turnstile>{Normal
+              ((\<lambda>Y' s' s. s' = s \<and> abrupt s = None) \<and>. G\<turnstile>init\<le>n \<and>.
+              (\<lambda>s. s\<Colon>\<preceq>(G, L)) \<and>.
+              (\<lambda>s. \<lparr>prg=G,cls=accC,lcl =L\<rparr>  
+                     \<turnstile>dom (locals (store s)) \<guillemotright>\<langle>c1 Finally c2\<rangle>\<^sub>s\<guillemotright> C))}
+             .c1 Finally c2. 
+             {\<lambda>Y s' s. Y = \<diamondsuit> \<and> G\<turnstile>s \<midarrow>c1 Finally c2\<rightarrow> s'}"
+    (is "G,A\<turnstile>{Normal ?P} .c1 Finally c2. {?R}")
+  proof (rule ax_derivs.Fin [where ?Q="?Q"])
+    from mgf_c1
+    show "G,A\<turnstile>{Normal ?P} .c1. {?Q}"
+    proof (rule MGFnD' [THEN conseq12],clarsimp)
+      fix s0
+      assume "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1 Finally c2\<rangle>\<^sub>s\<guillemotright> C"
+      thus "\<exists>C1. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1"
+	by cases (auto simp add: inj_term_simps)
+    qed
+  next
+    from mgf_c2
+    show "\<forall>abr. G,A\<turnstile>{?Q \<and>. (\<lambda>s. abr = abrupt s) ;. abupd (\<lambda>abr. None)} .c2.
+          {abupd (abrupt_if (abr \<noteq> None) abr) .; ?R}"
+    proof (rule MGFnD' [THEN conseq12, THEN allI],clarsimp)
+      fix s0 s1 s2 C1
+      assume da_c1:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1"
+      assume conf_s0: "Norm s0\<Colon>\<preceq>(G, L)"
+      assume eval_c1: "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1"
+      assume eval_c2: "G\<turnstile>abupd (\<lambda>abr. None) s1 \<midarrow>c2\<rightarrow> s2"
+      show "G\<turnstile>Norm s0 \<midarrow>c1 Finally c2
+               \<rightarrow> abupd (abrupt_if (\<exists>y. abrupt s1 = Some y) (abrupt s1)) s2"
+      proof -
+	obtain abr1 str1 where s1: "s1=(abr1,str1)"
+	  by (cases s1) simp
+	with eval_c1 eval_c2 obtain
+	  eval_c1': "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (abr1,str1)" and
+	  eval_c2': "G\<turnstile>Norm str1 \<midarrow>c2\<rightarrow> s2"
+	  by simp
+	obtain s3 where 
+	  s3: "s3 = (if \<exists>err. abr1 = Some (Error err) 
+	                then (abr1, str1)
+                        else abupd (abrupt_if (abr1 \<noteq> None) abr1) s2)"
+	  by simp
+	from eval_c1' conf_s0 wt_c1 _ wf 
+	have "error_free (abr1,str1)"
+	  by (rule eval_type_soundE) (insert da_c1,auto)
+	with s3 have eq_s3: "s3=abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
+	  by (simp add: error_free_def)
+	from eval_c1' eval_c2' s3
+	show ?thesis
+	  by (rule eval.Fin [elim_format]) (simp add: s1 eq_s3)
+      qed
+    qed 
+  qed
+qed
+      
+lemma Body_no_break:
+ assumes eval_init: "G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1" 
+   and      eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" 
+   and       jmpOk: "jumpNestingOkS {Ret} c"
+   and        wt_c: "\<lparr>prg=G, cls=C, lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>"
+   and        clsD: "class G D=Some d"
+   and          wf: "wf_prog G" 
+  shows "\<forall> l. abrupt s2 \<noteq> Some (Jump (Break l)) \<and> 
+              abrupt s2 \<noteq> Some (Jump (Cont l))"
+proof
+  fix l show "abrupt s2 \<noteq> Some (Jump (Break l)) \<and>  
+              abrupt s2 \<noteq> Some (Jump (Cont l))"
+  proof -
+    from clsD have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init D)\<Colon>\<surd>"
+      by auto
+    from eval_init wf
+    have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+      by - (rule eval_statement_no_jump [OF _ _ _ wt_init],auto)
+    from eval_c _ wt_c wf
+    show ?thesis
+      apply (rule jumpNestingOk_eval [THEN conjE, elim_format])
+      using jmpOk s1_no_jmp
+      apply auto
+      done
+  qed
+qed
 
-apply (rule ax_derivs.LVar [THEN conseq1], tactic "eval_Force_tac 1")
-
-apply (rule ax_derivs.AVar)
-apply  (erule MGFnD [THEN ax_NormalD])
-apply (tactic "forw_hyp_eval_Force_tac 1")
-
-apply (rule ax_derivs.InstInitV)
-
-apply (rule ax_derivs.NewC)
-apply (erule MGFn_InitD [THEN conseq2])
-apply (tactic "eval_Force_tac 1")
-
-apply (rule_tac Q = "(\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>In1r (init_comp_ty ty) \<succ>\<rightarrow> (Y',s')) \<and>. G\<turnstile>init\<le>n" in ax_derivs.NewA)
-apply  (simp add: init_comp_ty_def split add: split_if)
-apply   (rule conjI, clarsimp)
-apply   (erule MGFn_InitD [THEN conseq2])
-apply   (tactic "clarsimp_tac eval_css 1")
-apply  clarsimp
-apply  (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1")
-apply (tactic "forw_hyp_eval_Force_tac 1")
-
-apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast],tactic"eval_Force_tac 1")
-
-apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst],tactic"eval_Force_tac 1")
-apply (rule ax_derivs.Lit [THEN conseq1], tactic "eval_Force_tac 1")
-apply (rule ax_derivs.UnOp, tactic "forw_hyp_eval_Force_tac 1")
+lemma MGFn_Body:
+  assumes wf: "wf_prog G"
+  and     mgf_init: "G,A\<turnstile>{=:n} \<langle>Init D\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+  and     mgf_c: "G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+  shows  "G,(A\<Colon>state triple set)\<turnstile>{=:n} \<langle>Body D c\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+proof (rule MGFn_free_wt_da_NormalConformI [rule_format],clarsimp)
+  fix T L accC E
+  assume wt: "\<lparr>prg=G, cls=accC,lcl=L\<rparr>\<turnstile>\<langle>Body D c\<rangle>\<^sub>e\<Colon>T"
+  let ?Q="(\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>Init D\<rightarrow> s' \<and> jumpNestingOkS {Ret} c) 
+          \<and>. G\<turnstile>init\<le>n" 
+  show "G,A\<turnstile>{Normal
+               ((\<lambda>Y' s' s. s' = s \<and> fst s = None) \<and>. G\<turnstile>init\<le>n \<and>.
+                (\<lambda>s. s\<Colon>\<preceq>(G, L)) \<and>.
+                (\<lambda>s. \<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                       \<turnstile> dom (locals (store s)) \<guillemotright>\<langle>Body D c\<rangle>\<^sub>e\<guillemotright> E))}
+             Body D c-\<succ> 
+             {\<lambda>Y s' s. \<exists>v. Y = In1 v \<and> G\<turnstile>s \<midarrow>Body D c-\<succ>v\<rightarrow> s'}"
+    (is "G,A\<turnstile>{Normal ?P} Body D c-\<succ> {?R}")
+  proof (rule ax_derivs.Body [where ?Q="?Q"])
+    from mgf_init
+    show "G,A\<turnstile>{Normal ?P} .Init D. {?Q}"
+    proof (rule MGFnD' [THEN conseq12],clarsimp)
+      fix s0
+      assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>Body D c\<rangle>\<^sub>e\<guillemotright> E"
+      thus "jumpNestingOkS {Ret} c"
+	by cases simp
+    qed
+  next
+    from mgf_c
+    show "G,A\<turnstile>{?Q}.c.{\<lambda>s.. abupd (absorb Ret) .; ?R\<leftarrow>\<lfloor>the (locals s Result)\<rfloor>\<^sub>e}"
+    proof (rule MGFnD' [THEN conseq12],clarsimp)
+      fix s0 s1 s2
+      assume eval_init: "G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1"
+      assume eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2"
+      assume nestingOk: "jumpNestingOkS {Ret} c"
+      show "G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
+              \<rightarrow> abupd (absorb Ret) s2"
+      proof -
+	from wt obtain d where 
+          d: "class G D=Some d" and
+          wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
+	  by cases auto
+	obtain s3 where 
+	  s3: "s3= (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
+                           fst s2 = Some (Jump (Cont l))
+                       then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 
+                       else s2)"
+	  by simp
+	from eval_init eval_c nestingOk wt_c d wf
+	have eq_s3_s2: "s3=s2"
+	  by (rule Body_no_break [elim_format]) (simp add: s3)
+	from eval_init eval_c s3
+	show ?thesis
+	  by (rule eval.Body [elim_format]) (simp add: eq_s3_s2)
+      qed
+    qed
+  qed
+qed
 
-apply (rule ax_derivs.BinOp)
-apply   (erule MGFnD [THEN ax_NormalD])
-
-apply   (rule allI)
-apply   (case_tac "need_second_arg binop v1")
-apply     simp
-apply     (tactic "forw_hyp_eval_Force_tac 1")
-
-apply     simp
-apply     (rule ax_Normal_cases)
-apply       (rule ax_derivs.Skip [THEN conseq1])
-apply       clarsimp
-
-apply       (rule eval_BinOp_arg2_indepI)
-apply       simp
-apply       simp
-
-apply  (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
-apply  (tactic "eval_Force_tac 1")
-
-apply (rule ax_derivs.Super [THEN conseq1], tactic "eval_Force_tac 1")
-apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc],tactic"eval_Force_tac 1")
-
-apply (rule ax_derivs.Ass)
-apply  (erule MGFnD [THEN ax_NormalD])
-apply (tactic "forw_hyp_eval_Force_tac 1")
+(* To term *)
+lemma term_cases: "
+  \<lbrakk>\<And> v. P \<langle>v\<rangle>\<^sub>v; \<And> e. P \<langle>e\<rangle>\<^sub>e;\<And> c. P \<langle>c\<rangle>\<^sub>s;\<And> l. P \<langle>l\<rangle>\<^sub>l\<rbrakk>
+  \<Longrightarrow> P t"
+  apply (cases t)
+  apply (case_tac a)
+  apply auto
+  done
 
-apply (rule ax_derivs.Cond)
-apply  (erule MGFnD [THEN ax_NormalD])
-apply (rule allI)
-apply (rule ax_Normal_cases)
-prefer 2
-apply  (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
-apply  (tactic "eval_Force_tac 1")
-apply (case_tac "b")
-apply  (simp, tactic "forw_hyp_eval_Force_tac 1")
-apply (simp, tactic "forw_hyp_eval_Force_tac 1")
-
-apply (rule_tac Q = " (\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>Init pid_field_type\<rightarrow> s') \<and>. G\<turnstile>init\<le>n" in ax_derivs.Body)
- apply (erule MGFn_InitD [THEN conseq2])
- apply (tactic "eval_Force_tac 1")
-apply (tactic "forw_hyp_tac 1")
-apply (tactic {* clarsimp_tac (eval_css delsimps2 [split_paired_all]) 1 *})
-apply (erule (1) eval.Body)
-
-apply (rule ax_derivs.InstInitE)
-
-apply (rule ax_derivs.Callee)
-
-apply (rule ax_derivs.Skip [THEN conseq1], tactic "eval_Force_tac 1")
-
-apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr],tactic"eval_Force_tac 1")
-
-apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
-apply (tactic "clarsimp_tac eval_css 1")
-
-apply (rule ax_derivs.Comp)
-apply  (erule MGFnD [THEN ax_NormalD])
-apply (tactic "forw_hyp_eval_Force_tac 1")
-
-apply (rule ax_derivs.If)
-apply  (erule MGFnD [THEN ax_NormalD])
-apply (rule allI)
-apply (rule ax_Normal_cases)
-prefer 2
-apply  (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
-apply  (tactic "eval_Force_tac 1")
-apply (case_tac "b")
-apply  (simp, tactic "forw_hyp_eval_Force_tac 1")
-apply (simp, tactic "forw_hyp_eval_Force_tac 1")
-
-apply (rule ax_derivs.Do [THEN conseq1])
-apply (tactic {* force_tac (eval_css addsimps2 [thm "abupd_def2"]) 1 *})
-
-apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
-apply (tactic "clarsimp_tac eval_css 1")
-
-apply (rule_tac Q = " (\<lambda>Y' s' s. normal s \<and> (\<exists>s''. G\<turnstile>s \<midarrow>In1r stmt1\<succ>\<rightarrow> (Y',s'') \<and> G\<turnstile>s'' \<midarrow>sxalloc\<rightarrow> s')) \<and>. G\<turnstile>init\<le>n" in ax_derivs.Try)
-apply   (tactic "eval_Force_tac 3")
-apply  (tactic "forw_hyp_eval_Force_tac 2")
-apply (erule MGFnD [THEN ax_NormalD, THEN conseq2])
-apply (tactic "clarsimp_tac eval_css 1")
-apply (force elim: sxalloc_gext [THEN card_nyinitcls_gext])
-
-apply (rule ax_derivs.FinA)
-
-apply (rule ax_derivs.Nil [THEN conseq1], tactic "eval_Force_tac 1")
-
-apply (rule ax_derivs.Cons)
-apply  (erule MGFnD [THEN ax_NormalD])
-apply (tactic "forw_hyp_eval_Force_tac 1")
-done
+lemma MGFn_lemma:
+  assumes mgf_methds: 
+           "\<And> n. \<forall> C sig. G,(A::state triple set)\<turnstile>{=:n} \<langle>Methd C sig\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+  and wf: "wf_prog G"
+  shows "\<And> t. G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>}"
+proof (induct rule: full_nat_induct)
+  fix n t
+  assume hyp: "\<forall> m. Suc m \<le> n \<longrightarrow> (\<forall> t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>})"
+  show "G,A\<turnstile>{=:n} t\<succ> {G\<rightarrow>}"
+  proof -
+  { 
+    fix v e c es
+    have "G,A\<turnstile>{=:n} \<langle>v\<rangle>\<^sub>v\<succ> {G\<rightarrow>}" and 
+      "G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}" and
+      "G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}" and  
+      "G,A\<turnstile>{=:n} \<langle>es\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
+    proof (induct rule: var_expr_stmt.induct)
+      case (LVar v)
+      show "G,A\<turnstile>{=:n} \<langle>LVar v\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.LVar [THEN conseq1])
+	apply (clarsimp)
+	apply (rule eval.LVar)
+	done
+    next
+      case (FVar accC statDeclC stat e fn)
+      have "G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}".
+      from MGFn_Init [OF hyp] this wf 
+      show ?case
+	by (rule MGFn_FVar)
+    next
+      case (AVar e1 e2)
+      have mgf_e1: "G,A\<turnstile>{=:n} \<langle>e1\<rangle>\<^sub>e\<succ> {G\<rightarrow>}".
+      have mgf_e2: "G,A\<turnstile>{=:n} \<langle>e2\<rangle>\<^sub>e\<succ> {G\<rightarrow>}".
+      show "G,A\<turnstile>{=:n} \<langle>e1.[e2]\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.AVar)
+	apply  (rule MGFnD [OF mgf_e1, THEN ax_NormalD])
+	apply (rule allI)
+	apply (rule MGFnD' [OF mgf_e2, THEN conseq12])
+	apply (fastsimp intro: eval.AVar)
+	done
+    next
+      case (InsInitV c v)
+      show ?case
+	by (rule MGFn_NormalI) (rule ax_derivs.InsInitV)
+    next
+      case (NewC C)
+      show ?case
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.NewC)
+	apply (rule MGFn_InitD [OF hyp, THEN conseq2])
+	apply (fastsimp intro: eval.NewC)
+	done
+    next
+      case (NewA T e)
+      thus ?case
+	apply -
+	apply (rule MGFn_NormalI) 
+	apply (rule ax_derivs.NewA 
+               [where ?Q = "(\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>In1r (init_comp_ty T) 
+                              \<succ>\<rightarrow> (Y',s')) \<and>. G\<turnstile>init\<le>n"])
+	apply  (simp add: init_comp_ty_def split add: split_if)
+	apply  (rule conjI, clarsimp)
+	apply   (rule MGFn_InitD [OF hyp, THEN conseq2])
+	apply   (clarsimp intro: eval.Init)
+	apply  clarsimp
+	apply  (rule ax_derivs.Skip [THEN conseq1])
+	apply  (clarsimp intro: eval.Skip)
+	apply (erule MGFnD' [THEN conseq12])
+	apply (fastsimp intro: eval.NewA)
+	done
+    next
+      case (Cast C e)
+      thus ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast])
+	apply (fastsimp intro: eval.Cast)
+	done
+    next
+      case (Inst e C)
+      thus ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst])
+	apply (fastsimp intro: eval.Inst)
+	done
+    next
+      case (Lit v)
+      show ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Lit [THEN conseq1])
+	apply (fastsimp intro: eval.Lit)
+	done
+    next
+      case (UnOp unop e)
+      thus ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.UnOp)
+	apply (erule MGFnD' [THEN conseq12])
+	apply (fastsimp intro: eval.UnOp)
+	done
+    next
+      case (BinOp binop e1 e2)
+      thus ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.BinOp)
+	apply  (erule MGFnD [THEN ax_NormalD])
+	apply (rule allI)
+	apply (case_tac "need_second_arg binop__ v1")
+	apply  simp
+	apply  (erule MGFnD' [THEN conseq12])
+	apply  (fastsimp intro: eval.BinOp)
+	apply simp
+	apply (rule ax_Normal_cases)
+	apply  (rule ax_derivs.Skip [THEN conseq1])
+	apply  clarsimp
+	apply  (rule eval_BinOp_arg2_indepI)
+	apply   simp
+	apply  simp
+	apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
+	apply (fastsimp intro: eval.BinOp)
+	done
+    next
+      case Super
+      show ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Super [THEN conseq1])
+	apply (fastsimp intro: eval.Super)
+	done
+    next
+      case (Acc v)
+      thus ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc])
+	apply (fastsimp intro: eval.Acc simp add: split_paired_all)
+	done
+    next
+      case (Ass v e)
+      thus "G,A\<turnstile>{=:n} \<langle>v:=e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Ass)
+	apply  (erule MGFnD [THEN ax_NormalD])
+	apply (rule allI)
+	apply (erule MGFnD'[THEN conseq12])
+	apply (fastsimp intro: eval.Ass simp add: split_paired_all)
+	done
+    next
+      case (Cond e1 e2 e3)
+      thus "G,A\<turnstile>{=:n} \<langle>e1 ? e2 : e3\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Cond)
+	apply  (erule MGFnD [THEN ax_NormalD])
+	apply (rule allI)
+	apply (rule ax_Normal_cases)
+	prefer 2
+	apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
+	apply  (fastsimp intro: eval.Cond)
+	apply (case_tac "b")
+	apply  simp
+	apply  (erule MGFnD'[THEN conseq12])
+	apply  (fastsimp intro: eval.Cond)
+	apply simp
+	apply (erule MGFnD'[THEN conseq12])
+	apply (fastsimp intro: eval.Cond)
+	done
+    next
+      case (Call accC statT mode e mn pTs' ps)
+      have mgf_e:  "G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}".
+      have mgf_ps: "G,A\<turnstile>{=:n} \<langle>ps\<rangle>\<^sub>l\<succ> {G\<rightarrow>}".
+      from mgf_methds mgf_e mgf_ps wf
+      show "G,A\<turnstile>{=:n} \<langle>{accC,statT,mode}e\<cdot>mn({pTs'}ps)\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+	by (rule MGFn_Call)
+    next
+      case (Methd D mn)
+      from mgf_methds
+      show "G,A\<turnstile>{=:n} \<langle>Methd D mn\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+	by simp
+    next
+      case (Body D c)
+      have mgf_c: "G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}" .
+      from wf MGFn_Init [OF hyp] mgf_c
+      show "G,A\<turnstile>{=:n} \<langle>Body D c\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+	by (rule MGFn_Body)
+    next
+      case (InsInitE c e)
+      show ?case
+	by (rule MGFn_NormalI) (rule ax_derivs.InsInitE)
+    next
+      case (Callee l e)
+      show ?case
+	by (rule MGFn_NormalI) (rule ax_derivs.Callee)
+    next
+      case Skip
+      show ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Skip [THEN conseq1])
+	apply (fastsimp intro: eval.Skip)
+	done
+    next
+      case (Expr e)
+      thus ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr])
+	apply (fastsimp intro: eval.Expr)
+	done
+    next
+      case (Lab l c)
+      thus "G,A\<turnstile>{=:n} \<langle>l\<bullet> c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
+	apply (fastsimp intro: eval.Lab)
+	done
+    next
+      case (Comp c1 c2)
+      thus "G,A\<turnstile>{=:n} \<langle>c1;; c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Comp)
+	apply  (erule MGFnD [THEN ax_NormalD])
+	apply (erule MGFnD' [THEN conseq12])
+	apply (fastsimp intro: eval.Comp) 
+	done
+    next
+      case (If_ e c1 c2)
+      thus "G,A\<turnstile>{=:n} \<langle>If(e) c1 Else c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.If)
+	apply  (erule MGFnD [THEN ax_NormalD])
+	apply (rule allI)
+	apply (rule ax_Normal_cases)
+	prefer 2
+	apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
+	apply  (fastsimp intro: eval.If)
+	apply (case_tac "b")
+	apply  simp
+	apply  (erule MGFnD' [THEN conseq12])
+	apply  (fastsimp intro: eval.If)
+	apply simp
+	apply (erule MGFnD' [THEN conseq12])
+	apply (fastsimp intro: eval.If)
+	done
+    next
+      case (Loop l e c)
+      have mgf_e: "G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}".
+      have mgf_c: "G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}".
+      from mgf_e mgf_c wf
+      show "G,A\<turnstile>{=:n} \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+	by (rule MGFn_Loop)
+    next
+      case (Jmp j)
+      thus ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Jmp [THEN conseq1])
+	apply (auto intro: eval.Jmp simp add: abupd_def2)
+	done
+    next
+      case (Throw e)
+      thus ?case
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
+	apply (fastsimp intro: eval.Throw)
+	done
+    next
+      case (TryC c1 C vn c2)
+      thus "G,A\<turnstile>{=:n} \<langle>Try c1 Catch(C vn) c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Try [where 
+          ?Q = " (\<lambda>Y' s' s. normal s \<and> (\<exists>s''. G\<turnstile>s \<midarrow>\<langle>c1\<rangle>\<^sub>s\<succ>\<rightarrow> (Y',s'') \<and> 
+                            G\<turnstile>s'' \<midarrow>sxalloc\<rightarrow> s')) \<and>. G\<turnstile>init\<le>n"])
+	apply   (erule MGFnD [THEN ax_NormalD, THEN conseq2])
+	apply   (fastsimp elim: sxalloc_gext [THEN card_nyinitcls_gext])
+	apply  (erule MGFnD'[THEN conseq12])
+	apply  (fastsimp intro: eval.Try)
+	apply (fastsimp intro: eval.Try)
+	done
+    next
+      case (Fin c1 c2)
+      have mgf_c1: "G,A\<turnstile>{=:n} \<langle>c1\<rangle>\<^sub>s\<succ> {G\<rightarrow>}".
+      have mgf_c2: "G,A\<turnstile>{=:n} \<langle>c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}".
+      from wf mgf_c1 mgf_c2
+      show "G,A\<turnstile>{=:n} \<langle>c1 Finally c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+	by (rule MGFn_Fin)
+    next
+      case (FinA abr c)
+      show ?case
+	by (rule MGFn_NormalI) (rule ax_derivs.FinA)
+    next
+      case (Init C)
+      from hyp
+      show "G,A\<turnstile>{=:n} \<langle>Init C\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+	by (rule MGFn_Init)
+    next
+      case Nil_expr
+      show "G,A\<turnstile>{=:n} \<langle>[]\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Nil [THEN conseq1])
+	apply (fastsimp intro: eval.Nil)
+	done
+    next
+      case (Cons_expr e es)
+      thus "G,A\<turnstile>{=:n} \<langle>e# es\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
+	apply -
+	apply (rule MGFn_NormalI)
+	apply (rule ax_derivs.Cons)
+	apply  (erule MGFnD [THEN ax_NormalD])
+	apply (rule allI)
+	apply (erule MGFnD'[THEN conseq12])
+	apply (fastsimp intro: eval.Cons)
+	done
+    qed
+  }
+  thus ?thesis
+    by (cases rule: term_cases) auto
+  qed
+qed
 
 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>
@@ -643,70 +1401,55 @@
 apply assumption (* wf_prog G *)
 done
 
-declare splitI2 [intro!]
-ML_setup {*
-Addsimprocs [ eval_expr_proc, eval_var_proc, eval_exprs_proc, eval_stmt_proc]
-*}
-
-
 section "nested version"
 
-lemma nesting_lemma' [rule_format (no_asm)]: "[| !!A ts. ts <= A ==> P A ts; 
-  !!A pn. !b:bdy pn. P (insert (mgf_call pn) A) {mgf b} ==> P A {mgf_call pn}; 
-  !!A t. !pn:U. P A {mgf_call pn} ==> P A {mgf t};  
-          finite U; uA = mgf_call`U |] ==>  
-  !A. A <= uA --> n <= card uA --> card A = card uA - n --> (!t. P A {mgf t})"
-proof -
-  assume ax_derivs_asm:    "!!A ts. ts <= A ==> P A ts"
-  assume MGF_nested_Methd: "!!A pn. !b:bdy pn. P (insert (mgf_call pn) A) 
-                                                  {mgf b} ==> P A {mgf_call pn}"
-  assume MGF_asm:          "!!A t. !pn:U. P A {mgf_call pn} ==> P A {mgf t}"
-  assume "finite U" "uA = mgf_call`U"
-  then show ?thesis
-    apply -
-    apply (induct_tac "n")
-    apply  (tactic "ALLGOALS Clarsimp_tac")
-    apply  (tactic "dtac (permute_prems 0 1 card_seteq) 1")
-    apply    simp
-    apply   (erule finite_imageI)
-    apply  (simp add: MGF_asm ax_derivs_asm)
-    apply (rule MGF_asm)
-    apply (rule ballI)
-    apply (case_tac "mgf_call pn : A")
-    apply  (fast intro: ax_derivs_asm)
-    apply (rule MGF_nested_Methd)
-    apply (rule ballI)
-    apply (drule spec, erule impE, erule_tac [2] impE, erule_tac [3] spec)
-    apply   fast
-    apply (drule finite_subset)
-    apply (erule finite_imageI)
-    apply auto
-    apply arith
-  done
-qed
+lemma nesting_lemma' [rule_format (no_asm)]: 
+  assumes ax_derivs_asm: "\<And>A ts. ts \<subseteq> A \<Longrightarrow> P A ts" 
+  and MGF_nested_Methd: "\<And>A pn. \<forall>b\<in>bdy pn. P (insert (mgf_call pn) A) {mgf b}
+                                  \<Longrightarrow> P A {mgf_call pn}"
+  and MGF_asm: "\<And>A t. \<forall>pn\<in>U. P A {mgf_call pn} \<Longrightarrow> P A {mgf t}"
+  and finU: "finite U"
+  and uA: "uA = mgf_call`U"
+  shows "\<forall>A. A \<subseteq> uA \<longrightarrow> n \<le> card uA \<longrightarrow> card A = card uA - n 
+             \<longrightarrow> (\<forall>t. P A {mgf t})"
+using finU uA
+apply -
+apply (induct_tac "n")
+apply  (tactic "ALLGOALS Clarsimp_tac")
+apply  (tactic "dtac (permute_prems 0 1 card_seteq) 1")
+apply    simp
+apply   (erule finite_imageI)
+apply  (simp add: MGF_asm ax_derivs_asm)
+apply (rule MGF_asm)
+apply (rule ballI)
+apply (case_tac "mgf_call pn : A")
+apply  (fast intro: ax_derivs_asm)
+apply (rule MGF_nested_Methd)
+apply (rule ballI)
+apply (drule spec, erule impE, erule_tac [2] impE, erule_tac [3] spec)
+apply   fast
+apply (drule finite_subset)
+apply (erule finite_imageI)
+apply auto
+apply arith
+done
 
-lemma nesting_lemma [rule_format (no_asm)]: "[| !!A ts. ts <= A ==> P A ts; 
-  !!A pn. !b:bdy pn. P (insert (mgf (f pn)) A) {mgf b} ==> P A {mgf (f pn)}; 
-          !!A t. !pn:U. P A {mgf (f pn)} ==> P A {mgf t}; 
-          finite U |] ==> P {} {mgf t}"
-proof -
-  assume 2: "!!A pn. !b:bdy pn. P (insert (mgf (f pn)) A) {mgf b} ==> P A {mgf (f pn)}"
-  assume 3: "!!A t. !pn:U. P A {mgf (f pn)} ==> P A {mgf t}"
-  assume "!!A ts. ts <= A ==> P A ts" "finite U"
-  then show ?thesis
-    apply -
-    apply (rule_tac mgf = "mgf" in nesting_lemma')
-    apply (erule_tac [2] 2)
-    apply (rule_tac [2] 3)
-    apply (rule_tac [6] le_refl)
-    apply auto
-  done
-qed
+
+lemma nesting_lemma [rule_format (no_asm)]:
+  assumes ax_derivs_asm: "\<And>A ts. ts \<subseteq> A \<Longrightarrow> P A ts"
+  and MGF_nested_Methd: "\<And>A pn. \<forall>b\<in>bdy pn. P (insert (mgf (f pn)) A) {mgf b}
+                                  \<Longrightarrow> P A {mgf (f pn)}"
+  and MGF_asm: "\<And>A t. \<forall>pn\<in>U. P A {mgf (f pn)} \<Longrightarrow> P A {mgf t}"
+  and finU: "finite U"
+shows "P {} {mgf t}"
+using ax_derivs_asm MGF_nested_Methd MGF_asm finU
+by (rule nesting_lemma') (auto intro!: le_refl)
+
 
 lemma MGF_nested_Methd: "\<lbrakk>  
-  G,insert ({Normal \<doteq>} In1l (Methd  C sig) \<succ>{G\<rightarrow>}) A\<turnstile>  
-            {Normal \<doteq>} In1l (body G C sig) \<succ>{G\<rightarrow>}  
- \<rbrakk> \<Longrightarrow>  G,A\<turnstile>{Normal \<doteq>} In1l (Methd  C sig) \<succ>{G\<rightarrow>}"
+ G,insert ({Normal \<doteq>} \<langle>Methd  C sig\<rangle>\<^sub>e \<succ>{G\<rightarrow>}) A
+    \<turnstile>{Normal \<doteq>} \<langle>body G C sig\<rangle>\<^sub>e \<succ>{G\<rightarrow>}  
+ \<rbrakk> \<Longrightarrow>  G,A\<turnstile>{Normal \<doteq>}  \<langle>Methd  C sig\<rangle>\<^sub>e \<succ>{G\<rightarrow>}"
 apply (unfold MGF_def)
 apply (rule ax_MethdN)
 apply (erule conseq2)
@@ -717,8 +1460,8 @@
 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 
-                f = "\<lambda> (C,sig) . In1l (Methd C sig) " in nesting_lemma)
+                bdy = "\<lambda> (C,sig) .{\<langle>body G C sig\<rangle>\<^sub>e }" and 
+                f = "\<lambda> (C,sig) . \<langle>Methd C sig\<rangle>\<^sub>e " in nesting_lemma)
 apply    (erule ax_derivs.asm)
 apply   (clarsimp simp add: split_tupled_all)
 apply   (erule MGF_nested_Methd)
@@ -731,9 +1474,9 @@
 section "simultaneous version"
 
 lemma MGF_simult_Methd_lemma: "finite ms \<Longrightarrow>  
-  G,A\<union> (\<lambda>(C,sig). {Normal \<doteq>} In1l (Methd  C sig)\<succ> {G\<rightarrow>}) ` ms  
-     |\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} In1l (body G C sig)\<succ> {G\<rightarrow>}) ` ms \<Longrightarrow>  
-  G,A|\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} In1l (Methd  C sig)\<succ> {G\<rightarrow>}) ` ms"
+  G,A \<union> (\<lambda>(C,sig). {Normal \<doteq>} \<langle>Methd  C sig\<rangle>\<^sub>e\<succ> {G\<rightarrow>}) ` ms  
+      |\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} \<langle>body G C sig\<rangle>\<^sub>e\<succ> {G\<rightarrow>}) ` ms \<Longrightarrow>  
+  G,A|\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} \<langle>Methd  C sig\<rangle>\<^sub>e\<succ> {G\<rightarrow>}) ` ms"
 apply (unfold MGF_def)
 apply (rule ax_derivs.Methd [unfolded mtriples_def])
 apply (erule ax_finite_pointwise)
@@ -748,7 +1491,7 @@
 done
 
 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>}) 
+   G,({}::state triple set)|\<turnstile>(\<lambda>(C,sig). {Normal \<doteq>} \<langle>Methd C sig\<rangle>\<^sub>e\<succ> {G\<rightarrow>}) 
    ` Collect (split (is_methd G)) "
 apply (frule finite_is_methd [OF wf_ws_prog])
 apply (rule MGF_simult_Methd_lemma)
@@ -772,20 +1515,48 @@
 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 (blast dest: eval_to_evaln)
-done
-
-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 (1) MGF_deriv)
-done
-
+lemma MGF_complete:
+  assumes valid: "G,{}\<Turnstile>{P} t\<succ> {Q}"
+  and     mgf: "G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
+  and      wf: "wf_prog G"
+  shows "G,({}::state triple set)\<turnstile>{P::state assn} t\<succ> {Q}"
+proof (rule ax_no_hazard)
+  from mgf
+  have "G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y, s')}"  
+    by  (unfold MGF_def) 
+  thus "G,({}::state triple set)\<turnstile>{P \<and>. type_ok G t} t\<succ> {Q}"
+  proof (rule conseq12,clarsimp)
+    fix Y s Z Y' s'
+    assume P: "P Y s Z"
+    assume type_ok: "type_ok G t s"
+    assume eval_t: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y', s')"
+    show "Q Y' s' Z"
+    proof -
+      from eval_t type_ok wf 
+      obtain n where evaln: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s')"
+	by (rule eval_to_evaln [elim_format]) rules
+      from valid have 
+	valid_expanded:
+	"\<forall>n Y s Z. P Y s Z \<longrightarrow> type_ok G t s 
+                   \<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s') \<longrightarrow> Q Y' s' Z)"
+	by (simp add: ax_valids_def triple_valid_def)
+      from P type_ok evaln
+      show "Q Y' s' Z"
+	by (rule valid_expanded [rule_format])
+    qed
+  qed 
+qed
+   
+theorem ax_complete: 
+  assumes wf: "wf_prog G" 
+  and valid: "G,{}\<Turnstile>{P::state assn} t\<succ> {Q}"
+  shows "G,({}::state triple set)\<turnstile>{P} t\<succ> {Q}"
+proof -
+  from wf have "G,({}::state triple set)\<turnstile>{\<doteq>} t\<succ> {G\<rightarrow>}"
+    by (rule MGF_deriv)
+  from valid this wf
+  show ?thesis
+    by (rule MGF_complete)
+qed
+ 
 end
--- a/src/HOL/Bali/AxExample.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/AxExample.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -52,10 +52,13 @@
 
 theorem ax_test: "tprg,({}::'a triple set)\<turnstile> 
   {Normal (\<lambda>Y s Z::'a. heap_free four s \<and> \<not>initd Base s \<and> \<not> initd Ext s)} 
-  .test [Class Base]. {\<lambda>Y s Z. abrupt s = Some (Xcpt (Std IndOutBound))}"
+  .test [Class Base]. 
+  {\<lambda>Y s Z. abrupt s = Some (Xcpt (Std IndOutBound))}"
 apply (unfold test_def arr_viewed_from_def)
 apply (tactic "ax_tac 1" (*;;*))
-defer
+defer (* We begin with the last assertion, to synthesise the intermediate
+         assertions, like in the fashion of the weakest
+         precondition. *)
 apply  (tactic "ax_tac 1" (* Try *))
 defer
 apply    (tactic {* inst1_tac "Q1" 
@@ -108,11 +111,18 @@
 (* apply       (rule_tac [2] ax_derivs.Abrupt) *)
 defer
 apply      (simp (no_asm))
-apply      (tactic "ax_tac 1")
+apply      (tactic "ax_tac 1") (* Comp *)
+            (* The first statement in the  composition 
+                 ((Ext)z).vee = 1; Return Null 
+                will throw an exception (since z is null). So we can handle
+                Return Null with the Abrupt rule *)
+apply       (rule_tac [2] ax_derivs.Abrupt)
+             
+apply      (rule ax_derivs.Expr) (* Expr *)
 apply      (tactic "ax_tac 1") (* Ass *)
 prefer 2
 apply       (rule ax_subst_Var_allI)
-apply       (tactic {* inst1_tac "P'27" "\<lambda>a vs l vf. ?PP a vs l vf\<leftarrow>?x \<and>. ?p" *})
+apply       (tactic {* inst1_tac "P'29" "\<lambda>a vs l vf. ?PP a vs l vf\<leftarrow>?x \<and>. ?p" *})
 apply       (rule allI)
 apply       (tactic {* simp_tac (simpset() delloop "split_all_tac" delsimps [thm "peek_and_def2"]) 1 *})
 apply       (rule ax_derivs.Abrupt)
@@ -120,11 +130,10 @@
 apply      (tactic "ax_tac 1" (* FVar *))
 apply       (tactic "ax_tac 2", tactic "ax_tac 2", tactic "ax_tac 2")
 apply      (tactic "ax_tac 1")
-apply     clarsimp
 apply     (tactic {* inst1_tac "R14" "\<lambda>a'. Normal ((\<lambda>Vals:vs (x, s) Z. arr_inv s \<and> inited Ext (globs s) \<and> a' \<noteq> Null \<and> hd vs = Null) \<and>. heap_free two)" *})
-prefer 5
-apply     (rule ax_derivs.Done [THEN conseq1], force)
-apply    force
+apply     fastsimp
+prefer 4
+apply    (rule ax_derivs.Done [THEN conseq1],force)
 apply   (rule ax_subst_Val_allI)
 apply   (tactic {* inst1_tac "P'33" "\<lambda>u a. Normal (?PP a\<leftarrow>?x) u" *})
 apply   (simp (no_asm) del: peek_and_def2)
--- a/src/HOL/Bali/AxSem.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/AxSem.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -39,7 +39,7 @@
 \end{itemize}
 *}
 
-types  res = vals (* result entry *)
+types  res = vals --{* result entry *}
 syntax
   Val  :: "val      \<Rightarrow> res"
   Var  :: "var      \<Rightarrow> res"
@@ -59,7 +59,7 @@
   "\<lambda>Var:v . b"  == "(\<lambda>v. b) \<circ> the_In2"
   "\<lambda>Vals:v. b"  == "(\<lambda>v. b) \<circ> the_In3"
 
-  (* relation on result values, state and auxiliary variables *)
+  --{* relation on result values, state and auxiliary variables *}
 types 'a assn   =        "res \<Rightarrow> state \<Rightarrow> 'a \<Rightarrow> bool"
 translations
       "res"    <= (type) "AxSem.res"
@@ -377,7 +377,10 @@
 
 constdefs
   type_ok  :: "prog \<Rightarrow> term \<Rightarrow> state \<Rightarrow> bool"
- "type_ok G t s \<equiv> \<exists>L T C. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) \<and> s\<Colon>\<preceq>(G,L)"
+ "type_ok G t s \<equiv> 
+    \<exists>L T C A. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<and> 
+                            \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A )
+              \<and> s\<Colon>\<preceq>(G,L)"
 
 datatype    'a triple = triple "('a assn)" "term" "('a assn)" (** should be
 something like triple = \<forall>'a. triple ('a assn) term ('a assn)   **)
@@ -468,8 +471,10 @@
 
 lemma triple_valid_def2: "G\<Turnstile>n:{P} t\<succ> {Q} =  
  (\<forall>Y s Z. P Y s Z 
-  \<longrightarrow> (\<exists>L. (normal s \<longrightarrow> (\<exists>T C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T)) \<and> s\<Colon>\<preceq>(G,L)) \<longrightarrow> 
-  (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s')\<longrightarrow> Q Y' s' Z))"
+  \<longrightarrow> (\<exists>L. (normal s \<longrightarrow> (\<exists> C T A. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<and> 
+                   \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A)) \<and> 
+           s\<Colon>\<preceq>(G,L))
+  \<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s')\<longrightarrow> Q Y' s' Z))"
 apply (unfold triple_valid_def type_ok_def)
 apply (simp (no_asm))
 done
@@ -506,7 +511,7 @@
 
   Abrupt:  "G,A\<turnstile>{P\<leftarrow>(arbitrary3 t) \<and>. Not \<circ> normal} t\<succ> {P}"
 
-  (* variables *)
+  --{* variables *}
   LVar:  " G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Var (lvar vn s))} LVar vn=\<succ> {P}"
 
   FVar: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Q};
@@ -516,7 +521,7 @@
   AVar:  "\<lbrakk>G,A\<turnstile>{Normal P} e1-\<succ> {Q};
           \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} e2-\<succ> {\<lambda>Val:i:. avar G i a ..; R}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} e1.[e2]=\<succ> {R}"
-  (* expressions *)
+  --{* expressions *}
 
   NewC: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Alloc G (CInst C) Q}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}"
@@ -579,7 +584,7 @@
     \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} Body D c-\<succ> {R}"
   
-  (* expression lists *)
+  --{* expression lists *}
 
   Nil:                          "G,A\<turnstile>{Normal (P\<leftarrow>Vals [])} []\<doteq>\<succ> {P}"
 
@@ -587,7 +592,7 @@
           \<forall>v. G,A\<turnstile>{Q\<leftarrow>Val v} es\<doteq>\<succ> {\<lambda>Vals:vs:. R\<leftarrow>Vals (v#vs)}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} e#es\<doteq>\<succ> {R}"
 
-  (* statements *)
+  --{* statements *}
 
   Skip:                         "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit>)} .Skip. {P}"
 
@@ -612,9 +617,8 @@
   Loop: "\<lbrakk>G,A\<turnstile>{P} e-\<succ> {P'}; 
           G,A\<turnstile>{Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}\<rbrakk> \<Longrightarrow>
                             G,A\<turnstile>{P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}"
-(** Beware of polymorphic_Loop below: should be identical terms **)
   
-  Do: "G,A\<turnstile>{Normal (abupd (\<lambda>a. (Some (Jump j))) .; P\<leftarrow>\<diamondsuit>)} .Do j. {P}"
+  Jmp: "G,A\<turnstile>{Normal (abupd (\<lambda>a. (Some (Jump j))) .; P\<leftarrow>\<diamondsuit>)} .Jmp j. {P}"
 
   Throw:"\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} .Throw e. {Q}"
@@ -642,21 +646,13 @@
 @{text InsInitE}, @{text InsInitV}, @{text FinA} only used by the smallstep 
 semantics.
 *}
-  InstInitV: " G,A\<turnstile>{Normal P} InsInitV c v=\<succ> {Q}"
-  InstInitE: " G,A\<turnstile>{Normal P} InsInitE c e-\<succ> {Q}"
+  InsInitV: " G,A\<turnstile>{Normal P} InsInitV c v=\<succ> {Q}"
+  InsInitE: " G,A\<turnstile>{Normal P} InsInitE c e-\<succ> {Q}"
   Callee:    " G,A\<turnstile>{Normal P} Callee l e-\<succ> {Q}"
   FinA:      " G,A\<turnstile>{Normal P} .FinA a c. {Q}"
-axioms (** these terms are the same as above, but with generalized typing **)
-  polymorphic_conseq:
-        "\<forall>Y s Z . P  Y s Z  \<longrightarrow> (\<exists>P' Q'. G,A\<turnstile>{P'} t\<succ> {Q'} \<and> (\<forall>Y' s'. 
-        (\<forall>Y   Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>
-                                Q  Y' s' Z ))
-                                         \<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q }"
-
-  polymorphic_Loop:
-        "\<lbrakk>G,A\<turnstile>{P} e-\<succ> {P'}; 
-          G,A\<turnstile>{Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}\<rbrakk> \<Longrightarrow>
-                            G,A\<turnstile>{P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}"
+(*
+axioms 
+*)
 
 constdefs
  adapt_pre :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
@@ -743,50 +739,67 @@
 
 section "rules derived from conseq"
 
-lemma conseq12: "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q'};  
+text {* In the following rules we often have to give some type annotations like:
+ @{term "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q}"}.
+Given only the term above without annotations, Isabelle would infer a more 
+general type were we could have 
+different types of auxiliary variables in the assumption set (@{term A}) and 
+in the triple itself (@{term P} and @{term Q}). But 
+@{text "ax_derivs.Methd"} enforces the same type in the inductive definition of
+the derivation. So we have to restrict the types to be able to apply the
+rules. 
+*}
+lemma conseq12: "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'};  
  \<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>  
   Q Y' s' Z)\<rbrakk>  
   \<Longrightarrow>  G,A\<turnstile>{P ::'a assn} t\<succ> {Q }"
-apply (rule polymorphic_conseq)
+apply (rule ax_derivs.conseq)
 apply clarsimp
 apply blast
 done
 
-(*unused, but nice variant*)
-lemma conseq12': "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q'}; \<forall>s Y' s'.  
+-- {* Nice variant, since it is so symmetric we might be able to memorise it. *}
+lemma conseq12': "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'}; \<forall>s Y' s'.  
        (\<forall>Y Z. P' Y s Z \<longrightarrow> Q' Y' s' Z) \<longrightarrow>  
        (\<forall>Y Z. P  Y s Z \<longrightarrow> Q  Y' s' Z)\<rbrakk>  
-  \<Longrightarrow>  G,A\<turnstile>{P } t\<succ> {Q }"
+  \<Longrightarrow>  G,A\<turnstile>{P::'a assn } t\<succ> {Q }"
 apply (erule conseq12)
 apply fast
 done
 
-lemma conseq12_from_conseq12': "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q'};  
+lemma conseq12_from_conseq12': "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'};  
  \<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>  
   Q Y' s' Z)\<rbrakk>  
-  \<Longrightarrow>  G,A\<turnstile>{P } t\<succ> {Q }"
+  \<Longrightarrow>  G,A\<turnstile>{P::'a assn} t\<succ> {Q }"
 apply (erule conseq12')
 apply blast
 done
 
-lemma conseq1: "\<lbrakk>G,A\<turnstile>{P'} t\<succ> {Q}; P \<Rightarrow> P'\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q}"
+lemma conseq1: "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q}; P \<Rightarrow> P'\<rbrakk> 
+ \<Longrightarrow> G,A\<turnstile>{P::'a assn} t\<succ> {Q}"
 apply (erule conseq12)
 apply blast
 done
 
-lemma conseq2: "\<lbrakk>G,A\<turnstile>{P} t\<succ> {Q'}; Q' \<Rightarrow> Q\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
+lemma conseq2: "\<lbrakk>G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q'}; Q' \<Rightarrow> Q\<rbrakk> 
+\<Longrightarrow> G,A\<turnstile>{P::'a assn} t\<succ> {Q}"
 apply (erule conseq12)
 apply blast
 done
 
-lemma ax_escape: "\<lbrakk>\<forall>Y s Z. P Y s Z \<longrightarrow> G,A\<turnstile>{\<lambda>Y' s' Z'. (Y',s') = (Y,s)} t\<succ> {\<lambda>Y s Z'. Q Y s Z}\<rbrakk> \<Longrightarrow>  
-  G,A\<turnstile>{P} t\<succ> {Q}"
-apply (rule polymorphic_conseq)
+lemma ax_escape: 
+ "\<lbrakk>\<forall>Y s Z. P Y s Z 
+   \<longrightarrow> G,(A::'a triple set)\<turnstile>{\<lambda>Y' s' (Z'::'a). (Y',s') = (Y,s)} 
+                             t\<succ> 
+                            {\<lambda>Y s Z'. Q Y s Z}
+\<rbrakk> \<Longrightarrow>  G,A\<turnstile>{P::'a assn} t\<succ> {Q::'a assn}"
+apply (rule ax_derivs.conseq)
 apply force
 done
 
 (* unused *)
-lemma ax_constant: "\<lbrakk> C \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}\<rbrakk> \<Longrightarrow> G,A\<turnstile>{\<lambda>Y s Z. C \<and> P Y s Z} t\<succ> {Q}"
+lemma ax_constant: "\<lbrakk> C \<Longrightarrow> G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q}\<rbrakk> 
+\<Longrightarrow> G,A\<turnstile>{\<lambda>Y s Z. C \<and> P Y s Z} t\<succ> {Q}"
 apply (rule ax_escape (* unused *))
 apply clarify
 apply (rule conseq12)
@@ -799,7 +812,8 @@
 *)
 
 
-lemma ax_impossible [intro]: "G,A\<turnstile>{\<lambda>Y s Z. False} t\<succ> {Q}"
+lemma ax_impossible [intro]: 
+  "G,(A::'a triple set)\<turnstile>{\<lambda>Y s Z. False} t\<succ> {Q::'a assn}"
 apply (rule ax_escape)
 apply clarify
 done
@@ -808,34 +822,40 @@
 lemma ax_nochange_lemma: "\<lbrakk>P Y s; All (op = w)\<rbrakk> \<Longrightarrow> P w s"
 apply auto
 done
-lemma ax_nochange:"G,A\<turnstile>{\<lambda>Y s Z. (Y,s)=Z} t\<succ> {\<lambda>Y s Z. (Y,s)=Z} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {P}"
+
+lemma ax_nochange:
+ "G,(A::(res \<times> state) triple set)\<turnstile>{\<lambda>Y s Z. (Y,s)=Z} t\<succ> {\<lambda>Y s Z. (Y,s)=Z} 
+  \<Longrightarrow> G,A\<turnstile>{P::(res \<times> state) assn} t\<succ> {P}"
 apply (erule conseq12)
 apply auto
 apply (erule (1) ax_nochange_lemma)
 done
 
 (* unused *)
-lemma ax_trivial: "G,A\<turnstile>{P}  t\<succ> {\<lambda>Y s Z. True}"
-apply (rule polymorphic_conseq(* unused *))
+lemma ax_trivial: "G,(A::'a triple set)\<turnstile>{P::'a assn}  t\<succ> {\<lambda>Y s Z. True}"
+apply (rule ax_derivs.conseq(* unused *))
 apply auto
 done
 
 (* unused *)
-lemma ax_disj: "\<lbrakk>G,A\<turnstile>{P1} t\<succ> {Q1}; G,A\<turnstile>{P2} t\<succ> {Q2}\<rbrakk> \<Longrightarrow>  
-  G,A\<turnstile>{\<lambda>Y s Z. P1 Y s Z \<or> P2 Y s Z} t\<succ> {\<lambda>Y s Z. Q1 Y s Z \<or> Q2 Y s Z}"
+lemma ax_disj: 
+ "\<lbrakk>G,(A::'a triple set)\<turnstile>{P1::'a assn} t\<succ> {Q1}; G,A\<turnstile>{P2::'a assn} t\<succ> {Q2}\<rbrakk> 
+  \<Longrightarrow>  G,A\<turnstile>{\<lambda>Y s Z. P1 Y s Z \<or> P2 Y s Z} t\<succ> {\<lambda>Y s Z. Q1 Y s Z \<or> Q2 Y s Z}"
 apply (rule ax_escape (* unused *))
 apply safe
 apply  (erule conseq12, fast)+
 done
 
 (* unused *)
-lemma ax_supd_shuffle: "(\<exists>Q. G,A\<turnstile>{P} .c1. {Q} \<and> G,A\<turnstile>{Q ;. f} .c2. {R}) =  
+lemma ax_supd_shuffle: 
+ "(\<exists>Q. G,(A::'a triple set)\<turnstile>{P::'a assn} .c1. {Q} \<and> G,A\<turnstile>{Q ;. f} .c2. {R}) =  
        (\<exists>Q'. G,A\<turnstile>{P} .c1. {f .; Q'} \<and> G,A\<turnstile>{Q'} .c2. {R})"
 apply (best elim!: conseq1 conseq2)
 done
 
-lemma ax_cases: "\<lbrakk>G,A\<turnstile>{P \<and>.       C} t\<succ> {Q};  
-                       G,A\<turnstile>{P \<and>. Not \<circ> C} t\<succ> {Q}\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
+lemma ax_cases: "
+ \<lbrakk>G,(A::'a triple set)\<turnstile>{P \<and>.       C} t\<succ> {Q::'a assn};  
+                   G,A\<turnstile>{P \<and>. Not \<circ> C} t\<succ> {Q}\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
 apply (unfold peek_and_def)
 apply (rule ax_escape)
 apply clarify
@@ -849,13 +869,15 @@
 apply  force+
 *)
 
-lemma ax_adapt: "G,A\<turnstile>{P} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}"
+lemma ax_adapt: "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q} 
+  \<Longrightarrow> G,A\<turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}"
 apply (unfold adapt_pre_def)
 apply (erule conseq12)
 apply fast
 done
 
-lemma adapt_pre_adapts: "G,A\<Turnstile>{P} t\<succ> {Q} \<longrightarrow> G,A\<Turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}"
+lemma adapt_pre_adapts: "G,(A::'a triple set)\<Turnstile>{P::'a assn} t\<succ> {Q} 
+\<longrightarrow> G,A\<Turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}"
 apply (unfold adapt_pre_def)
 apply (simp add: ax_valids_def triple_valid_def2)
 apply fast
@@ -872,72 +894,44 @@
 apply (simp add: ax_valids_def triple_valid_def2)
 oops
 
-(*
-Goal "\<forall>(A::'a triple set) t. G,A\<Turnstile>{P} t\<succ> {Q} \<longrightarrow> G,A\<Turnstile>{P'} t\<succ> {Q'} \<Longrightarrow>  
-  wf_prog G \<Longrightarrow> G,(A::'a triple set)\<turnstile>{P} t\<succ> {Q::'a assn} \<Longrightarrow> G,A\<turnstile>{P'} t\<succ> {Q'::'a assn}"
-b y fatac ax_sound 1 1;
-b y asm_full_simp_tac (simpset() addsimps [ax_valids_def,triple_valid_def2]) 1;
-b y rtac ax_no_hazard 1; 
-b y etac conseq12 1;
-b y Clarify_tac 1;
-b y case_tac "\<forall>Z. \<not>P Y s Z" 1;
-b y smp_tac 2 1;
-b y etac thin_rl 1;
-b y etac thin_rl 1;
-b y clarsimp_tac (claset(), simpset() addsimps [type_ok_def]) 1;
-b y subgoal_tac "G|\<Turnstile>n:A" 1;
-b y smp_tac 1 1;
-b y smp_tac 3 1;
-b y etac impE 1;
- back();
- b y Fast_tac 1;
-b y 
-b y rotate_tac 2 1;
-b y etac thin_rl 1;
-b y  etac thin_rl 2;
-b y  etac thin_rl 2;
-b y  Clarify_tac 2;
-b y  dtac spec 2;
-b y  EVERY'[dtac spec, mp_tac] 2;
-b y  thin_tac "\<forall>n Y s Z. ?PP n Y s Z" 2;
-b y  thin_tac "P' Y s Z" 2;
-b y  Blast_tac 2;
-b y smp_tac 3 1;
-b y case_tac "\<forall>Z. \<not>P Y s Z" 1;
-b y dres_inst_tac [("x","In1r Skip")] spec 1;
-b y Full_simp_tac 1;
-*)
-
 lemma peek_and_forget1_Normal: 
- "G,A\<turnstile>{Normal P} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{Normal (P \<and>. p)} t\<succ> {Q}"
+ "G,(A::'a triple set)\<turnstile>{Normal P} t\<succ> {Q::'a assn} 
+ \<Longrightarrow> G,A\<turnstile>{Normal (P \<and>. p)} t\<succ> {Q}"
 apply (erule conseq1)
 apply (simp (no_asm))
 done
 
-lemma peek_and_forget1: "G,A\<turnstile>{P} t\<succ> {Q} \<Longrightarrow> G,A\<turnstile>{P \<and>. p} t\<succ> {Q}"
+lemma peek_and_forget1: 
+"G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q} 
+ \<Longrightarrow> G,A\<turnstile>{P \<and>. p} t\<succ> {Q}"
 apply (erule conseq1)
 apply (simp (no_asm))
 done
 
 lemmas ax_NormalD = peek_and_forget1 [of _ _ _ _ _ normal] 
 
-lemma peek_and_forget2: "G,A\<turnstile>{P} t\<succ> {Q \<and>. p} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
+lemma peek_and_forget2: 
+"G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q \<and>. p} 
+\<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}"
 apply (erule conseq2)
 apply (simp (no_asm))
 done
 
-lemma ax_subst_Val_allI: "\<forall>v. G,A\<turnstile>{(P'               v )\<leftarrow>Val v} t\<succ> {Q v} \<Longrightarrow>  
-      \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In1 w))\<leftarrow>Val v} t\<succ> {Q v}"
+lemma ax_subst_Val_allI: 
+"\<forall>v. G,(A::'a triple set)\<turnstile>{(P'               v )\<leftarrow>Val v} t\<succ> {(Q v)::'a assn}
+ \<Longrightarrow>  \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In1 w))\<leftarrow>Val v} t\<succ> {Q v}"
 apply (force elim!: conseq1)
 done
 
-lemma ax_subst_Var_allI: "\<forall>v. G,A\<turnstile>{(P'               v )\<leftarrow>Var v} t\<succ> {Q v} \<Longrightarrow>  
-      \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In2 w))\<leftarrow>Var v} t\<succ> {Q v}"
+lemma ax_subst_Var_allI: 
+"\<forall>v. G,(A::'a triple set)\<turnstile>{(P'               v )\<leftarrow>Var v} t\<succ> {(Q v)::'a assn}
+ \<Longrightarrow>  \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In2 w))\<leftarrow>Var v} t\<succ> {Q v}"
 apply (force elim!: conseq1)
 done
 
-lemma ax_subst_Vals_allI: "(\<forall>v. G,A\<turnstile>{(     P'          v )\<leftarrow>Vals v} t\<succ> {Q v}) \<Longrightarrow>  
-       \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In3 w))\<leftarrow>Vals v} t\<succ> {Q v}"
+lemma ax_subst_Vals_allI: 
+"(\<forall>v. G,(A::'a triple set)\<turnstile>{(     P'          v )\<leftarrow>Vals v} t\<succ> {(Q v)::'a assn})
+ \<Longrightarrow>  \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In3 w))\<leftarrow>Vals v} t\<succ> {Q v}"
 apply (force elim!: conseq1)
 done
 
@@ -1171,36 +1165,43 @@
 
 section "introduction rules for Alloc and SXAlloc"
 
-lemma ax_SXAlloc_Normal: "G,A\<turnstile>{P} .c. {Normal Q} \<Longrightarrow> G,A\<turnstile>{P} .c. {SXAlloc G Q}"
+lemma ax_SXAlloc_Normal: 
+ "G,(A::'a triple set)\<turnstile>{P::'a assn} .c. {Normal Q} 
+ \<Longrightarrow> G,A\<turnstile>{P} .c. {SXAlloc G Q}"
 apply (erule conseq2)
 apply (clarsimp elim!: sxalloc_elim_cases simp add: split_tupled_all)
 done
 
 lemma ax_Alloc: 
-  "G,A\<turnstile>{P} t\<succ> {Normal (\<lambda>Y (x,s) Z. (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
- Q (Val (Addr a)) (Norm(init_obj G (CInst C) (Heap a) s)) Z)) \<and>. 
-    heap_free (Suc (Suc 0))}
+  "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> 
+     {Normal (\<lambda>Y (x,s) Z. (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
+      Q (Val (Addr a)) (Norm(init_obj G (CInst C) (Heap a) s)) Z)) \<and>. 
+      heap_free (Suc (Suc 0))}
    \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Alloc G (CInst C) Q}"
 apply (erule conseq2)
 apply (auto elim!: halloc_elim_cases)
 done
 
 lemma ax_Alloc_Arr: 
- "G,A\<turnstile>{P} t\<succ> {\<lambda>Val:i:. Normal (\<lambda>Y (x,s) Z. \<not>the_Intg i<0 \<and>  
-  (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
-  Q (Val (Addr a)) (Norm (init_obj G (Arr T (the_Intg i)) (Heap a) s)) Z)) \<and>. 
-   heap_free (Suc (Suc 0))} \<Longrightarrow>  
+ "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> 
+   {\<lambda>Val:i:. Normal (\<lambda>Y (x,s) Z. \<not>the_Intg i<0 \<and>  
+    (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
+    Q (Val (Addr a)) (Norm (init_obj G (Arr T (the_Intg i)) (Heap a) s)) Z)) \<and>.
+    heap_free (Suc (Suc 0))} 
+ \<Longrightarrow>  
  G,A\<turnstile>{P} t\<succ> {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T(the_Intg i)) Q}"
 apply (erule conseq2)
 apply (auto elim!: halloc_elim_cases)
 done
 
 lemma ax_SXAlloc_catch_SXcpt: 
- "\<lbrakk>G,A\<turnstile>{P} t\<succ> {(\<lambda>Y (x,s) Z. x=Some (Xcpt (Std xn)) \<and>  
-  (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
-  Q Y (Some (Xcpt (Loc a)),init_obj G (CInst (SXcpt xn)) (Heap a) s) Z))  
-  \<and>. heap_free (Suc (Suc 0))}\<rbrakk> \<Longrightarrow>  
-  G,A\<turnstile>{P} t\<succ> {SXAlloc G (\<lambda>Y s Z. Q Y s Z \<and> G,s\<turnstile>catch SXcpt xn)}"
+ "\<lbrakk>G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> 
+     {(\<lambda>Y (x,s) Z. x=Some (Xcpt (Std xn)) \<and>  
+      (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow>  
+      Q Y (Some (Xcpt (Loc a)),init_obj G (CInst (SXcpt xn)) (Heap a) s) Z))  
+      \<and>. heap_free (Suc (Suc 0))}\<rbrakk> 
+ \<Longrightarrow>  
+ G,A\<turnstile>{P} t\<succ> {SXAlloc G (\<lambda>Y s Z. Q Y s Z \<and> G,s\<turnstile>catch SXcpt xn)}"
 apply (erule conseq2)
 apply (auto elim!: sxalloc_elim_cases halloc_elim_cases)
 done
--- a/src/HOL/Bali/AxSound.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/AxSound.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Bali/AxSound.thy
     ID:         $Id$
-    Author:     David von Oheimb
+    Author:     David von Oheimb and Norbert Schirmer
     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 *)
 header {* Soundness proof for Axiomatic semantics of Java expressions and 
@@ -22,14 +22,24 @@
 
 defs  triple_valid2_def: "G\<Turnstile>n\<Colon>t \<equiv> case t of {P} t\<succ> {Q} \<Rightarrow>
  \<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>L. s\<Colon>\<preceq>(G,L) 
- \<longrightarrow> (\<forall>T C. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) \<longrightarrow>
+ \<longrightarrow> (\<forall>T C A. (normal s \<longrightarrow> (\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<and> 
+                            \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A)) \<longrightarrow>
  (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> Q Y' s' Z \<and> s'\<Colon>\<preceq>(G,L))))"
 
+text {* This definition differs from the ordinary  @{text triple_valid_def} 
+manly in the conclusion: We also ensures conformance of the result state. So
+we don't have to apply the type soundness lemma all the time during
+induction. This definition is only introduced for the soundness
+proof of the axiomatic semantics, in the end we will conclude to 
+the ordinary definition.
+*}
+ 
 defs  ax_valids2_def:    "G,A|\<Turnstile>\<Colon>ts \<equiv>  \<forall>n. (\<forall>t\<in>A. G\<Turnstile>n\<Colon>t) \<longrightarrow> (\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t)"
 
 lemma triple_valid2_def2: "G\<Turnstile>n\<Colon>{P} t\<succ> {Q} =  
  (\<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s')\<longrightarrow>  
-  (\<forall>L. s\<Colon>\<preceq>(G,L) \<longrightarrow> (\<forall>T C. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) \<longrightarrow>  
+  (\<forall>L. s\<Colon>\<preceq>(G,L) \<longrightarrow> (\<forall>T C A. (normal s \<longrightarrow> (\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<and> 
+                            \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A)) \<longrightarrow>
   Q Y' s' Z \<and> s'\<Colon>\<preceq>(G,L)))))"
 apply (unfold triple_valid2_def)
 apply (simp (no_asm) add: split_paired_All)
@@ -48,10 +58,18 @@
 apply (tactic "smp_tac 3 1")
 apply (case_tac "normal s")
 apply  clarsimp
-apply  (blast dest: evaln_eval eval_type_sound [THEN conjunct1])
-apply clarsimp
+apply  (elim conjE impE)
+apply    blast
+
+apply    (tactic "smp_tac 2 1")
+apply    (drule evaln_eval)
+apply    (drule (1) eval_type_sound [THEN conjunct1],simp, assumption+)
+apply    simp
+
+apply    clarsimp
 done
 
+
 lemma ax_valids2_eq: "wf_prog G \<Longrightarrow> G,A|\<Turnstile>\<Colon>ts = G,A|\<Turnstile>ts"
 apply (unfold ax_valids_def ax_valids2_def)
 apply (force simp add: triple_valid2_eq)
@@ -73,9 +91,9 @@
   \<Longrightarrow> G\<Turnstile>Suc n\<Colon>{Normal P} Methd C sig-\<succ> {Q}"
 apply (simp (no_asm_use) add: triple_valid2_def2)
 apply (intro strip, tactic "smp_tac 3 1", clarify)
-apply (erule wt_elim_cases, erule evaln_elim_cases)
+apply (erule wt_elim_cases, erule da_elim_cases, erule evaln_elim_cases)
 apply (unfold body_def Let_def)
-apply clarsimp
+apply (clarsimp simp add: inj_term_simps)
 apply blast
 done
 
@@ -91,380 +109,2555 @@
 section "soundness"
 
 lemma Methd_sound: 
-"\<lbrakk>G,A\<union>  {{P} Methd-\<succ> {Q} | ms}|\<Turnstile>\<Colon>{{P} body G-\<succ> {Q} | ms}\<rbrakk> \<Longrightarrow> 
-  G,A|\<Turnstile>\<Colon>{{P} Methd-\<succ> {Q} | ms}"
-apply (unfold ax_valids2_def mtriples_def)
-apply (rule allI)
-apply (induct_tac "n")
-apply  (clarify, tactic {* pair_tac "x" 1 *}, simp (no_asm))
-apply  (fast intro: Methd_triple_valid2_0)
-apply (clarify, tactic {* pair_tac "xa" 1 *}, simp (no_asm))
-apply (drule triples_valid2_Suc)
-apply (erule (1) notE impE)
-apply (drule_tac x = na in spec)
-apply (rule Methd_triple_valid2_SucI)
-apply (simp (no_asm_use) add: ball_Un)
-apply auto
-done
+  assumes recursive: "G,A\<union>  {{P} Methd-\<succ> {Q} | ms}|\<Turnstile>\<Colon>{{P} body G-\<succ> {Q} | ms}"
+  shows "G,A|\<Turnstile>\<Colon>{{P} Methd-\<succ> {Q} | ms}"
+proof -
+  {
+    fix n
+    assume recursive: "\<And> n. \<forall>t\<in>(A \<union> {{P} Methd-\<succ> {Q} | ms}). G\<Turnstile>n\<Colon>t
+                              \<Longrightarrow>  \<forall>t\<in>{{P} body G-\<succ> {Q} | ms}.  G\<Turnstile>n\<Colon>t"
+    have "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t \<Longrightarrow> \<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}.  G\<Turnstile>n\<Colon>t"
+    proof (induct n)
+      case 0
+      show "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}.  G\<Turnstile>0\<Colon>t"
+      proof -
+	{
+	  fix C sig
+	  assume "(C,sig) \<in> ms" 
+	  have "G\<Turnstile>0\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
+	    by (rule Methd_triple_valid2_0)
+	}
+	thus ?thesis
+	  by (simp add: mtriples_def split_def)
+      qed
+    next
+      case (Suc m)
+      have hyp: "\<forall>t\<in>A. G\<Turnstile>m\<Colon>t \<Longrightarrow> \<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}.  G\<Turnstile>m\<Colon>t".
+      have prem: "\<forall>t\<in>A. G\<Turnstile>Suc m\<Colon>t" .
+      show "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}.  G\<Turnstile>Suc m\<Colon>t"
+      proof -
+	{
+	  fix C sig
+	  assume m: "(C,sig) \<in> ms" 
+	  have "G\<Turnstile>Suc m\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
+	  proof -
+	    from prem have prem_m: "\<forall>t\<in>A. G\<Turnstile>m\<Colon>t"
+	      by (rule triples_valid2_Suc)
+	    hence "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}.  G\<Turnstile>m\<Colon>t"
+	      by (rule hyp)
+	    with prem_m
+	    have "\<forall>t\<in>(A \<union> {{P} Methd-\<succ> {Q} | ms}). G\<Turnstile>m\<Colon>t"
+	      by (simp add: ball_Un)
+	    hence "\<forall>t\<in>{{P} body G-\<succ> {Q} | ms}.  G\<Turnstile>m\<Colon>t"
+	      by (rule recursive)
+	    with m have "G\<Turnstile>m\<Colon>{Normal (P C sig)} body G C sig-\<succ> {Q C sig}"
+	      by (auto simp add: mtriples_def split_def)
+	    thus ?thesis
+	      by (rule Methd_triple_valid2_SucI)
+	  qed
+	}
+	thus ?thesis
+	  by (simp add: mtriples_def split_def)
+      qed
+    qed
+  }
+  with recursive show ?thesis
+    by (unfold ax_valids2_def) blast
+qed
 
 
 lemma valids2_inductI: "\<forall>s t n Y' s'. G\<turnstile>s\<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> t = c \<longrightarrow>    
   Ball A (triple_valid2 G n) \<longrightarrow> (\<forall>Y Z. P Y s Z \<longrightarrow>  
-  (\<forall>L. s\<Colon>\<preceq>(G,L) \<longrightarrow> (\<forall>T C. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) \<longrightarrow>  
-  Q Y' s' Z \<and> s'\<Colon>\<preceq>(G, L)))) \<Longrightarrow>  
+  (\<forall>L. s\<Colon>\<preceq>(G,L) \<longrightarrow> 
+    (\<forall>T C A. (normal s \<longrightarrow> (\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) \<and> 
+                            \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A) \<longrightarrow>
+    Q Y' s' Z \<and> s'\<Colon>\<preceq>(G, L)))) \<Longrightarrow>  
   G,A|\<Turnstile>\<Colon>{ {P} c\<succ> {Q}}"
 apply (simp (no_asm) add: ax_valids2_def triple_valid2_def2)
 apply clarsimp
 done
 
-ML_setup {*
-Delsimprocs [evaln_expr_proc,evaln_var_proc,evaln_exprs_proc,evaln_stmt_proc]
-*}
+lemma da_good_approx_evalnE [consumes 4]:
+  assumes evaln: "G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s1)"
+     and     wt: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T"
+     and     da: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>t\<guillemotright> A"
+     and     wf: "wf_prog G"
+     and   elim: "\<lbrakk>normal s1 \<Longrightarrow> nrm A \<subseteq> dom (locals (store s1));
+                  \<And> l. \<lbrakk>abrupt s1 = Some (Jump (Break l)); normal s0\<rbrakk>
+                        \<Longrightarrow> brk A l \<subseteq> dom (locals (store s1));
+                   \<lbrakk>abrupt s1 = Some (Jump Ret);normal s0\<rbrakk>
+                   \<Longrightarrow>Result \<in> dom (locals (store s1))
+                  \<rbrakk> \<Longrightarrow> P"
+  shows "P"
+proof -
+  from evaln have "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v, s1)"
+    by (rule evaln_eval)
+  from this wt da wf elim show P
+    by (rule da_good_approxE') rules+
+qed
 
-lemma Loop_sound: "\<lbrakk>G,A|\<Turnstile>\<Colon>{ {P} e-\<succ> {P'}};  
-       G,A|\<Turnstile>\<Colon>{ {Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}}\<rbrakk> \<Longrightarrow>  
-       G,A|\<Turnstile>\<Colon>{ {P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}}"
-apply (rule valids2_inductI)
-apply ((rule allI)+, rule impI, tactic {* pair_tac "s" 1*}, tactic {* pair_tac "s'" 1*})
-apply (erule evaln.induct)
-apply  simp_all (* takes half a minute *)
-apply  clarify
-apply  (erule_tac V = "G,A|\<Turnstile>\<Colon>{ {?P'} .c. {?P}}" in thin_rl)
-apply  (simp_all (no_asm_use) add: ax_valids2_def triple_valid2_def2)
-apply  (tactic "smp_tac 1 1", tactic "smp_tac 3 1", force)
-apply clarify
-apply (rule wt_elim_cases, assumption)
-apply (tactic "smp_tac 1 1", tactic "smp_tac 1 1", tactic "smp_tac 3 1", 
-       tactic "smp_tac 2 1", tactic "smp_tac 1 1")
-apply (erule impE,simp (no_asm),blast)
-apply (simp add: imp_conjL split_tupled_all split_paired_All)
-apply (case_tac "the_Bool b")
-apply  clarsimp
-apply  (case_tac "a")
-apply (simp_all)
-apply clarsimp
-apply  (erule_tac V = "c = l\<bullet> While(e) c \<longrightarrow> ?P" in thin_rl)
-apply (blast intro: conforms_absorb)
-apply blast+
+lemma validI: 
+   assumes I: "\<And> n s0 L accC T C v s1 Y Z.
+               \<lbrakk>\<forall>t\<in>A. G\<Turnstile>n\<Colon>t; s0\<Colon>\<preceq>(G,L); 
+               normal s0 \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T;
+               normal s0 \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>t\<guillemotright>C;
+               G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1); P Y s0 Z\<rbrakk> \<Longrightarrow> Q v s1 Z \<and> s1\<Colon>\<preceq>(G,L)" 
+  shows "G,A|\<Turnstile>\<Colon>{ {P} t\<succ> {Q} }"
+apply (simp add: ax_valids2_def triple_valid2_def2)
+apply (intro allI impI)
+apply (case_tac "normal s")
+apply   clarsimp 
+apply   (rule I,(assumption|simp)+)
+
+apply   (rule I,auto)
+done
+  
+
+
+
+ML "Addsimprocs [wt_expr_proc,wt_var_proc,wt_exprs_proc,wt_stmt_proc]"
+
+lemma valid_stmtI: 
+   assumes I: "\<And> n s0 L accC C s1 Y Z.
+             \<lbrakk>\<forall>t\<in>A. G\<Turnstile>n\<Colon>t; s0\<Colon>\<preceq>(G,L); 
+              normal s0\<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>;
+              normal s0\<Longrightarrow>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>C;
+              G\<turnstile>s0 \<midarrow>c\<midarrow>n\<rightarrow> s1; P Y s0 Z\<rbrakk> \<Longrightarrow> Q \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G,L)" 
+  shows "G,A|\<Turnstile>\<Colon>{ {P} \<langle>c\<rangle>\<^sub>s\<succ> {Q} }"
+apply (simp add: ax_valids2_def triple_valid2_def2)
+apply (intro allI impI)
+apply (case_tac "normal s")
+apply   clarsimp 
+apply   (rule I,(assumption|simp)+)
+
+apply   (rule I,auto)
 done
 
-declare subst_Bool_def2 [simp del]
+lemma valid_stmt_NormalI: 
+   assumes I: "\<And> n s0 L accC C s1 Y Z.
+               \<lbrakk>\<forall>t\<in>A. G\<Turnstile>n\<Colon>t; s0\<Colon>\<preceq>(G,L); normal s0; \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>;
+               \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>C;
+               G\<turnstile>s0 \<midarrow>c\<midarrow>n\<rightarrow> s1; (Normal P) Y s0 Z\<rbrakk> \<Longrightarrow> Q \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G,L)" 
+  shows "G,A|\<Turnstile>\<Colon>{ {Normal P} \<langle>c\<rangle>\<^sub>s\<succ> {Q} }"
+apply (simp add: ax_valids2_def triple_valid2_def2)
+apply (intro allI impI)
+apply (elim exE conjE)
+apply (rule I)
+by auto
+
+lemma valid_var_NormalI: 
+   assumes I: "\<And> n s0 L accC T C vf s1 Y Z.
+               \<lbrakk>\<forall>t\<in>A. G\<Turnstile>n\<Colon>t; s0\<Colon>\<preceq>(G,L); normal s0; 
+                \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>=T;
+                \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>t\<rangle>\<^sub>v\<guillemotright>C;
+                G\<turnstile>s0 \<midarrow>t=\<succ>vf\<midarrow>n\<rightarrow> s1; (Normal P) Y s0 Z\<rbrakk> 
+               \<Longrightarrow> Q (In2 vf) s1 Z \<and> s1\<Colon>\<preceq>(G,L)"
+   shows "G,A|\<Turnstile>\<Colon>{ {Normal P} \<langle>t\<rangle>\<^sub>v\<succ> {Q} }"
+apply (simp add: ax_valids2_def triple_valid2_def2)
+apply (intro allI impI)
+apply (elim exE conjE)
+apply simp
+apply (rule I)
+by auto
+
+lemma valid_expr_NormalI: 
+   assumes I: "\<And> n s0 L accC T C v s1 Y Z.
+               \<lbrakk>\<forall>t\<in>A. G\<Turnstile>n\<Colon>t; s0\<Colon>\<preceq>(G,L); normal s0; 
+                \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>-T;
+                \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>t\<rangle>\<^sub>e\<guillemotright>C;
+                G\<turnstile>s0 \<midarrow>t-\<succ>v\<midarrow>n\<rightarrow> s1; (Normal P) Y s0 Z\<rbrakk> 
+               \<Longrightarrow> Q (In1 v) s1 Z \<and> s1\<Colon>\<preceq>(G,L)"
+   shows "G,A|\<Turnstile>\<Colon>{ {Normal P} \<langle>t\<rangle>\<^sub>e\<succ> {Q} }"
+apply (simp add: ax_valids2_def triple_valid2_def2)
+apply (intro allI impI)
+apply (elim exE conjE)
+apply simp
+apply (rule I)
+by auto
+
+lemma valid_expr_list_NormalI: 
+   assumes I: "\<And> n s0 L accC T C vs s1 Y Z.
+               \<lbrakk>\<forall>t\<in>A. G\<Turnstile>n\<Colon>t; s0\<Colon>\<preceq>(G,L); normal s0; 
+                \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>\<doteq>T;
+                \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>t\<rangle>\<^sub>l\<guillemotright>C;
+                G\<turnstile>s0 \<midarrow>t\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s1; (Normal P) Y s0 Z\<rbrakk> 
+                \<Longrightarrow> Q (In3 vs) s1 Z \<and> s1\<Colon>\<preceq>(G,L)"
+   shows "G,A|\<Turnstile>\<Colon>{ {Normal P} \<langle>t\<rangle>\<^sub>l\<succ> {Q} }"
+apply (simp add: ax_valids2_def triple_valid2_def2)
+apply (intro allI impI)
+apply (elim exE conjE)
+apply simp
+apply (rule I)
+by auto
+
+lemma validE [consumes 5]: 
+  assumes valid: "G,A|\<Turnstile>\<Colon>{ {P} t\<succ> {Q} }"
+   and    P: "P Y s0 Z"
+   and    valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+   and    conf: "s0\<Colon>\<preceq>(G,L)"
+   and    eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)"
+   and    wt: "normal s0 \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T"
+   and    da: "normal s0 \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>t\<guillemotright>C"
+   and    elim: "\<lbrakk>Q v s1 Z; s1\<Colon>\<preceq>(G,L)\<rbrakk> \<Longrightarrow> concl" 
+  shows "concl"
+using prems
+by (simp add: ax_valids2_def triple_valid2_def2) fast
+(* why consumes 5?. If I want to apply this lemma in a context wgere
+   \<not> normal s0 holds,
+   I can chain "\<not> normal s0" as fact number 6 and apply the rule with
+   cases. Auto will then solve premise 6 and 7.
+*)
+
 lemma all_empty: "(!x. P) = P"
 by simp
-lemma sound_valid2_lemma: 
-"\<lbrakk>\<forall>v n. Ball A (triple_valid2 G n) \<longrightarrow> P v n; Ball A (triple_valid2 G n)\<rbrakk>
- \<Longrightarrow>P v n"
-by blast
-ML {*
-val fullsimptac = full_simp_tac(simpset() delsimps [thm "all_empty"]);
-val sound_prepare_tac = EVERY'[REPEAT o thin_tac "?x \<in> ax_derivs G",
- full_simp_tac (simpset()addsimps[thm "ax_valids2_def",thm "triple_valid2_def2",
-                                  thm "imp_conjL"] delsimps[thm "all_empty"]),
- Clarify_tac];
-val sound_elim_tac = EVERY'[eresolve_tac (thms "evaln_elim_cases"), 
-        TRY o eresolve_tac (thms "wt_elim_cases"), fullsimptac, Clarify_tac];
-val sound_valid2_tac = REPEAT o FIRST'[smp_tac 1, 
-                  datac (thm "sound_valid2_lemma") 1];
-val sound_forw_hyp_tac = 
- EVERY'[smp_tac 3 
-          ORELSE' EVERY'[dtac spec,dtac spec, dtac spec,etac impE, Fast_tac] 
-          ORELSE' EVERY'[dtac spec,dtac spec,etac impE, Fast_tac],
-        fullsimptac, 
-        smp_tac 2,TRY o smp_tac 1,
-        TRY o EVERY'[etac impE, TRY o rtac impI, 
-        atac ORELSE' (EVERY' [REPEAT o rtac exI,Blast_tac]),
-        fullsimptac, Clarify_tac, TRY o smp_tac 1]]
-*}
-(* ### rtac conjI,rtac HOL.refl *)
-lemma Call_sound: 
-"\<lbrakk>wf_prog G; G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q}}; \<forall>a. G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>Val a} ps\<doteq>\<succ> {R a}};
-  \<forall>a vs invC declC l. G,A|\<Turnstile>\<Colon>{ {(R a\<leftarrow>Vals vs \<and>.  
-   (\<lambda>s. declC = invocation_declclass 
-                    G mode (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> \<and>
-         invC = invocation_class mode (store s) a statT \<and>
-            l = locals (store s)) ;.  
-   init_lvars G declC \<lparr>name=mn,parTs=pTs\<rparr> mode a vs) \<and>.  
-   (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)}  
-   Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}}\<rbrakk> \<Longrightarrow>  
-  G,A|\<Turnstile>\<Colon>{ {Normal P} {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 (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 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 
-apply   (force split add: split_if_asm)
-apply  (simp del: if_raise_if_None)
-apply  (tactic "smp_tac 2 1")
-apply (simp only: init_lvars_def2 invmode_Static_eq)
-apply (clarsimp simp del: resTy_mthd)
-apply  (drule spec,erule swap,erule conforms_set_locals [OF _ lconf_empty])
-apply clarsimp
-apply (drule Null_staticD)
-apply (drule eval_gext', drule (1) conf_gext, frule (3) DynT_propI)
-apply (erule impE) apply blast
-apply (subgoal_tac 
- "G\<turnstile>invmode (mhd (statDeclC,m)) e
-     \<rightarrow>invocation_class (invmode m e) s2 a' statT\<preceq>statT")
-defer   apply simp
-apply (drule (3) DynT_mheadsD,simp,simp)
-apply (clarify, drule wf_mdeclD1, clarify)
-apply (frule ty_expr_is_type) apply simp
-apply (subgoal_tac "invmode (mhd (statDeclC,m)) e = IntVir \<longrightarrow> a' \<noteq> Null")
-defer   apply simp
-apply (frule (2) wt_MethdI)
-apply clarify
-apply (drule (2) conforms_init_lvars)
-apply   (simp) 
-apply   (assumption)+
-apply   simp
-apply   (assumption)+
-apply   (rule impI) apply simp
-apply   simp
-apply   simp
-apply   (rule Ball_weaken)
-apply     assumption
-apply     (force simp add: is_acc_type_def)
-apply (tactic "smp_tac 2 1")
-apply simp
-apply (tactic "smp_tac 1 1")
-apply (erule_tac V = "?P \<longrightarrow> ?Q" in thin_rl) 
-apply (erule impE)
-apply   (rule exI)+ 
-apply   (subgoal_tac "is_static dm = (static m)") 
-prefer 2  apply (simp add: member_is_static_simp)
-apply   (simp only: )
-apply   (simp only: sig.simps)
-apply (force dest!: evaln_eval eval_gext' elim: conforms_return 
-             del: impCE simp add: init_lvars_def2)
-done
 
 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
+             da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>t\<guillemotright> A" and
         conf_s0: "s0\<Colon>\<preceq>(G,L)" and
              wf: "wf_prog G"                         
   shows "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)
+  from evaln have "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
+    by (rule evaln_eval)
+  from this wt da wf conf_s0 show ?thesis
+    by (rule 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}};  
-  \<forall>l. G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars empty}  
-            .init c. {set_lvars l .; R}}\<rbrakk> \<Longrightarrow>  
-      G,A|\<Turnstile>\<Colon>{ {Normal (P \<and>. Not \<circ> initd C)} .Init C. {R}}"
-apply (tactic "EVERY'[sound_prepare_tac, sound_elim_tac,sound_valid2_tac] 1")
-apply (tactic {* instantiate_tac [("l24","\<lambda> n Y Z sa Y' s' L y a b aa ba ab bb. locals b")]*})
-apply (clarsimp simp add: split_paired_Ex)
-apply (drule spec, drule spec, drule spec, erule impE)
-apply  (erule_tac V = "All ?P" in thin_rl, fast)
-apply clarsimp
-apply (tactic "smp_tac 2 1", drule spec, erule impE, 
-       erule (3) conforms_init_class_obj)
-apply (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
-apply (drule spec)
-apply (drule spec)
-apply (drule spec)
-apply  (erule impE)
-apply ( fast)
-apply (simp (no_asm_use) del: empty_def2)
-apply (tactic "smp_tac 2 1")
-apply (drule spec, erule impE, erule conforms_set_locals, rule lconf_empty)
-apply (erule impE,rule impI,rule exI,erule wf_cdecl_wt_init)
-apply clarsimp
-apply (erule (1) conforms_return)
-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"
-by fast
-
-lemma all4_conjunct2: 
-  "\<forall>a vs D l. (P' a vs D l \<and> P a vs D l) \<Longrightarrow> \<forall>a vs D l. P a vs D l"
-by fast
-
-
-lemmas sound_lemmas = Init_sound Loop_sound Methd_sound
-
-lemma ax_sound2: "wf_prog G \<Longrightarrow> G,A|\<turnstile>ts \<Longrightarrow> G,A|\<Turnstile>\<Colon>ts"
-apply (erule ax_derivs.induct)
-prefer 22 (* Call *)
-apply (erule (1) Call_sound) apply simp apply force apply force 
-
-apply (tactic {* TRYALL (eresolve_tac (thms "sound_lemmas") THEN_ALL_NEW 
-    eresolve_tac [asm_rl, thm "all_conjunct2", thm "all4_conjunct2"]) *})
-
-apply(tactic "COND (has_fewer_prems(30+9)) (ALLGOALS sound_prepare_tac) no_tac")
-
-               (*empty*)
-apply        fast (* insert *)
-apply       fast (* asm *)
-(*apply    fast *) (* cut *)
-apply     fast (* weaken *)
-apply    (tactic "smp_tac 3 1", clarify, tactic "smp_tac 1 1",
-          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 *)
-
-prefer 28 apply (simp add: evaln_InsInitV)
-prefer 28 apply (simp add: evaln_InsInitE)
-prefer 28 apply (simp add: evaln_Callee)
-prefer 28 apply (simp add: evaln_FinA)
-
-(* 27 subgoals *)
-apply (tactic {* sound_elim_tac 1 *})
-apply (tactic {* ALLGOALS sound_elim_tac *})(* LVar, Lit, Super, Nil, Skip,Do *)
-apply (tactic {* ALLGOALS (asm_simp_tac (noAll_simpset() 
-                          delsimps [thm "all_empty"])) *})    (* Done *)
-(* for FVar *)
-apply (frule wf_ws_prog) 
-apply (frule ty_expr_is_type [THEN type_is_class, 
-                              THEN accfield_declC_is_class])
-apply (simp (no_asm_use), simp (no_asm_use), simp (no_asm_use))
-apply (frule_tac [4] wt_init_comp_ty) (* for NewA*)
-apply (tactic "ALLGOALS sound_valid2_tac")
-apply (tactic "TRYALL sound_forw_hyp_tac") (* UnOp, Cast, Inst, Acc, Expr *)
-apply (tactic {* TRYALL (EVERY'[dtac spec, TRY o EVERY'[rotate_tac ~1, 
-  dtac spec], dtac conjunct2, smp_tac 1, 
-  TRY o dres_inst_tac [("P","P'")] (thm "subst_Bool_the_BoolI")]) *})
-apply (frule_tac [15] x = x1 in conforms_NormI)  (* for Fin *)
-
-(* 15 subgoals *)
-(* FVar *)
-apply (tactic "sound_forw_hyp_tac 1")
-apply (clarsimp simp add: fvar_def2 Let_def split add: split_if_asm)
-
-(* AVar *)
-(*
-apply (drule spec, drule spec, erule impE, fast)
-apply (simp)
-apply (tactic "smp_tac 2 1")
-apply (tactic "smp_tac 1 1")
-apply (erule impE)
-apply (rule impI)
-apply (rule exI)+
-apply blast
-apply (clarsimp simp add: avar_def2)
-*)
-apply (tactic "sound_forw_hyp_tac 1")
-apply (clarsimp simp add: avar_def2)
-
-(* NewC *)
-apply (clarsimp simp add: is_acc_class_def)
-apply (erule (1) halloc_conforms, simp, simp)
-
-(* NewA *)
-apply (tactic "sound_forw_hyp_tac 1")
-apply (rule conjI,blast)
-apply (erule (1) halloc_conforms, simp, simp, simp add: is_acc_type_def)
-
-(* BinOp *)
-apply (tactic "sound_forw_hyp_tac 1")
-apply (case_tac "need_second_arg binop v1")
-apply   fastsimp
-apply   simp
-
-(* Ass *)
-apply (tactic "sound_forw_hyp_tac 1")
-apply (case_tac "aa")
-prefer 2
-apply  clarsimp
-apply (drule (3) evaln_type_sound)
-apply (drule (3) evaln_eval)
-apply (frule (3) eval_type_sound)
-apply clarsimp
-apply (frule wf_ws_prog)
-apply (drule (2) conf_widen)
-apply (drule_tac "s1.0" = b in eval_gext')
-apply (clarsimp simp add: assign_conforms_def)
-
-
-(* Cond *)
-apply (tactic "smp_tac 3 1") apply (tactic "smp_tac 2 1") 
-apply (tactic "smp_tac 1 1") apply (erule impE) 
-apply (rule impI,rule exI) 
-apply (rule_tac x = "if the_Bool b then T1 else T2" in exI)
-apply (force split add: split_if)
-apply assumption
-
-(* Body *)
-apply (tactic "sound_forw_hyp_tac 1")
-apply (rule conforms_absorb,assumption)
-
-(* Lab *)
-apply (tactic "sound_forw_hyp_tac 1")
-apply (rule conforms_absorb,assumption)
-
-(* If *)
-apply (tactic "sound_forw_hyp_tac 1")
-apply (tactic "sound_forw_hyp_tac 1")
-apply (force split add: split_if)
-
-(* Throw *)
-apply (drule (3) evaln_type_sound)
-apply clarsimp
-apply (drule (3) Throw_lemma)
-apply clarsimp
-
-(* Try *)
-apply (frule (1) sxalloc_type_sound)
-apply (erule sxalloc_elim_cases2)
-apply  (tactic "smp_tac 3 1")
-apply  (clarsimp split add: option.split_asm)
-apply (clarsimp split add: option.split_asm)
-apply (tactic "smp_tac 1 1")
-apply (simp only: split add: split_if_asm)
-prefer 2
-apply  (tactic "smp_tac 3 1", erule_tac V = "All ?P" in thin_rl, clarsimp)
-apply (drule spec, erule_tac V = "All ?P" in thin_rl, drule spec, drule spec, 
-       erule impE, force)
-apply (frule (2) Try_lemma)
-apply clarsimp
-apply (fast elim!: conforms_deallocL)
-
-(* Fin *)
-apply (tactic "sound_forw_hyp_tac 1")
-apply (case_tac "x1", force)
-apply clarsimp
-apply (drule (3) evaln_eval, drule (4) Fin_lemma)
-done
+corollary dom_locals_evaln_mono_elim [consumes 1]: 
+  assumes   
+  evaln: "G\<turnstile> s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)" and
+    hyps: "\<lbrakk>dom (locals (store s0)) \<subseteq> dom (locals (store s1));
+           \<And> vv s val. \<lbrakk>v=In2 vv; normal s1\<rbrakk> 
+                        \<Longrightarrow> dom (locals (store s)) 
+                             \<subseteq> dom (locals (store ((snd vv) val s)))\<rbrakk> \<Longrightarrow> P"
+ shows "P"
+proof -
+  from evaln have "G\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" by (rule evaln_eval)
+  from this hyps show ?thesis
+    by (rule dom_locals_eval_mono_elim) rules+
+qed
 
 
 
-declare subst_Bool_def2 [simp]
+lemma evaln_no_abrupt: 
+   "\<And>s s'. \<lbrakk>G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (w,s'); normal s'\<rbrakk> \<Longrightarrow> normal s"
+by (erule evaln_cases,auto)
+
+declare inj_term_simps [simp]
+lemma ax_sound2: 
+  assumes    wf: "wf_prog G" 
+    and   deriv: "G,A|\<turnstile>ts"
+  shows "G,A|\<Turnstile>\<Colon>ts"
+using deriv
+proof (induct)
+  case (empty A)
+  show ?case
+    by (simp add: ax_valids2_def triple_valid2_def2)
+next
+  case (insert A t ts)
+  have valid_t: "G,A|\<Turnstile>\<Colon>{t}" . 
+  moreover have valid_ts: "G,A|\<Turnstile>\<Colon>ts" .
+  {
+    fix n assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    have "G\<Turnstile>n\<Colon>t" and "\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t"
+    proof -
+      from valid_A valid_t show "G\<Turnstile>n\<Colon>t"
+	by (simp add: ax_valids2_def)
+    next
+      from valid_A valid_ts show "\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t"
+	by (unfold ax_valids2_def) blast
+    qed
+    hence "\<forall>t'\<in>insert t ts. G\<Turnstile>n\<Colon>t'"
+      by simp
+  }
+  thus ?case
+    by (unfold ax_valids2_def) blast
+next
+  case (asm A ts)
+  have "ts \<subseteq> A" .
+  then show "G,A|\<Turnstile>\<Colon>ts"
+    by (auto simp add: ax_valids2_def triple_valid2_def)
+next
+  case (weaken A ts ts')
+  have "G,A|\<Turnstile>\<Colon>ts'" .
+  moreover have "ts \<subseteq> ts'" .
+  ultimately show "G,A|\<Turnstile>\<Colon>ts"
+    by (unfold ax_valids2_def triple_valid2_def) blast
+next
+  case (conseq A P Q t)
+  have con: "\<forall>Y s Z. P Y s Z \<longrightarrow> 
+              (\<exists>P' Q'.
+                  (G,A\<turnstile>{P'} t\<succ> {Q'} \<and> G,A|\<Turnstile>\<Colon>{ {P'} t\<succ> {Q'} }) \<and>
+                  (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow> Q Y' s' Z))".
+  show "G,A|\<Turnstile>\<Colon>{ {P} t\<succ> {Q} }"
+  proof (rule validI)
+    fix n s0 L accC T C v s1 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t" 
+    assume conf: "s0\<Colon>\<preceq>(G,L)"
+    assume wt: "normal s0 \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T"
+    assume da: "normal s0 
+                 \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>t\<guillemotright> C"
+    assume eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s1)"
+    assume P: "P Y s0 Z"
+    show "Q v s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof -
+      from valid_A conf wt da eval P con
+      have "Q v s1 Z"
+	apply (simp add: ax_valids2_def triple_valid2_def2)
+	apply (tactic "smp_tac 3 1")
+	apply clarify
+	apply (tactic "smp_tac 1 1")
+	apply (erule allE,erule allE, erule mp)
+	apply (intro strip)
+	apply (tactic "smp_tac 3 1")
+	apply (tactic "smp_tac 2 1")
+	apply (tactic "smp_tac 1 1")
+	by blast
+      moreover have "s1\<Colon>\<preceq>(G, L)"
+      proof (cases "normal s0")
+	case True
+	from eval wt [OF True] da [OF True] conf wf 
+	show ?thesis
+	  by (rule evaln_type_sound [elim_format]) simp
+      next
+	case False
+	with eval have "s1=s0"
+	  by auto
+	with conf show ?thesis by simp
+      qed
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (hazard A P Q t)
+  show "G,A|\<Turnstile>\<Colon>{ {P \<and>. Not \<circ> type_ok G t} t\<succ> {Q} }"
+    by (simp add: ax_valids2_def triple_valid2_def2 type_ok_def) fast
+next
+  case (Abrupt A P t)
+  show "G,A|\<Turnstile>\<Colon>{ {P\<leftarrow>arbitrary3 t \<and>. Not \<circ> normal} t\<succ> {P} }"
+  proof (rule validI)
+    fix n s0 L accC T C v s1 Y Z 
+    assume conf_s0: "s0\<Colon>\<preceq>(G, L)"
+    assume eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s1)"
+    assume "(P\<leftarrow>arbitrary3 t \<and>. Not \<circ> normal) Y s0 Z"
+    then obtain P: "P (arbitrary3 t) s0 Z" and abrupt_s0: "\<not> normal s0"
+      by simp
+    from eval abrupt_s0 obtain "s1=s0" and "v=arbitrary3 t"
+      by auto
+    with P conf_s0
+    show "P v s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+      by simp
+  qed
+next
+  case (LVar A P vn)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal (\<lambda>s.. P\<leftarrow>In2 (lvar vn s))} LVar vn=\<succ> {P} }"
+  proof (rule valid_var_NormalI)
+    fix n s0 L accC T C vf s1 Y Z
+    assume conf_s0: "s0\<Colon>\<preceq>(G, L)"
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>LVar vn\<Colon>=T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>LVar vn\<rangle>\<^sub>v\<guillemotright> C"
+    assume eval: "G\<turnstile>s0 \<midarrow>LVar vn=\<succ>vf\<midarrow>n\<rightarrow> s1" 
+    assume P: "(Normal (\<lambda>s.. P\<leftarrow>In2 (lvar vn s))) Y s0 Z"
+    show "P (In2 vf) s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof 
+      from eval normal_s0 obtain "s1=s0" "vf=lvar vn (store s0)"
+	by (fastsimp elim: evaln_elim_cases)
+      with P show "P (In2 vf) s1 Z"
+	by simp
+    next
+      from eval wt da conf_s0 wf
+      show "s1\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+    qed
+  qed
+next
+  case (FVar A statDeclC P Q R accC e fn stat)
+  have valid_init: "G,A|\<Turnstile>\<Colon>{ {Normal P} .Init statDeclC. {Q} }" .
+  have valid_e: "G,A|\<Turnstile>\<Colon>{ {Q} e-\<succ> {\<lambda>Val:a:. fvar statDeclC stat fn a ..; R} }" .
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} {accC,statDeclC,stat}e..fn=\<succ> {R} }"
+  proof (rule valid_var_NormalI)
+    fix n s0 L accC' T V vf s3 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile>{accC,statDeclC,stat}e..fn\<Colon>=T"
+    assume da: "\<lparr>prg=G,cls=accC',lcl=L\<rparr>
+                  \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>{accC,statDeclC,stat}e..fn\<rangle>\<^sub>v\<guillemotright> V"
+    assume eval: "G\<turnstile>s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>vf\<midarrow>n\<rightarrow> s3"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<lfloor>vf\<rfloor>\<^sub>v s3 Z \<and> s3\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt 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'" and
+        stat: "stat=is_static f" and
+	T: "T=(type f)"
+	by (cases) (auto simp add: member_is_static_simp)
+      from da eq_accC
+      have da_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V"
+	by cases simp
+      from eval obtain a s1 s2 s2' where
+	eval_init: "G\<turnstile>s0 \<midarrow>Init statDeclC\<midarrow>n\<rightarrow> s1" and 
+        eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s2" and 
+	fvar: "(vf,s2')=fvar statDeclC stat fn a s2" and
+	s3: "s3 = check_field_access G accC statDeclC fn stat a s2'"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases) 
+      have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
+      proof -
+	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)
+	thus ?thesis by simp
+      qed
+      obtain I where 
+	da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                    \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init statDeclC\<rangle>\<^sub>s\<guillemotright> I"
+	by (auto intro: da_Init [simplified] assigned.select_convs)
+      from valid_init P valid_A conf_s0 eval_init wt_init da_init
+      obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
+	by (rule validE)
+      obtain 
+	R: "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and 
+        conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+	conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+      proof (cases "normal s1")
+	case True
+	obtain V' where 
+	  da_e':
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V'"
+	proof -
+	  from eval_init 
+	  have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
+	    by (rule dom_locals_evaln_mono_elim)
+	  with da_e show ?thesis
+	    by (rule da_weakenE)
+	qed
+	with valid_e Q valid_A conf_s1 eval_e wt_e
+	obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
+	  by (rule validE) (simp add: fvar [symmetric])
+	moreover
+	from eval_e wt_e da_e' conf_s1 wf
+	have "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+	  by (rule evaln_type_sound [elim_format]) simp
+	ultimately show ?thesis ..
+      next
+	case False
+	with valid_e Q valid_A conf_s1 eval_e
+	obtain  "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
+	  by (cases rule: validE) (simp add: fvar [symmetric])+
+	moreover from False eval_e have "\<not> normal s2"
+	  by auto
+	hence "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+	  by auto
+	ultimately show ?thesis ..
+      qed
+      from accfield wt_e eval_init eval_e conf_s2 conf_a fvar stat s3 wf
+      have eq_s3_s2': "s3=s2'"  
+	using normal_s0 by (auto dest!: error_free_field_access evaln_eval)
+      moreover
+      from eval wt da conf_s0 wf
+      have "s3\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis using Q by simp
+    qed
+  qed
+next   
+  case (AVar A P Q R e1 e2)
+  have valid_e1: "G,A|\<Turnstile>\<Colon>{ {Normal P} e1-\<succ> {Q} }" .
+  have valid_e2: "\<And> a. G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>In1 a} e2-\<succ> {\<lambda>Val:i:. avar G i a ..; R} }"
+    using AVar.hyps by simp
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} e1.[e2]=\<succ> {R} }"
+  proof (rule valid_var_NormalI)
+    fix n s0 L accC T V vf s2' Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1.[e2]\<Colon>=T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1.[e2]\<rangle>\<^sub>v\<guillemotright> V"
+    assume eval: "G\<turnstile>s0 \<midarrow>e1.[e2]=\<succ>vf\<midarrow>n\<rightarrow> s2'"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z \<and> s2'\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain 
+	wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T.[]" and
+        wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-PrimT Integer" 
+	by (rule wt_elim_cases) simp
+      from da obtain E1 where
+	da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
+	da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V"
+	by (rule da_elim_cases) simp
+      from eval obtain s1 a i s2 where
+	eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1" and
+	eval_e2: "G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2" and
+	avar: "avar G i a s2 =(vf, s2')"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_e1 P valid_A conf_s0 eval_e1 wt_e1 da_e1
+      obtain Q: "Q \<lfloor>a\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
+	by (rule validE)
+      from Q have Q': "\<And> v. (Q\<leftarrow>In1 a) v s1 Z"
+	by simp
+      have "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z"
+      proof (cases "normal s1")
+	case True
+	obtain V' where 
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V'"
+	proof -
+	  from eval_e1  wt_e1 da_e1 wf True
+	  have "nrm E1 \<subseteq> dom (locals (store s1))"
+	    by (cases rule: da_good_approx_evalnE) rules
+	  with da_e2 show ?thesis
+	    by (rule da_weakenE)
+	qed
+	with valid_e2 Q' valid_A conf_s1 eval_e2 wt_e2 
+	show ?thesis
+	  by (rule validE) (simp add: avar)
+      next
+	case False
+	with valid_e2 Q' valid_A conf_s1 eval_e2
+	show ?thesis
+	  by (cases rule: validE) (simp add: avar)+
+      qed
+      moreover
+      from eval wt da conf_s0 wf
+      have "s2'\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (NewC A C P Q)
+  have valid_init: "G,A|\<Turnstile>\<Colon>{ {Normal P} .Init C. {Alloc G (CInst C) Q} }".
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} NewC C-\<succ> {Q} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC T E v s2 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>NewC C\<Colon>-T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>NewC C\<rangle>\<^sub>e\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>NewC C-\<succ>v\<midarrow>n\<rightarrow> s2"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain is_cls_C: "is_class G C" 
+	by (rule wt_elim_cases) (auto dest: is_acc_classD)
+      hence wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>Init C\<Colon>\<surd>" 
+	by auto
+      obtain I where 
+	da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init C\<rangle>\<^sub>s\<guillemotright> I"
+	by (auto intro: da_Init [simplified] assigned.select_convs)
+      from eval obtain s1 a where
+	eval_init: "G\<turnstile>s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1" and 
+        alloc: "G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2" and
+	v: "v=Addr a"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_init P valid_A conf_s0 eval_init wt_init da_init
+      obtain "(Alloc G (CInst C) Q) \<diamondsuit> s1 Z" 
+	by (rule validE)
+      with alloc v have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (NewA A P Q R T e)
+  have valid_init: "G,A|\<Turnstile>\<Colon>{ {Normal P} .init_comp_ty T. {Q} }" .
+  have valid_e: "G,A|\<Turnstile>\<Colon>{ {Q} e-\<succ> {\<lambda>Val:i:. abupd (check_neg i) .; 
+                                            Alloc G (Arr T (the_Intg i)) R}}" .
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} New T[e]-\<succ> {R} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC arrT E v s3 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>New T[e]\<Colon>-arrT"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>New T[e]\<rangle>\<^sub>e\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>New T[e]-\<succ>v\<midarrow>n\<rightarrow> s3"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z \<and> s3\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain
+	wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and 
+	wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Integer" 
+	by (rule wt_elim_cases) (auto intro: wt_init_comp_ty )
+      from da obtain
+	da_e:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+	by cases simp
+      from eval obtain s1 i s2 a where
+	eval_init: "G\<turnstile>s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1" and 
+        eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>i\<midarrow>n\<rightarrow> s2" and
+        alloc: "G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3" and
+        v: "v=Addr a"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      obtain I where
+	da_init:
+	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>init_comp_ty T\<rangle>\<^sub>s\<guillemotright> I"
+      proof (cases "\<exists>C. T = Class C")
+	case True
+	thus ?thesis
+	  by - (rule that, (auto intro: da_Init [simplified] 
+                                        assigned.select_convs
+                              simp add: init_comp_ty_def))
+	 (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
+      next
+	case False
+	thus ?thesis
+	  by - (rule that, (auto intro: da_Skip [simplified] 
+                                      assigned.select_convs
+                           simp add: init_comp_ty_def))
+         (* simplified: to rewrite \<langle>Skip\<rangle> to In1r (Skip) *)
+      qed
+      with valid_init P valid_A conf_s0 eval_init wt_init 
+      obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
+	by (rule validE)
+      obtain E' where
+       "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
+      proof -
+	from eval_init 
+	have "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+	  by (rule dom_locals_evaln_mono_elim)
+	with da_e show ?thesis
+	  by (rule da_weakenE)
+      qed
+      with valid_e Q valid_A conf_s1 eval_e wt_e
+      have "(\<lambda>Val:i:. abupd (check_neg i) .; 
+                      Alloc G (Arr T (the_Intg i)) R) \<lfloor>i\<rfloor>\<^sub>e s2 Z"
+	by (rule validE)
+      with alloc v have "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
+	by simp
+      moreover 
+      from eval wt da conf_s0 wf
+      have "s3\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Cast A P Q T e)
+  have valid_e: "G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> 
+                 {\<lambda>Val:v:. \<lambda>s.. abupd (raise_if (\<not> G,s\<turnstile>v fits T) ClassCast) .;
+                  Q\<leftarrow>In1 v} }" .
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} Cast T e-\<succ> {Q} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC castT E v s2 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Cast T e\<Colon>-castT"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>Cast T e\<rangle>\<^sub>e\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>Cast T e-\<succ>v\<midarrow>n\<rightarrow> s2"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain eT where 
+	wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
+	by cases simp
+      from da obtain
+	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+	by cases simp
+      from eval obtain s1 where
+	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1" and
+        s2: "s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits T) ClassCast) s1"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_e P valid_A conf_s0 eval_e wt_e da_e
+      have "(\<lambda>Val:v:. \<lambda>s.. abupd (raise_if (\<not> G,s\<turnstile>v fits T) ClassCast) .;
+                  Q\<leftarrow>In1 v) \<lfloor>v\<rfloor>\<^sub>e s1 Z"
+	by (rule validE)
+      with s2 have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Inst A P Q T e)
+  assume valid_e: "G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ>
+               {\<lambda>Val:v:. \<lambda>s.. Q\<leftarrow>In1 (Bool (v \<noteq> Null \<and> G,s\<turnstile>v fits RefT T))} }"
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} e InstOf T-\<succ> {Q} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC instT E v s1 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e InstOf T\<Colon>-instT"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e InstOf T\<rangle>\<^sub>e\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>e InstOf T-\<succ>v\<midarrow>n\<rightarrow> s1"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain eT where 
+	wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
+	by cases simp
+      from da obtain
+	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+	by cases simp
+      from eval obtain a where
+	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
+        v: "v = Bool (a \<noteq> Null \<and> G,store s1\<turnstile>a fits RefT T)"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_e P valid_A conf_s0 eval_e wt_e da_e
+      have "(\<lambda>Val:v:. \<lambda>s.. Q\<leftarrow>In1 (Bool (v \<noteq> Null \<and> G,s\<turnstile>v fits RefT T))) 
+              \<lfloor>a\<rfloor>\<^sub>e s1 Z"
+	by (rule validE)
+      with v have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s1\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Lit A P v)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal (P\<leftarrow>In1 v)} Lit v-\<succ> {P} }"
+  proof (rule valid_expr_NormalI)
+    fix n L s0 s1 v'  Y Z
+    assume conf_s0: "s0\<Colon>\<preceq>(G, L)"
+    assume normal_s0: " normal s0"
+    assume eval: "G\<turnstile>s0 \<midarrow>Lit v-\<succ>v'\<midarrow>n\<rightarrow> s1"
+    assume P: "(Normal (P\<leftarrow>In1 v)) Y s0 Z"
+    show "P \<lfloor>v'\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof -
+      from eval have "s1=s0" and  "v'=v"
+	using normal_s0 by (auto elim: evaln_elim_cases)
+      with P conf_s0 show ?thesis by simp
+    qed
+  qed
+next
+  case (UnOp A P Q e unop)
+  assume valid_e: "G,A|\<Turnstile>\<Colon>{ {Normal P}e-\<succ>{\<lambda>Val:v:. Q\<leftarrow>In1 (eval_unop unop v)} }"
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} UnOp unop e-\<succ> {Q} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC T E v s1 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>UnOp unop e\<Colon>-T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>UnOp unop e\<rangle>\<^sub>e\<guillemotright>E"
+    assume eval: "G\<turnstile>s0 \<midarrow>UnOp unop e-\<succ>v\<midarrow>n\<rightarrow> s1"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain eT where 
+	wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
+	by cases simp
+      from da obtain
+	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+	by cases simp
+      from eval obtain ve where
+	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
+        v: "v = eval_unop unop ve"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_e P valid_A conf_s0 eval_e wt_e da_e
+      have "(\<lambda>Val:v:. Q\<leftarrow>In1 (eval_unop unop v)) \<lfloor>ve\<rfloor>\<^sub>e s1 Z"
+	by (rule validE)
+      with v have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s1\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (BinOp A P Q R binop e1 e2)
+  assume valid_e1: "G,A|\<Turnstile>\<Colon>{ {Normal P} e1-\<succ> {Q} }" 
+  have valid_e2: "\<And> v1.  G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>In1 v1}
+              (if need_second_arg binop v1 then In1l e2 else In1r Skip)\<succ>
+              {\<lambda>Val:v2:. R\<leftarrow>In1 (eval_binop binop v1 v2)} }"
+    using BinOp.hyps by simp
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} BinOp binop e1 e2-\<succ> {R} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC T E v s2 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>BinOp binop e1 e2\<Colon>-T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>BinOp binop e1 e2\<rangle>\<^sub>e\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>BinOp binop e1 e2-\<succ>v\<midarrow>n\<rightarrow> s2"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain e1T e2T where
+        wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-e1T" and
+        wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-e2T" and
+	wt_binop: "wt_binop G binop e1T e2T" 
+	by cases simp
+      have wt_Skip: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>Skip\<Colon>\<surd>"
+	by simp
+      (*
+      obtain S where
+	daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                   \<turnstile> dom (locals (store s1)) \<guillemotright>In1r Skip\<guillemotright> S"
+	by (auto intro: da_Skip [simplified] assigned.select_convs) *)
+      from da obtain E1 where
+	da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1"
+	by cases simp+
+      from eval obtain v1 s1 v2 where
+	eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<midarrow>n\<rightarrow> s1" and
+	eval_e2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)
+                        \<succ>\<midarrow>n\<rightarrow> (\<lfloor>v2\<rfloor>\<^sub>e, s2)" and
+        v: "v=eval_binop binop v1 v2"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_e1 P valid_A conf_s0 eval_e1 wt_e1 da_e1
+      obtain Q: "Q \<lfloor>v1\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
+	by (rule validE)
+      from Q have Q': "\<And> v. (Q\<leftarrow>In1 v1) v s1 Z"
+	by simp
+      have "(\<lambda>Val:v2:. R\<leftarrow>In1 (eval_binop binop v1 v2)) \<lfloor>v2\<rfloor>\<^sub>e s2 Z"
+      proof (cases "normal s1")
+	case True
+	from eval_e1 wt_e1 da_e1 conf_s0 wf
+	have conf_v1: "G,store s1\<turnstile>v1\<Colon>\<preceq>e1T" 
+	  by (rule evaln_type_sound [elim_format]) (insert True,simp)
+	from eval_e1 
+	have "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
+	  by (rule evaln_eval)
+	from da wt_e1 wt_e2 wt_binop conf_s0 True this conf_v1 wf
+	obtain E2 where
+	  da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) 
+                   \<guillemotright>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<guillemotright> E2"
+	  by (rule da_e2_BinOp [elim_format]) rules
+	from wt_e2 wt_Skip obtain T2 
+	  where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<Colon>T2"
+	  by (cases "need_second_arg binop v1") auto
+	note ve=validE [OF valid_e2,OF  Q' valid_A conf_s1 eval_e2 this da_e2]
+	(* chaining Q', without extra OF causes unification error *)
+	thus ?thesis
+	  by (rule ve)
+      next
+	case False
+	note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2]
+	with False show ?thesis
+	  by rules
+      qed
+      with v have "R \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Super A P)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal (\<lambda>s.. P\<leftarrow>In1 (val_this s))} Super-\<succ> {P} }"
+  proof (rule valid_expr_NormalI)
+    fix n L s0 s1 v  Y Z
+    assume conf_s0: "s0\<Colon>\<preceq>(G, L)"
+    assume normal_s0: " normal s0"
+    assume eval: "G\<turnstile>s0 \<midarrow>Super-\<succ>v\<midarrow>n\<rightarrow> s1"
+    assume P: "(Normal (\<lambda>s.. P\<leftarrow>In1 (val_this s))) Y s0 Z"
+    show "P \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof -
+      from eval have "s1=s0" and  "v=val_this (store s0)"
+	using normal_s0 by (auto elim: evaln_elim_cases)
+      with P conf_s0 show ?thesis by simp
+    qed
+  qed
+next
+  case (Acc A P Q var)
+  have valid_var: "G,A|\<Turnstile>\<Colon>{ {Normal P} var=\<succ> {\<lambda>Var:(v, f):. Q\<leftarrow>In1 v} }" .
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} Acc var-\<succ> {Q} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC T E v s1 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Acc var\<Colon>-T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>Acc var\<rangle>\<^sub>e\<guillemotright>E"
+    assume eval: "G\<turnstile>s0 \<midarrow>Acc var-\<succ>v\<midarrow>n\<rightarrow> s1"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain 
+	wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=T" 
+	by cases simp
+      from da obtain V where 
+	da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
+	by (cases "\<exists> n. var=LVar n") (insert da.LVar,auto elim!: da_elim_cases)
+      from eval obtain w upd where
+	eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(v, upd)\<midarrow>n\<rightarrow> s1"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_var P valid_A conf_s0 eval_var wt_var da_var
+      have "(\<lambda>Var:(v, f):. Q\<leftarrow>In1 v) \<lfloor>(v, upd)\<rfloor>\<^sub>v s1 Z"
+	by (rule validE)
+      then have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s1\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Ass A P Q R e var)
+  have valid_var: "G,A|\<Turnstile>\<Colon>{ {Normal P} var=\<succ> {Q} }" .
+  have valid_e: "\<And> vf. 
+                  G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>In2 vf} e-\<succ> {\<lambda>Val:v:. assign (snd vf) v .; R} }"
+    using Ass.hyps by simp
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} var:=e-\<succ> {R} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC T E v s3 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var:=e\<Colon>-T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>var:=e\<rangle>\<^sub>e\<guillemotright>E"
+    assume eval: "G\<turnstile>s0 \<midarrow>var:=e-\<succ>v\<midarrow>n\<rightarrow> s3"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z \<and> s3\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain varT  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>-T" 
+	by cases simp
+      from eval obtain w upd s1 s2 where
+	eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(w, upd)\<midarrow>n\<rightarrow> s1" and
+        eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s2" and
+	s3: "s3=assign upd v s2"
+	using normal_s0 by (auto elim: evaln_elim_cases)
+      have "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
+      proof (cases "\<exists> vn. var = LVar vn")
+	case False
+	with da obtain V where
+	  da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                      \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V" and
+	  da_e:   "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+	  by cases simp+
+	from valid_var P valid_A conf_s0 eval_var wt_var da_var
+	obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
+	  by (rule validE) 
+	hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
+	  by simp
+	have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+	proof (cases "normal s1")
+	  case True
+	  obtain E' where 
+	    da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
+	  proof -
+	    from eval_var wt_var da_var wf True
+	    have "nrm V \<subseteq>  dom (locals (store s1))"
+	      by (cases rule: da_good_approx_evalnE) rules
+	    with da_e show ?thesis
+	      by (rule da_weakenE) 
+	  qed
+	  note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
+	  show ?thesis
+	    by (rule ve)
+	next
+	  case False
+	  note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
+	  with False show ?thesis
+	    by rules
+	qed
+	with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
+	  by simp
+      next
+	case True
+	then obtain vn where 
+	  vn: "var = LVar vn" 
+	  by auto
+	with da obtain E where
+	    da_e:   "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+	  by cases simp+
+	from da.LVar vn obtain  V where
+	  da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                      \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
+	  by auto
+	from valid_var P valid_A conf_s0 eval_var wt_var da_var
+	obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
+	  by (rule validE) 
+	hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
+	  by simp
+	have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+	proof (cases "normal s1")
+	  case True
+	  obtain E' where
+	    da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                       \<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
+	  proof -
+	    from eval_var
+	    have "dom (locals (store s0)) \<subseteq> dom (locals (store (s1)))"
+	      by (rule dom_locals_evaln_mono_elim)
+	    with da_e show ?thesis
+	      by (rule da_weakenE)
+	  qed
+	  note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
+	  show ?thesis
+	    by (rule ve)
+	next
+	  case False
+	  note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
+	  with False show ?thesis
+	    by rules
+	qed
+	with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
+	  by simp
+      qed
+      moreover
+      from eval wt da conf_s0 wf
+      have "s3\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Cond A P P' Q e0 e1 e2)
+  have valid_e0: "G,A|\<Turnstile>\<Colon>{ {Normal P} e0-\<succ> {P'} }" .
+  have valid_then_else:"\<And> b.  G,A|\<Turnstile>\<Colon>{ {P'\<leftarrow>=b} (if b then e1 else e2)-\<succ> {Q} }"
+    using Cond.hyps by simp
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} e0 ? e1 : e2-\<succ> {Q} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC T E v s2 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e0 ? e1 : e2\<Colon>-T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e0 ? e1:e2\<rangle>\<^sub>e\<guillemotright>E"
+    assume eval: "G\<turnstile>s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<midarrow>n\<rightarrow> s2"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain T1 T2 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" 
+	by cases simp
+      from da obtain E0 E1 E2 where
+        da_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e0\<rangle>\<^sub>e\<guillemotright> E0" and
+        da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                 \<turnstile>(dom (locals (store s0)) \<union> assigns_if True e0)\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
+        da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                 \<turnstile>(dom (locals (store s0)) \<union> assigns_if False e0)\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> E2"
+	by cases simp+
+      from eval obtain b s1 where
+	eval_e0: "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1" and
+        eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<midarrow>n\<rightarrow> s2"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_e0 P valid_A conf_s0 eval_e0 wt_e0 da_e0
+      obtain "P' \<lfloor>b\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
+	by (rule validE)
+      hence P': "\<And> v. (P'\<leftarrow>=(the_Bool b)) v s1 Z"
+	by (cases "normal s1") auto
+      have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+      proof (cases "normal s1")
+	case True
+	note normal_s1=this
+	from wt_e1 wt_e2 obtain T' where
+	  wt_then_else: 
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then e1 else e2)\<Colon>-T'"
+	  by (cases "the_Bool b") simp+
+	have s0_s1: "dom (locals (store s0)) 
+                      \<union> assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
+	proof -
+	  from eval_e0 
+	  have eval_e0': "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<rightarrow> s1"
+	    by (rule evaln_eval)
+	  hence
+	    "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+	    by (rule dom_locals_eval_mono_elim)
+          moreover
+	  from eval_e0' True wt_e0 
+	  have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
+	    by (rule assigns_if_good_approx') 
+	  ultimately show ?thesis by (rule Un_least)
+	qed
+	obtain E' where
+	  da_then_else:
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+              \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>if the_Bool b then e1 else e2\<rangle>\<^sub>e\<guillemotright> E'"
+	proof (cases "the_Bool b")
+	  case True
+	  with that da_e1 s0_s1 show ?thesis
+	    by simp (erule da_weakenE,auto)
+	next
+	  case False
+	  with that da_e2 s0_s1 show ?thesis
+	    by simp (erule da_weakenE,auto)
+	qed
+	with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
+	show ?thesis
+	  by (rule validE)
+      next
+	case False
+	with valid_then_else P' valid_A conf_s1 eval_then_else
+	show ?thesis
+	  by (cases rule: validE) rules+
+      qed
+      moreover
+      from eval wt da conf_s0 wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Call A P Q R S accC' args e mn mode pTs' statT)
+  have valid_e: "G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q} }" .
+  have valid_args: "\<And> a. G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>In1 a} args\<doteq>\<succ> {R a} }"
+    using Call.hyps by simp
+  have valid_methd: "\<And> a vs invC declC l.
+        G,A|\<Turnstile>\<Colon>{ {R a\<leftarrow>In3 vs \<and>.
+                 (\<lambda>s. declC =
+                    invocation_declclass G mode (store s) a statT
+                     \<lparr>name = mn, parTs = pTs'\<rparr> \<and>
+                    invC = invocation_class mode (store s) a statT \<and>
+                    l = locals (store s)) ;.
+                 init_lvars G declC \<lparr>name = mn, parTs = pTs'\<rparr> mode a vs \<and>.
+                 (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)}
+            Methd declC \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ> {set_lvars l .; S} }"
+    using Call.hyps by simp
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} {accC',statT,mode}e\<cdot>mn( {pTs'}args)-\<succ> {S} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC T E v s5 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>{accC',statT,mode}e\<cdot>mn( {pTs'}args)\<Colon>-T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))
+                   \<guillemotright>\<langle>{accC',statT,mode}e\<cdot>mn( {pTs'}args)\<rangle>\<^sub>e\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>{accC',statT,mode}e\<cdot>mn( {pTs'}args)-\<succ>v\<midarrow>n\<rightarrow> s5"
+    assume P: "(Normal P) Y s0 Z"
+    show "S \<lfloor>v\<rfloor>\<^sub>e s5 Z \<and> s5\<Colon>\<preceq>(G, L)"
+    proof -
+      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 =(resTy statM)" and
+        eq_accC_accC': "accC=accC'"
+	by cases fastsimp+
+      from da obtain C where
+	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
+	da_args: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E" 
+	by cases simp
+      from eval eq_accC_accC' obtain a s1 vs s2 s3 s3' s4 invDeclC where
+	evaln_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
+        evaln_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
+	invDeclC: "invDeclC = invocation_declclass 
+                G mode (store s2) a statT \<lparr>name=mn,parTs=pTs'\<rparr>" and
+        s3: "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" and
+	evaln_methd:
+           "G\<turnstile>s3' \<midarrow>Methd invDeclC  \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<midarrow>n\<rightarrow> s4" and
+        s5: "s5=(set_lvars (locals (store s2))) s4"
+	using normal_s0 by (auto elim: evaln_elim_cases)
+
+      from evaln_e
+      have eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+	by (rule evaln_eval)
+      
+      from eval_e _ wt_e wf
+      have s1_no_return: "abrupt s1 \<noteq> Some (Jump Ret)"
+	by (rule eval_expression_no_jump 
+                 [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
+	   (insert normal_s0,auto)
+
+      from valid_e P valid_A conf_s0 evaln_e wt_e da_e
+      obtain "Q \<lfloor>a\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
+	by (rule validE)
+      hence Q: "\<And> v. (Q\<leftarrow>In1 a) v s1 Z"
+	by simp
+      obtain 
+	R: "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" and 
+	conf_s2: "s2\<Colon>\<preceq>(G,L)" and 
+	s2_no_return: "abrupt s2 \<noteq> Some (Jump Ret)"
+      proof (cases "normal s1")
+	case True
+	obtain E' where 
+	  da_args':
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
+	proof -
+	  from evaln_e wt_e da_e wf True
+	  have "nrm C \<subseteq>  dom (locals (store s1))"
+	    by (cases rule: da_good_approx_evalnE) rules
+	  with da_args show ?thesis
+	    by (rule da_weakenE) 
+	qed
+	with valid_args Q valid_A conf_s1 evaln_args wt_args 
+	obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)" 
+	  by (rule validE)
+	moreover
+	from evaln_args
+	have e: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
+	  by (rule evaln_eval)
+	from this s1_no_return wt_args wf
+	have "abrupt s2 \<noteq> Some (Jump Ret)"
+	  by (rule eval_expression_list_no_jump 
+                 [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
+	ultimately show ?thesis ..
+      next
+	case False
+	with valid_args Q valid_A conf_s1 evaln_args
+	obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)" 
+	  by (cases rule: validE) rules+
+	moreover
+	from False evaln_args have "s2=s1"
+	  by auto
+	with s1_no_return have "abrupt s2 \<noteq> Some (Jump Ret)"
+	  by simp
+	ultimately show ?thesis ..
+      qed
+
+      obtain invC where
+	invC: "invC = invocation_class mode (store s2) a statT"
+	by simp
+      with s3
+      have invC': "invC = (invocation_class mode (store s3) a statT)"
+	by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
+      obtain l where
+	l: "l = locals (store s2)"
+	by simp
+
+      from eval wt da conf_s0 wf
+      have conf_s5: "s5\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      let "PROP ?R" = "\<And> v.
+             (R a\<leftarrow>In3 vs \<and>.
+                 (\<lambda>s. invDeclC = invocation_declclass G mode (store s) a statT
+                                  \<lparr>name = mn, parTs = pTs'\<rparr> \<and>
+                       invC = invocation_class mode (store s) a statT \<and>
+                          l = locals (store s)) ;.
+                  init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> mode a vs \<and>.
+                  (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)
+               ) v s3' Z"
+      {
+	assume abrupt_s3: "\<not> normal s3"
+	have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
+	proof -
+	  from abrupt_s3 check have eq_s3'_s3: "s3'=s3"
+	    by (auto simp add: check_method_access_def Let_def)
+	  with R s3 invDeclC invC l abrupt_s3
+	  have R': "PROP ?R"
+	    by auto
+	  have conf_s3': "s3'\<Colon>\<preceq>(G, empty)"
+	   (* we need an arbirary environment (here empty) that s2' conforms to
+              to apply validE *)
+	  proof -
+	    from s2_no_return s3
+	    have "abrupt s3 \<noteq> Some (Jump Ret)"
+	      by (cases s2) (auto simp add: init_lvars_def2 split: split_if_asm)
+	    moreover
+	    obtain abr2 str2 where s2: "s2=(abr2,str2)"
+	      by (cases s2) simp
+	    from s3 s2 conf_s2 have "(abrupt s3,str2)\<Colon>\<preceq>(G, L)"
+	      by (auto simp add: init_lvars_def2 split: split_if_asm)
+	    ultimately show ?thesis
+	      using s3 s2 eq_s3'_s3
+	      apply (simp add: init_lvars_def2)
+	      apply (rule conforms_set_locals [OF _ wlconf_empty])
+	      by auto
+	  qed
+	  from valid_methd R' valid_A conf_s3' evaln_methd abrupt_s3 eq_s3'_s3
+	  have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
+	    by (cases rule: validE) simp+
+	  with s5 l show ?thesis
+	    by simp
+	qed
+      } note abrupt_s3_lemma = this
+
+      have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
+      proof (cases "normal s2")
+	case False
+	with s3 have abrupt_s3: "\<not> normal s3"
+	  by (cases s2) (simp add: init_lvars_def2)
+	thus ?thesis
+	  by (rule abrupt_s3_lemma)
+      next
+	case True
+	note normal_s2 = this
+	with evaln_args 
+	have normal_s1: "normal s1"
+	  by (rule evaln_no_abrupt)
+	obtain E' where 
+	  da_args':
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
+	proof -
+	  from evaln_e wt_e da_e wf normal_s1
+	  have "nrm C \<subseteq>  dom (locals (store s1))"
+	    by (cases rule: da_good_approx_evalnE) rules
+	  with da_args show ?thesis
+	    by (rule da_weakenE) 
+	qed
+	from evaln_args
+	have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
+	  by (rule evaln_eval)
+	from evaln_e wt_e da_e conf_s0 wf
+	have conf_a: "G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+	  by (rule evaln_type_sound [elim_format]) (insert normal_s1,simp)
+	with normal_s1 normal_s2 eval_args 
+	have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
+	  by (auto dest: eval_gext intro: conf_gext)
+	from evaln_args wt_args da_args' conf_s1 wf
+	have conf_args: "list_all2 (conf G (store s2)) vs pTs"
+	  by (rule evaln_type_sound [elim_format]) (insert normal_s2,simp)
+	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)
+	show ?thesis
+	proof (cases "normal s3")
+	  case False
+	  thus ?thesis
+	    by (rule abrupt_s3_lemma)
+	next
+	  case True
+	  note normal_s3 = this
+	  with s3 have notNull: "mode = IntVir \<longrightarrow> a \<noteq> Null"
+	    by (cases s2) (auto simp add: init_lvars_def2)
+	  from conf_s2 conf_a_s2 wf notNull 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)
+	  
+	  with dynT_prop R s3 invDeclC invC l 
+	  have R': "PROP ?R"
+	    by auto
 
+	  from dynT_prop wf wt_e statM' mode invC invDeclC dynM
+	  obtain 
+            dynM: "dynlookup G statT invC  \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+	    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 (cases rule: DynT_mheadsE) simp
+	  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)
+	    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 (simp add: normal_s2)
+	    apply (cases s2, simp add: L' init_lvars_def2 s3
+	                     cong add: lname.case_cong ename.case_cong)
+	    done
+	  with eq_s3'_s3 have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
+	  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,
+                  auto simp add: callee_lcl_def  
+                       cong add: lname.case_cong ename.case_cong)
+	  with dynM' iscls_invDeclC invDeclC'
+	  have
+	    wt_methd:
+	    "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+               \<turnstile>(Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<Colon>-mthdT"
+	    by (auto intro: wt.Methd)
+	  obtain M where 
+	    da_methd:
+	    "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr> 
+	       \<turnstile> dom (locals (store s3')) 
+                   \<guillemotright>\<langle>Methd invDeclC \<lparr>name=mn,parTs=pTs'\<rparr>\<rangle>\<^sub>e\<guillemotright> M"
+	  proof -
+	    from wf_dynM
+	    obtain M' where
+	      da_body: 
+	      "\<lparr>prg=G, cls=invDeclC
+               ,lcl=callee_lcl invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> (mthd dynM)
+               \<rparr> \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'" and
+              res: "Result \<in> nrm M'"
+	      by (rule wf_mdeclE) rules
+	    from da_body is_static_eq L' have
+	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
+                 \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'"
+	      by (simp add: callee_lcl_def  
+                  cong add: lname.case_cong ename.case_cong)
+	    moreover have "parameters (mthd dynM) \<subseteq>  dom (locals (store s3'))"
+	    proof -
+	      from is_static_eq 
+	      have "(invmode (mthd dynM) e) = (invmode statM e)"
+		by (simp add: invmode_def)
+	      with s3 dynM' is_static_eq normal_s2 mode 
+	      have "parameters (mthd dynM) = dom (locals (store s3))"
+		using dom_locals_init_lvars 
+                  [of "mthd dynM" G invDeclC "\<lparr>name=mn,parTs=pTs'\<rparr>" e a vs s2]
+		by simp
+	      thus ?thesis using eq_s3'_s3 by simp
+	    qed
+	    ultimately obtain M2 where
+	      da:
+	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
+                \<turnstile> dom (locals (store s3')) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M2" and
+              M2: "nrm M' \<subseteq> nrm M2"
+	      by (rule da_weakenE)
+	    from res M2 have "Result \<in> nrm M2"
+	      by blast
+	    moreover from wf_dynM
+	    have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
+	      by (rule wf_mdeclE)
+	    ultimately
+	    obtain M3 where
+	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3')) 
+                     \<guillemotright>\<langle>Body (declclass dynM) (stmt (mbody (mthd dynM)))\<rangle>\<guillemotright> M3"
+	      using da
+	      by (rules intro: da.Body assigned.select_convs)
+	    from _ this [simplified]
+	    show ?thesis
+	      by (rule da.Methd [simplified,elim_format])
+	         (auto intro: dynM')
+	  qed
+	  from valid_methd R' valid_A conf_s3' evaln_methd wt_methd da_methd
+	  have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
+	    by (cases rule: validE) rules+
+	  with s5 l show ?thesis
+	    by simp
+	qed
+      qed
+      with conf_s5 show ?thesis by rules
+    qed
+  qed
+next
+-- {* 
+\par
+*} (* dummy text command to break paragraph for latex;
+              large paragraphs exhaust memory of debian pdflatex *)
+  case (Methd A P Q ms)
+  have valid_body: "G,A \<union> {{P} Methd-\<succ> {Q} | ms}|\<Turnstile>\<Colon>{{P} body G-\<succ> {Q} | ms}".
+  show "G,A|\<Turnstile>\<Colon>{{P} Methd-\<succ> {Q} | ms}"
+    by (rule Methd_sound)
+next
+  case (Body A D P Q R c)
+  have valid_init: "G,A|\<Turnstile>\<Colon>{ {Normal P} .Init D. {Q} }".
+  have valid_c: "G,A|\<Turnstile>\<Colon>{ {Q} .c. 
+              {\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>In1 (the (locals s Result))} }".
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} Body D c-\<succ> {R} }"
+  proof (rule valid_expr_NormalI)
+    fix n s0 L accC T E v s4 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Body D c\<Colon>-T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>Body D c\<rangle>\<^sub>e\<guillemotright>E"
+    assume eval: "G\<turnstile>s0 \<midarrow>Body D c-\<succ>v\<midarrow>n\<rightarrow> s4"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<lfloor>v\<rfloor>\<^sub>e s4 Z \<and> s4\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain 
+	iscls_D: "is_class G D" and
+        wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Init D\<Colon>\<surd>" and
+        wt_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>" 
+	by cases auto
+      obtain I where 
+	da_init:"\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init D\<rangle>\<^sub>s\<guillemotright> I"
+	by (auto intro: da_Init [simplified] assigned.select_convs)
+      from da obtain C where
+	da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C" and
+	jmpOk: "jumpNestingOkS {Ret} c" 
+	by cases simp
+      from eval obtain s1 s2 s3 where
+	eval_init: "G\<turnstile>s0 \<midarrow>Init D\<midarrow>n\<rightarrow> s1" and
+        eval_c: "G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2" and
+	v: "v = the (locals (store s2) Result)" and
+        s3: "s3 =(if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or> 
+                         abrupt s2 = Some (Jump (Cont l))
+                  then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)"and
+        s4: "s4 = abupd (absorb Ret) s3"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      obtain C' where 
+	da_c': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C'"
+      proof -
+	from eval_init 
+	have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
+	  by (rule dom_locals_evaln_mono_elim)
+	with da_c show ?thesis by (rule da_weakenE)
+      qed
+      from valid_init P valid_A conf_s0 eval_init wt_init da_init
+      obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
+	by (rule validE)
+      from valid_c Q valid_A conf_s1 eval_c wt_c da_c' 
+      have R: "(\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>In1 (the (locals s Result))) 
+                \<diamondsuit> s2 Z"
+	by (rule validE)
+      have "s3=s2"
+      proof -
+	from eval_init [THEN evaln_eval] wf
+	have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+	  by - (rule eval_statement_no_jump [OF _ _ _ wt_init],
+                insert normal_s0,auto)
+	from eval_c [THEN evaln_eval] _ wt_c wf
+	have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
+	  by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
+	moreover note s3
+	ultimately show ?thesis 
+	  by (force split: split_if)
+      qed
+      with R v s4 
+      have "R \<lfloor>v\<rfloor>\<^sub>e s4 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s4\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Nil A P)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal (P\<leftarrow>\<lfloor>[]\<rfloor>\<^sub>l)} []\<doteq>\<succ> {P} }"
+  proof (rule valid_expr_list_NormalI)
+    fix s0 s1 vs n L Y Z
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume eval: "G\<turnstile>s0 \<midarrow>[]\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s1"
+    assume P: "(Normal (P\<leftarrow>\<lfloor>[]\<rfloor>\<^sub>l)) Y s0 Z"
+    show "P \<lfloor>vs\<rfloor>\<^sub>l s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof -
+      from eval obtain "vs=[]" "s1=s0"
+	using normal_s0 by (auto elim: evaln_elim_cases)
+      with P conf_s0 show ?thesis
+	by simp
+    qed
+  qed
+next
+  case (Cons A P Q R e es)
+  have valid_e: "G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q} }".
+  have valid_es: "\<And> v. G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>\<lfloor>v\<rfloor>\<^sub>e} es\<doteq>\<succ> {\<lambda>Vals:vs:. R\<leftarrow>\<lfloor>(v # vs)\<rfloor>\<^sub>l} }"
+    using Cons.hyps by simp
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} e # es\<doteq>\<succ> {R} }"
+  proof (rule valid_expr_list_NormalI)
+    fix n s0 L accC T E v s2 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e # es\<Colon>\<doteq>T"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>e # es\<rangle>\<^sub>l\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>e # es\<doteq>\<succ>v\<midarrow>n\<rightarrow> s2"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<lfloor>v\<rfloor>\<^sub>l s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt 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"
+	by cases simp
+      from da obtain E1 where
+	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E1" and
+	da_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E" 
+	by cases simp
+      from eval obtain s1 ve vs where
+	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
+	eval_es: "G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
+	v: "v=ve#vs"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_e P valid_A conf_s0 eval_e wt_e da_e 
+      obtain Q: "Q \<lfloor>ve\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
+	by (rule validE)
+      from Q have Q': "\<And> v. (Q\<leftarrow>\<lfloor>ve\<rfloor>\<^sub>e) v s1 Z"
+	by simp
+      have "(\<lambda>Vals:vs:. R\<leftarrow>\<lfloor>(ve # vs)\<rfloor>\<^sub>l) \<lfloor>vs\<rfloor>\<^sub>l s2 Z"
+      proof (cases "normal s1")
+	case True
+	obtain E' where 
+	  da_es': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E'"
+	proof -
+	  from eval_e wt_e da_e wf True
+	  have "nrm E1 \<subseteq> dom (locals (store s1))"
+	    by (cases rule: da_good_approx_evalnE) rules
+	  with da_es show ?thesis
+	    by (rule da_weakenE)
+	qed
+	from valid_es Q' valid_A conf_s1 eval_es wt_es da_es'
+	show ?thesis
+	  by (rule validE)
+      next
+	case False
+	with valid_es Q' valid_A conf_s1 eval_es 
+	show ?thesis
+	  by (cases rule: validE) rules+
+      qed
+      with v have "R \<lfloor>v\<rfloor>\<^sub>l s2 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Skip A P)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal (P\<leftarrow>\<diamondsuit>)} .Skip. {P} }"
+  proof (rule valid_stmt_NormalI)
+    fix s0 s1 n L Y Z
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume eval: "G\<turnstile>s0 \<midarrow>Skip\<midarrow>n\<rightarrow> s1"
+    assume P: "(Normal (P\<leftarrow>\<diamondsuit>)) Y s0 Z"
+    show "P \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof -
+      from eval obtain "s1=s0"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      with P conf_s0 show ?thesis
+	by simp
+    qed
+  qed
+next
+  case (Expr A P Q e)
+  have valid_e: "G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q\<leftarrow>\<diamondsuit>} }".
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} .Expr e. {Q} }"
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC C s1 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Expr e\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>Expr e\<rangle>\<^sub>s\<guillemotright> C"
+    assume eval: "G\<turnstile>s0 \<midarrow>Expr e\<midarrow>n\<rightarrow> s1"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain eT where 
+	wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
+	by cases simp
+      from da obtain E where
+	da_e: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright>E"
+	by cases simp
+      from eval obtain v where
+	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_e P valid_A conf_s0 eval_e wt_e da_e
+      obtain Q: "(Q\<leftarrow>\<diamondsuit>) \<lfloor>v\<rfloor>\<^sub>e s1 Z" and "s1\<Colon>\<preceq>(G,L)"
+	by (rule validE)
+      thus ?thesis by simp
+    qed
+  qed
+next
+  case (Lab A P Q c l)
+  have valid_c: "G,A|\<Turnstile>\<Colon>{ {Normal P} .c. {abupd (absorb l) .; Q} }".
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} .l\<bullet> c. {Q} }"
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC C s2 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>l\<bullet> c\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>l\<bullet> c\<rangle>\<^sub>s\<guillemotright> C"
+    assume eval: "G\<turnstile>s0 \<midarrow>l\<bullet> c\<midarrow>n\<rightarrow> s2"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
+    proof -
+      from wt obtain 
+	wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
+	by cases simp
+      from da obtain E where
+	da_c: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>E"
+	by cases simp
+      from eval obtain s1 where
+	eval_c: "G\<turnstile>s0 \<midarrow>c\<midarrow>n\<rightarrow> s1" and
+	s2: "s2 = abupd (absorb l) s1"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from valid_c P valid_A conf_s0 eval_c wt_c da_c
+      obtain Q: "(abupd (absorb l) .; Q) \<diamondsuit> s1 Z" 
+	by (rule validE)
+      with s2 have "Q \<diamondsuit> s2 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Comp A P Q R c1 c2)
+  have valid_c1: "G,A|\<Turnstile>\<Colon>{ {Normal P} .c1. {Q} }" .
+  have valid_c2: "G,A|\<Turnstile>\<Colon>{ {Q} .c2. {R} }" .
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} .c1;; c2. {R} }"
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC C s2 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(c1;; c2)\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c1;;c2\<rangle>\<^sub>s\<guillemotright>C"
+    assume eval: "G\<turnstile>s0 \<midarrow>c1;; c2\<midarrow>n\<rightarrow> s2"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
+    proof -
+      from eval  obtain s1 where
+	eval_c1: "G\<turnstile>s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1" and
+	eval_c2: "G\<turnstile>s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from wt 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 cases simp
+      from da obtain C1 C2 where 
+	da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and 
+	da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>nrm C1 \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2" 
+	by cases simp
+      from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1  
+      obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
+	by (rule validE) 
+      have "R \<diamondsuit> s2 Z"
+      proof (cases "normal s1")
+	case True
+	obtain C2' where 
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
+	proof -
+	  from eval_c1 wt_c1 da_c1 wf True
+	  have "nrm C1 \<subseteq> dom (locals (store s1))"
+	    by (cases rule: da_good_approx_evalnE) rules
+	  with da_c2 show ?thesis
+	    by (rule da_weakenE)
+	qed
+	with valid_c2 Q valid_A conf_s1 eval_c2 wt_c2 
+	show ?thesis
+	  by (rule validE)
+      next
+	case False
+	from valid_c2 Q valid_A conf_s1 eval_c2 False
+	show ?thesis
+	  by (cases rule: validE) rules+
+      qed
+      moreover
+      from eval wt da conf_s0 wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (If A P P' Q c1 c2 e)
+  have valid_e: "G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {P'} }" .
+  have valid_then_else: "\<And> b. G,A|\<Turnstile>\<Colon>{ {P'\<leftarrow>=b} .(if b then c1 else c2). {Q} }" 
+    using If.hyps by simp
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} .If(e) c1 Else c2. {Q} }"
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC C s2 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>If(e) c1 Else c2\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                    \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>If(e) c1 Else c2\<rangle>\<^sub>s\<guillemotright>C"
+    assume eval: "G\<turnstile>s0 \<midarrow>If(e) c1 Else c2\<midarrow>n\<rightarrow> s2"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
+    proof -
+      from eval obtain b s1 where
+	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1" and
+	eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2"
+	using normal_s0 by (auto elim: evaln_elim_cases)
+      from wt 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 cases (simp split: split_if)
+      from da obtain E S where
+	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
+	da_then_else: 
+	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+             (dom (locals (store s0)) \<union> assigns_if (the_Bool b) e)
+               \<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> S"
+	by cases (cases "the_Bool b",auto)
+      from valid_e P valid_A conf_s0 eval_e wt_e da_e
+      obtain "P' \<lfloor>b\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
+	by (rule validE)
+      hence P': "\<And>v. (P'\<leftarrow>=the_Bool b) v s1 Z"
+	by (cases "normal s1") auto
+      have "Q \<diamondsuit> s2 Z"
+      proof (cases "normal s1")
+	case True
+	have s0_s1: "dom (locals (store s0)) 
+                      \<union> assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+	proof -
+	  from eval_e 
+	  have eval_e': "G\<turnstile>s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
+	    by (rule evaln_eval)
+	  hence
+	    "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+	    by (rule dom_locals_eval_mono_elim)
+          moreover
+	  from eval_e' True wt_e
+	  have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+	    by (rule assigns_if_good_approx') 
+	  ultimately show ?thesis by (rule Un_least)
+	qed
+	with da_then_else
+	obtain S' where
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+              \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> S'"
+	  by (rule da_weakenE)
+	with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
+	show ?thesis
+	  by (rule validE)
+      next
+	case False
+	with valid_then_else P' valid_A conf_s1 eval_then_else
+	show ?thesis
+	  by (cases rule: validE) rules+
+      qed
+      moreover
+      from eval wt da conf_s0 wf
+      have "s2\<Colon>\<preceq>(G, L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Loop A P P' c e l)
+  have valid_e: "G,A|\<Turnstile>\<Colon>{ {P} e-\<succ> {P'} }".
+  have valid_c: "G,A|\<Turnstile>\<Colon>{ {Normal (P'\<leftarrow>=True)} 
+                         .c. 
+                         {abupd (absorb (Cont l)) .; P} }" .
+  show "G,A|\<Turnstile>\<Colon>{ {P} .l\<bullet> While(e) c. {P'\<leftarrow>=False\<down>=\<diamondsuit>} }"
+  proof (rule valid_stmtI)
+    fix n s0 L accC C s3 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume wt: "normal s0 \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>l\<bullet> While(e) c\<Colon>\<surd>"
+    assume da: "normal s0 \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                    \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<guillemotright> C"
+    assume eval: "G\<turnstile>s0 \<midarrow>l\<bullet> While(e) c\<midarrow>n\<rightarrow> s3"
+    assume P: "P Y s0 Z"
+    show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
+    proof -
+        --{* From the given hypothesises @{text valid_e} and @{text valid_c} 
+           we can only reach the state after unfolding the loop once, i.e. 
+	   @{term "P \<diamondsuit> s2 Z"}, where @{term s2} is the state after executing
+           @{term c}. To gain validity of the further execution of while, to
+           finally get @{term "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3 Z"} we have to get 
+           a hypothesis about the subsequent unfoldings (the whole loop again),
+           too. We can achieve this, by performing induction on the 
+           evaluation relation, with all
+           the necessary preconditions to apply @{text valid_e} and 
+           @{text valid_c} in the goal.
+        *}
+      {
+	fix t s s' v 
+	assume "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
+	hence "\<And> Y' T E. 
+                \<lbrakk>t =  \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s; \<forall>t\<in>A. G\<Turnstile>n\<Colon>t; P Y' s Z; s\<Colon>\<preceq>(G, L);
+                 normal s \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T; 
+                 normal s \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>E
+                \<rbrakk>\<Longrightarrow> (P'\<leftarrow>=False\<down>=\<diamondsuit>) v s' Z"
+	  (is "PROP ?Hyp n t s v s'")
+	proof (induct)
+	  case (Loop b c' e' l' n' s0' s1' s2' s3' Y' T E)
+	  have while: "(\<langle>l'\<bullet> While(e') c'\<rangle>\<^sub>s::term) = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s" .
+          hence eqs: "l'=l" "e'=e" "c'=c" by simp_all
+	  have valid_A: "\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t". 
+	  have P: "P Y' (Norm s0') Z".
+	  have conf_s0': "Norm s0'\<Colon>\<preceq>(G, L)" .
+          have wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<Colon>T"
+	    using Loop.prems eqs by simp
+	  have da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
+                    dom (locals (store ((Norm s0')::state)))\<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<guillemotright>E"
+	    using Loop.prems eqs by simp
+	  have evaln_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<midarrow>n'\<rightarrow> s1'" 
+	    using Loop.hyps eqs by simp
+	  show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
+	  proof -
+	    from wt  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 cases (simp add: eqs)
+	    from da obtain E S where
+	      da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                     \<turnstile> dom (locals (store ((Norm s0')::state))) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
+	      da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                     \<turnstile> (dom (locals (store ((Norm s0')::state))) 
+                            \<union> assigns_if True e) \<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> S"
+	      by cases (simp add: eqs)
+	    from evaln_e 
+	    have eval_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<rightarrow> s1'"
+	      by (rule evaln_eval)
+	    from valid_e P valid_A conf_s0' evaln_e wt_e da_e
+	    obtain P': "P' \<lfloor>b\<rfloor>\<^sub>e s1' Z" and conf_s1': "s1'\<Colon>\<preceq>(G,L)"
+	      by (rule validE)
+	    show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
+	    proof (cases "normal s1'")
+	      case True
+	      note normal_s1'=this
+	      show ?thesis
+	      proof (cases "the_Bool b")
+		case True
+		with P' normal_s1' have P'': "(Normal (P'\<leftarrow>=True)) \<lfloor>b\<rfloor>\<^sub>e s1' Z"
+		  by auto
+		from True Loop.hyps obtain
+		  eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and 
+		  eval_while:  
+		     "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
+		  by (simp add: eqs)
+		from True Loop.hyps have
+		  hyp: "PROP ?Hyp n' \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s 
+                          (abupd (absorb (Cont l')) s2') \<diamondsuit> s3'"
+		  apply (simp only: True if_True eqs)
+		  apply (elim conjE)
+		  apply (tactic "smp_tac 3 1")
+		  apply fast
+		  done
+		from eval_e
+		have s0'_s1': "dom (locals (store ((Norm s0')::state))) 
+                                  \<subseteq> dom (locals (store s1'))"
+		  by (rule dom_locals_eval_mono_elim)
+		obtain S' where
+		  da_c':
+		   "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(dom (locals (store s1')))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> S'" 
+		proof -
+		  note s0'_s1'
+		  moreover
+		  from eval_e normal_s1' wt_e 
+		  have "assigns_if True e \<subseteq> dom (locals (store s1'))"
+		    by (rule assigns_if_good_approx' [elim_format]) 
+                       (simp add: True)
+		  ultimately 
+		  have "dom (locals (store ((Norm s0')::state)))
+                           \<union> assigns_if True e \<subseteq> dom (locals (store s1'))"
+		    by (rule Un_least)
+		  with da_c show ?thesis
+		    by (rule da_weakenE)
+		qed
+		with valid_c P'' valid_A conf_s1' eval_c wt_c
+		obtain "(abupd (absorb (Cont l)) .; P) \<diamondsuit> s2' Z" and 
+                  conf_s2': "s2'\<Colon>\<preceq>(G,L)"
+		  by (rule validE)
+		hence P_s2': "P \<diamondsuit> (abupd (absorb (Cont l)) s2') Z"
+		  by simp
+		from conf_s2'
+		have conf_absorb: "abupd (absorb (Cont l)) s2' \<Colon>\<preceq>(G, L)"
+		  by (cases s2') (auto intro: conforms_absorb)
+		moreover
+		obtain E' where 
+		  da_while':
+		   "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+		     dom (locals(store (abupd (absorb (Cont l)) s2')))
+                      \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<guillemotright> E'"
+		proof -
+		  note s0'_s1'
+		  also 
+		  from eval_c 
+		  have "G\<turnstile>s1' \<midarrow>c\<rightarrow> s2'"
+		    by (rule evaln_eval)
+		  hence "dom (locals (store s1')) \<subseteq> dom (locals (store s2'))"
+		    by (rule dom_locals_eval_mono_elim)
+		  also 
+		  have "\<dots>\<subseteq>dom (locals (store (abupd (absorb (Cont l)) s2')))"
+		    by simp
+		  finally
+		  have "dom (locals (store ((Norm s0')::state))) \<subseteq> \<dots>" .
+		  with da show ?thesis
+		    by (rule da_weakenE)
+		qed
+		from valid_A P_s2' conf_absorb wt da_while'
+		show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z" 
+		  using hyp by (simp add: eqs)
+	      next
+		case False
+		with Loop.hyps obtain "s3'=s1'"
+		  by simp
+		with P' False show ?thesis
+		  by auto
+	      qed 
+	    next
+	      case False
+	      note abnormal_s1'=this
+	      have "s3'=s1'"
+	      proof -
+		from False obtain abr where abr: "abrupt s1' = Some abr"
+		  by (cases s1') auto
+		from eval_e _ wt_e wf
+		have no_jmp: "\<And> j. abrupt s1' \<noteq> Some (Jump j)"
+		  by (rule eval_expression_no_jump 
+                       [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
+		     simp
+		show ?thesis
+		proof (cases "the_Bool b")
+		  case True  
+		  with Loop.hyps obtain
+		    eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and 
+		    eval_while:  
+		     "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
+		    by (simp add: eqs)
+		  from eval_c abr have "s2'=s1'" by auto
+		  moreover from calculation no_jmp 
+		  have "abupd (absorb (Cont l)) s2'=s2'"
+		    by (cases s1') (simp add: absorb_def)
+		  ultimately show ?thesis
+		    using eval_while abr
+		    by auto
+		next
+		  case False
+		  with Loop.hyps show ?thesis by simp
+		qed
+	      qed
+	      with P' False show ?thesis
+		by auto
+	    qed
+	  qed
+	next
+	  case (Abrupt n' s t' abr Y' T E)
+	  have t': "t' = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s".
+	  have conf: "(Some abr, s)\<Colon>\<preceq>(G, L)".
+	  have P: "P Y' (Some abr, s) Z".
+	  have valid_A: "\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t". 
+	  show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) (arbitrary3 t') (Some abr, s) Z"
+	  proof -
+	    have eval_e: 
+	      "G\<turnstile>(Some abr,s) \<midarrow>\<langle>e\<rangle>\<^sub>e\<succ>\<midarrow>n'\<rightarrow> (arbitrary3 \<langle>e\<rangle>\<^sub>e,(Some abr,s))"
+	      by auto
+	    from valid_e P valid_A conf eval_e 
+	    have "P' (arbitrary3 \<langle>e\<rangle>\<^sub>e) (Some abr,s) Z"
+	      by (cases rule: validE [where ?P="P"]) simp+
+	    with t' show ?thesis
+	      by auto
+	  qed
+	qed (simp_all)
+      } note generalized=this
+      from eval _ valid_A P conf_s0 wt da
+      have "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3 Z"
+	by (rule generalized)  simp_all
+      moreover
+      have "s3\<Colon>\<preceq>(G, L)"
+      proof (cases "normal s0")
+	case True
+	from eval wt [OF True] da [OF True] conf_s0 wf
+	show ?thesis
+	  by (rule evaln_type_sound [elim_format]) simp
+      next
+	case False
+	with eval have "s3=s0"
+	  by auto
+	with conf_s0 show ?thesis 
+	  by simp
+      qed
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Jmp A P j)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal (abupd (\<lambda>a. Some (Jump j)) .; P\<leftarrow>\<diamondsuit>)} .Jmp j. {P} }"
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC C s1 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Jmp j\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                    \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>Jmp j\<rangle>\<^sub>s\<guillemotright>C"
+    assume eval: "G\<turnstile>s0 \<midarrow>Jmp j\<midarrow>n\<rightarrow> s1"
+    assume P: "(Normal (abupd (\<lambda>a. Some (Jump j)) .; P\<leftarrow>\<diamondsuit>)) Y s0 Z"
+    show "P \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G,L)"
+    proof -
+      from eval obtain s where  
+	s: "s0=Norm s" "s1=(Some (Jump j), s)" 
+	using normal_s0 by (auto elim: evaln_elim_cases)
+      with P have "P \<diamondsuit> s1 Z"
+	by simp
+      moreover 
+      from eval wt da conf_s0 wf
+      have "s1\<Colon>\<preceq>(G,L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Throw A P Q e)
+  have valid_e: "G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>} }".
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} .Throw e. {Q} }"
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC C s2 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Throw e\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                    \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>Throw e\<rangle>\<^sub>s\<guillemotright>C"
+    assume eval: "G\<turnstile>s0 \<midarrow>Throw e\<midarrow>n\<rightarrow> s2"
+    assume P: "(Normal P) Y s0 Z"
+    show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
+    proof -
+      from eval obtain s1 a where
+	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
+	s2: "s2 = abupd (throw a) s1"
+	using normal_s0 by (auto elim: evaln_elim_cases)
+      from wt obtain T where
+	wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
+	by cases simp
+      from da obtain E where
+	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+	by cases simp
+      from valid_e P valid_A conf_s0 eval_e wt_e da_e 
+      obtain "(\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>) \<lfloor>a\<rfloor>\<^sub>e s1 Z"
+	by (rule validE)
+      with s2 have "Q \<diamondsuit> s2 Z"
+	by simp
+      moreover 
+      from eval wt da conf_s0 wf
+      have "s2\<Colon>\<preceq>(G,L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Try A C P Q R c1 c2 vn)
+  have valid_c1: "G,A|\<Turnstile>\<Colon>{ {Normal P} .c1. {SXAlloc G Q} }".
+  have valid_c2: "G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. G,s\<turnstile>catch C) ;. new_xcpt_var vn} 
+                           .c2. 
+                          {R} }".
+  have Q_R: "(Q \<and>. (\<lambda>s. \<not> G,s\<turnstile>catch C)) \<Rightarrow> R" .
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} .Try c1 Catch(C vn) c2. {R} }"
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC E s3 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Try c1 Catch(C vn) c2\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                    \<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>Try c1 Catch(C vn) c2\<rangle>\<^sub>s\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>Try c1 Catch(C vn) c2\<midarrow>n\<rightarrow> s3"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
+    proof -
+      from eval obtain s1 s2 where
+	eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1" and
+        sxalloc: "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" and
+        s3: "if G,s2\<turnstile>catch C 
+                then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3 
+                else s3 = s2"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from wt 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(VName vn\<mapsto>Class C)\<rparr>\<turnstile>c2\<Colon>\<surd>"
+	by cases simp
+      from da obtain C1 C2 where
+	da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
+	da_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
+                   \<turnstile> (dom (locals (store s0)) \<union> {VName vn}) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
+	by cases simp
+      from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1
+      obtain sxQ: "(SXAlloc G Q) \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
+	by (rule validE)
+      from sxalloc sxQ
+      have Q: "Q \<diamondsuit> s2 Z"
+	by auto
+      have "R \<diamondsuit> s3 Z"
+      proof (cases "\<exists> x. abrupt s1 = Some (Xcpt x)")
+	case False
+	from sxalloc wf
+	have "s2=s1"
+	  by (rule sxalloc_type_sound [elim_format])
+	     (insert False, auto split: option.splits abrupt.splits )
+	with False 
+	have no_catch: "\<not>  G,s2\<turnstile>catch C"
+	  by (simp add: catch_def)
+	moreover
+	from no_catch s3
+	have "s3=s2"
+	  by simp
+	ultimately show ?thesis
+	  using Q Q_R by simp
+      next
+	case True
+	note exception_s1 = this
+	show ?thesis
+	proof (cases "G,s2\<turnstile>catch C") 
+	  case False
+	  with s3
+	  have "s3=s2"
+	    by simp
+	  with False Q Q_R show ?thesis
+	    by simp
+	next
+	  case True
+	  with s3 have eval_c2: "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3"
+	    by simp
+	  from conf_s1 sxalloc wf 
+	  have conf_s2: "s2\<Colon>\<preceq>(G, L)" 
+	    by (auto dest: sxalloc_type_sound 
+                    split: option.splits abrupt.splits)
+	  from exception_s1 sxalloc wf
+	  obtain a 
+	    where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
+	    by (auto dest!: sxalloc_type_sound 
+                            split: option.splits abrupt.splits)
+	  with True
+	  have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class C"
+	    by (cases s2) simp
+	  with xcpt_s2 conf_s2 wf
+	  have conf_new_xcpt: "new_xcpt_var vn s2 \<Colon>\<preceq>(G, L(VName vn\<mapsto>Class C))"
+	    by (auto dest: Try_lemma)
+	  obtain C2' where
+	    da_c2':
+	    "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
+              \<turnstile> (dom (locals (store (new_xcpt_var vn s2)))) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
+	  proof -
+	    have "(dom (locals (store s0)) \<union> {VName vn}) 
+                    \<subseteq> dom (locals (store (new_xcpt_var vn s2)))"
+            proof -
+	      from eval_c1 
+              have "dom (locals (store s0)) 
+                      \<subseteq> dom (locals (store s1))"
+		by (rule dom_locals_evaln_mono_elim)
+              also
+              from sxalloc
+              have "\<dots> \<subseteq> dom (locals (store s2))"
+		by (rule dom_locals_sxalloc_mono)
+              also 
+              have "\<dots> \<subseteq> dom (locals (store (new_xcpt_var vn s2)))" 
+		by (cases s2) (simp add: new_xcpt_var_def, blast) 
+              also
+              have "{VName vn} \<subseteq> \<dots>"
+		by (cases s2) simp
+              ultimately show ?thesis
+		by (rule Un_least)
+            qed
+	    with da_c2 show ?thesis
+	      by (rule da_weakenE)
+	  qed
+	  from Q eval_c2 True 
+	  have "(Q \<and>. (\<lambda>s. G,s\<turnstile>catch C) ;. new_xcpt_var vn) 
+                   \<diamondsuit> (new_xcpt_var vn s2) Z"
+	    by auto
+	  from valid_c2 this valid_A conf_new_xcpt eval_c2 wt_c2 da_c2'
+	  show "R \<diamondsuit> s3 Z"
+	    by (rule validE)
+	qed
+      qed
+      moreover 
+      from eval wt da conf_s0 wf
+      have "s3\<Colon>\<preceq>(G,L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Fin A P Q R c1 c2)
+  have valid_c1: "G,A|\<Turnstile>\<Colon>{ {Normal P} .c1. {Q} }".
+  have valid_c2: "\<And> abr. G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. abr = fst s) ;. abupd (\<lambda>x. None)} 
+                                  .c2.
+                                  {abupd (abrupt_if (abr \<noteq> None) abr) .; R} }"
+    using Fin.hyps by simp
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} .c1 Finally c2. {R} }"
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC E s3 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1 Finally c2\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                    \<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>c1 Finally c2\<rangle>\<^sub>s\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>c1 Finally c2\<midarrow>n\<rightarrow> s3"
+    assume P: "(Normal P) Y s0 Z"
+    show "R \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
+    proof -
+      from eval obtain s1 abr1 s2 where
+	eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> (abr1, s1)" and
+        eval_c2: "G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n\<rightarrow> s2" and
+        s3: "s3 = (if \<exists>err. abr1 = Some (Error err) 
+                      then (abr1, s1)
+                      else abupd (abrupt_if (abr1 \<noteq> None) abr1) s2)"
+	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+      from wt 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 cases simp
+      from da obtain C1 C2 where
+	da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
+	da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
+	by cases simp
+      from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1
+      obtain Q: "Q \<diamondsuit> (abr1,s1) Z" and conf_s1: "(abr1,s1)\<Colon>\<preceq>(G,L)" 
+	by (rule validE)
+      from Q 
+      have Q': "(Q \<and>. (\<lambda>s. abr1 = fst s) ;. abupd (\<lambda>x. None)) \<diamondsuit> (Norm s1) Z"
+	by auto
+      from eval_c1 wt_c1 da_c1 conf_s0 wf
+      have  "error_free (abr1,s1)"
+	by (rule evaln_type_sound  [elim_format]) (insert normal_s0,simp)
+      with s3 have s3': "s3 = abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
+	by (simp add: error_free_def)
+      from conf_s1 
+      have conf_Norm_s1: "Norm s1\<Colon>\<preceq>(G,L)"
+	by (rule conforms_NormI)
+      obtain C2' where 
+	da_c2': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                   \<turnstile> dom (locals (store ((Norm s1)::state))) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
+      proof -
+	from eval_c1 
+	have "dom (locals (store s0)) \<subseteq> dom (locals (store (abr1,s1)))"
+          by (rule dom_locals_evaln_mono_elim)
+	hence "dom (locals (store s0)) 
+                 \<subseteq> dom (locals (store ((Norm s1)::state)))"
+	  by simp
+	with da_c2 show ?thesis
+	  by (rule da_weakenE)
+      qed
+      from valid_c2 Q' valid_A conf_Norm_s1 eval_c2 wt_c2 da_c2'
+      have "(abupd (abrupt_if (abr1 \<noteq> None) abr1) .; R) \<diamondsuit> s2 Z"
+	by (rule validE)
+      with s3' have "R \<diamondsuit> s3 Z"
+	by simp
+      moreover
+      from eval wt da conf_s0 wf
+      have "s3\<Colon>\<preceq>(G,L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (Done A C P)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal (P\<leftarrow>\<diamondsuit> \<and>. initd C)} .Init C. {P} }" 
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC E s3 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Init C\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                    \<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>Init C\<rangle>\<^sub>s\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s3"
+    assume P: "(Normal (P\<leftarrow>\<diamondsuit> \<and>. initd C)) Y s0 Z"
+    show "P \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
+    proof -
+      from P have inited: "inited C (globs (store s0))"
+	by simp
+      with eval have "s3=s0"
+	using normal_s0 by (auto elim: evaln_elim_cases)
+      with P conf_s0 show ?thesis
+	by simp
+    qed
+  qed
+next
+  case (Init A C P Q R c)
+  have c: "the (class G C) = c".
+  have valid_super: 
+        "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} }".
+  have valid_init: 
+        "\<And> l.  G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty} 
+                        .init c.
+                        {set_lvars l .; R} }"
+    using Init.hyps by simp
+  show "G,A|\<Turnstile>\<Colon>{ {Normal (P \<and>. Not \<circ> initd C)} .Init C. {R} }"
+  proof (rule valid_stmt_NormalI)
+    fix n s0 L accC E s3 Y Z
+    assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
+    assume conf_s0:  "s0\<Colon>\<preceq>(G,L)"  
+    assume normal_s0: "normal s0"
+    assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Init C\<Colon>\<surd>"
+    assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                    \<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>Init C\<rangle>\<^sub>s\<guillemotright> E"
+    assume eval: "G\<turnstile>s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s3"
+    assume P: "(Normal (P \<and>. Not \<circ> initd C)) Y s0 Z"
+    show "R \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
+    proof -
+      from P have not_inited: "\<not> inited C (globs (store s0))" by simp
+      with eval c obtain s1 s2 where
+	eval_super: 
+	"G\<turnstile>Norm ((init_class_obj G C) (store s0)) 
+           \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1" and
+	eval_init: "G\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2" and
+        s3: "s3 = (set_lvars (locals (store s1))) s2"
+	using normal_s0 by (auto elim!: evaln_elim_cases)
+      from wt c have
+	cls_C: "class G C = Some c"
+	by cases auto
+      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)
+      obtain S where
+	da_super:
+	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+          \<turnstile> dom (locals (store ((Norm 
+                            ((init_class_obj G C) (store s0)))::state))) 
+               \<guillemotright>\<langle>if C = Object then Skip else Init (super c)\<rangle>\<^sub>s\<guillemotright> S"
+      proof (cases "C=Object")
+	case True 
+	with da_Skip show ?thesis
+	  using that by (auto intro: assigned.select_convs)
+      next
+	case False 
+	with da_Init show ?thesis
+	  by - (rule that, auto intro: assigned.select_convs)
+      qed
+      from normal_s0 conf_s0 wf cls_C not_inited
+      have conf_init_cls: "(Norm ((init_class_obj G C) (store s0)))\<Colon>\<preceq>(G, L)"
+	by (auto intro: conforms_init_class_obj)	
+      from P 
+      have P': "(Normal (P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C)))
+                   Y (Norm ((init_class_obj G C) (store s0))) Z"
+	by auto
+
+      from valid_super P' valid_A conf_init_cls eval_super wt_super da_super
+      obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
+	by (rule validE)
+      
+      from cls_C wf have wt_init: "\<lparr>prg=G, cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd>"
+	by (rule wf_prog_cdecl [THEN wf_cdecl_wt_init])
+      from cls_C wf obtain I where 
+	"\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<^sub>s\<guillemotright> I"
+	by (rule wf_prog_cdecl [THEN wf_cdeclE,simplified]) blast
+       (*  simplified: to rewrite \<langle>init c\<rangle> to In1r (init c) *) 
+      then obtain I' where
+	da_init:
+	"\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>dom (locals (store ((set_lvars empty) s1))) 
+            \<guillemotright>\<langle>init c\<rangle>\<^sub>s\<guillemotright> I'"
+	by (rule da_weakenE) simp
+      have conf_s1_empty: "(set_lvars empty) s1\<Colon>\<preceq>(G, empty)"
+      proof -
+	from eval_super have
+	  "G\<turnstile>Norm ((init_class_obj G C) (store s0)) 
+             \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1"
+	  by (rule evaln_eval)
+	from this wt_super wf
+	have s1_no_ret: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+	  by - (rule eval_statement_no_jump 
+                 [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>"], auto split: split_if)
+        with conf_s1
+	show ?thesis
+	  by (cases s1) (auto intro: conforms_set_locals)
+      qed
+      
+      obtain l where l: "l = locals (store s1)"
+	by simp
+      with Q 
+      have Q': "(Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty)
+                  \<diamondsuit> ((set_lvars empty) s1) Z"
+	by auto
+      from valid_init Q' valid_A conf_s1_empty eval_init wt_init da_init
+      have "(set_lvars l .; R) \<diamondsuit> s2 Z"
+	by (rule validE)
+      with s3 l have "R \<diamondsuit> s3 Z"
+	by simp
+      moreover 
+      from eval wt da conf_s0 wf
+      have "s3\<Colon>\<preceq>(G,L)"
+	by (rule evaln_type_sound [elim_format]) simp
+      ultimately show ?thesis ..
+    qed
+  qed
+next
+  case (InsInitV A P Q c v)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} InsInitV c v=\<succ> {Q} }"
+  proof (rule valid_var_NormalI)
+    fix s0 vf n s1 L Z
+    assume "normal s0"
+    moreover
+    assume "G\<turnstile>s0 \<midarrow>InsInitV c v=\<succ>vf\<midarrow>n\<rightarrow> s1"
+    ultimately have "False" 
+      by (cases s0) (simp add: evaln_InsInitV) 
+    thus "Q \<lfloor>vf\<rfloor>\<^sub>v s1 Z \<and> s1\<Colon>\<preceq>(G, L)"..
+  qed
+next
+  case (InsInitE A P Q c e)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} InsInitE c e-\<succ> {Q} }"
+  proof (rule valid_expr_NormalI)
+    fix s0 v n s1 L Z
+    assume "normal s0"
+    moreover
+    assume "G\<turnstile>s0 \<midarrow>InsInitE c e-\<succ>v\<midarrow>n\<rightarrow> s1"
+    ultimately have "False" 
+      by (cases s0) (simp add: evaln_InsInitE) 
+    thus "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"..
+  qed
+next
+  case (Callee A P Q e l)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} Callee l e-\<succ> {Q} }"
+  proof (rule valid_expr_NormalI)
+    fix s0 v n s1 L Z
+    assume "normal s0"
+    moreover
+    assume "G\<turnstile>s0 \<midarrow>Callee l e-\<succ>v\<midarrow>n\<rightarrow> s1"
+    ultimately have "False" 
+      by (cases s0) (simp add: evaln_Callee) 
+    thus "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"..
+  qed
+next
+  case (FinA A P Q a c)
+  show "G,A|\<Turnstile>\<Colon>{ {Normal P} .FinA a c. {Q} }"
+  proof (rule valid_stmt_NormalI)
+    fix s0 v n s1 L Z
+    assume "normal s0"
+    moreover
+    assume "G\<turnstile>s0 \<midarrow>FinA a c\<midarrow>n\<rightarrow> s1"
+    ultimately have "False" 
+      by (cases s0) (simp add: evaln_FinA) 
+    thus "Q \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G, L)"..
+  qed
+qed
+declare inj_term_simps [simp del]
+    
 theorem ax_sound: 
  "wf_prog G \<Longrightarrow> G,(A::'a triple set)|\<turnstile>(ts::'a triple set) \<Longrightarrow> G,A|\<Turnstile>ts"
 apply (subst ax_valids2_eq [symmetric])
@@ -472,5 +2665,9 @@
 apply (erule (1) ax_sound2)
 done
 
+lemma sound_valid2_lemma: 
+"\<lbrakk>\<forall>v n. Ball A (triple_valid2 G n) \<longrightarrow> P v n; Ball A (triple_valid2 G n)\<rbrakk>
+ \<Longrightarrow>P v n"
+by blast
 
 end
--- a/src/HOL/Bali/Basis.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Basis.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -247,6 +247,12 @@
 	"In1l e" == "In1 (Inl e)"
 	"In1r c" == "In1 (Inr c)"
 
+syntax the_In1l :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'al"
+       the_In1r :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'ar"
+translations
+   "the_In1l" == "the_Inl \<circ> the_In1"
+   "the_In1r" == "the_Inr \<circ> the_In1"
+
 ML {*
 fun sum3_instantiate thm = map (fn s => simplify(simpset()delsimps[not_None_eq])
  (read_instantiate [("t","In"^s^" ?x")] thm)) ["1l","2","3","1r"]
--- a/src/HOL/Bali/Conform.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Conform.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -115,6 +115,13 @@
 apply (simp add: conf_def)
 done
 
+lemma conf_Boolean: "G,s\<turnstile>v\<Colon>\<preceq>PrimT Boolean \<Longrightarrow> \<exists> b. v=Bool b"
+by (cases v)
+   (auto simp: conf_def obj_ty_def 
+         dest: widen_Boolean2 
+        split: obj_tag.splits)
+
+
 lemma conf_litval [rule_format (no_asm)]: 
   "typeof (\<lambda>a. None) v = Some T \<longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T"
 apply (unfold conf_def)
@@ -249,6 +256,97 @@
 apply force
 done
 
+section "weak value list conformance"
+
+text {* Only if the value is defined it has to conform to its type. 
+        This is the contribution of the definite assignment analysis to 
+        the notion of conformance. The definite assignment analysis ensures
+        that the program only attempts to access local variables that 
+        actually have a defined value in the state. 
+        So conformance must only ensure that the
+        defined values are of the right type, and not also that the value
+        is defined. 
+*}
+
+  
+constdefs
+
+  wlconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool"
+                                          ("_,_\<turnstile>_[\<sim>\<Colon>\<preceq>]_" [71,71,71,71] 70)
+           "G,s\<turnstile>vs[\<sim>\<Colon>\<preceq>]Ts \<equiv> \<forall>n. \<forall>T\<in>Ts n: \<forall> v\<in>vs n: G,s\<turnstile>v\<Colon>\<preceq>T"
+
+lemma wlconfD: "\<lbrakk>G,s\<turnstile>vs[\<sim>\<Colon>\<preceq>]Ts; Ts n = Some T; vs n = Some v\<rbrakk> \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T"
+by (auto simp: wlconf_def)
+
+
+lemma wlconf_cong [simp]: "\<And>s. G,set_locals x s\<turnstile>l[\<sim>\<Colon>\<preceq>]L = G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L"
+by (auto simp: wlconf_def)
+
+lemma wlconf_lupd [simp]: "G,lupd(vn\<mapsto>v)s\<turnstile>l[\<sim>\<Colon>\<preceq>]L = G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L"
+by (auto simp: wlconf_def)
+
+
+lemma wlconf_upd: "\<lbrakk>G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L; G,s\<turnstile>v\<Colon>\<preceq>T; L vn = Some T\<rbrakk> \<Longrightarrow>  
+  G,s\<turnstile>l(vn\<mapsto>v)[\<sim>\<Colon>\<preceq>]L"
+by (auto simp: wlconf_def)
+
+lemma wlconf_ext: "\<lbrakk>G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L; G,s\<turnstile>v\<Colon>\<preceq>T\<rbrakk> \<Longrightarrow> G,s\<turnstile>l(vn\<mapsto>v)[\<sim>\<Colon>\<preceq>]L(vn\<mapsto>T)"
+by (auto simp: wlconf_def)
+
+lemma wlconf_map_sum [simp]: 
+ "G,s\<turnstile>l1 (+) l2[\<sim>\<Colon>\<preceq>]L1 (+) L2 = (G,s\<turnstile>l1[\<sim>\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>l2[\<sim>\<Colon>\<preceq>]L2)"
+apply (unfold wlconf_def)
+apply safe
+apply (case_tac [3] "n")
+apply (force split add: sum.split)+
+done
+
+lemma wlconf_ext_list [rule_format (no_asm)]: "
+ \<And>X. \<lbrakk>G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L\<rbrakk> \<Longrightarrow> 
+      \<forall>vs Ts. distinct vns \<longrightarrow> length Ts = length vns 
+      \<longrightarrow> list_all2 (conf G s) vs Ts \<longrightarrow> G,s\<turnstile>l(vns[\<mapsto>]vs)[\<sim>\<Colon>\<preceq>]L(vns[\<mapsto>]Ts)"
+apply (unfold wlconf_def)
+apply (induct_tac "vns")
+apply  clarsimp
+apply clarsimp
+apply (frule list_all2_lengthD)
+apply clarsimp
+done
+
+
+lemma wlconf_deallocL: "\<lbrakk>G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L(vn\<mapsto>T); L vn = None\<rbrakk> \<Longrightarrow> G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L"
+apply (simp only: wlconf_def)
+apply safe
+apply (drule spec)
+apply (drule ospec)
+defer
+apply (drule ospec )
+apply auto
+done 
+
+
+lemma wlconf_gext [elim]: "\<lbrakk>G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L; s\<le>|s'\<rbrakk> \<Longrightarrow> G,s'\<turnstile>l[\<sim>\<Colon>\<preceq>]L"
+apply (simp only: wlconf_def)
+apply fast
+done
+
+lemma wlconf_empty [simp, intro!]: "G,s\<turnstile>vs[\<sim>\<Colon>\<preceq>]empty"
+apply (unfold wlconf_def)
+apply force
+done
+
+lemma wlconf_empty_vals: "G,s\<turnstile>empty[\<sim>\<Colon>\<preceq>]ts"
+  by (simp add: wlconf_def)
+
+lemma wlconf_init_vals [intro!]: 
+	" \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<sim>\<Colon>\<preceq>]fs"
+apply (unfold wlconf_def)
+apply force
+done
+
+lemma lconf_wlconf:
+ "G,s\<turnstile>l[\<Colon>\<preceq>]L \<Longrightarrow> G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L"
+by (force simp add: lconf_def wlconf_def)
 
 section "object conformance"
 
@@ -259,18 +357,8 @@
                            (case r of 
 		              Heap a \<Rightarrow> is_type G (obj_ty obj) 
                             | Stat C \<Rightarrow> True)"
-(*
-lemma oconf_def2:  "G,s\<turnstile>\<lparr>tag=oi,values=fs\<rparr>\<Colon>\<preceq>\<surd>r =  
-  (G,s\<turnstile>fs[\<Colon>\<preceq>]var_tys G oi r \<and> 
-  (case r of Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=fs\<rparr>) | Stat C \<Rightarrow> True))"
-by (simp add: oconf_def Let_def)
-*)
-(*
-lemma oconf_def2:  "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r =  
-  (G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and> 
-  (case r of Heap a \<Rightarrow> is_type G (obj_ty obj) | Stat C \<Rightarrow> True))"
-by (simp add: oconf_def Let_def)
-*)
+
+
 lemma oconf_is_type: "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>Heap a \<Longrightarrow> is_type G (obj_ty obj)"
 by (auto simp: oconf_def Let_def)
 
@@ -297,33 +385,16 @@
            split add: sum.split_asm obj_tag.split_asm)
 done
 
-(*
-lemma oconf_init_obj_lemma: 
-"\<lbrakk>\<And>C c. class G C = Some c \<Longrightarrow> unique (fields G C);  
-  \<And>C c f fld. \<lbrakk>class G C = Some c; table_of (fields G C) f = Some fld \<rbrakk> 
-            \<Longrightarrow> is_type G (type fld);  
-  (case r of 
-     Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=fs\<rparr>) 
-  | Stat C \<Rightarrow> is_class G C)
-\<rbrakk> \<Longrightarrow>  G,s\<turnstile>\<lparr>tag=oi, values=init_vals (var_tys G oi r)\<rparr>\<Colon>\<preceq>\<surd>r"
-apply (auto simp add: oconf_def)
-apply (drule_tac var_tys_Some_eq [THEN iffD1]) 
-defer
-apply (subst obj_ty_eq)
-apply(auto dest!: fields_table_SomeD split add: sum.split_asm obj_tag.split_asm)
-done
-*)
-
-
 section "state conformance"
 
 constdefs
 
   conforms :: "state \<Rightarrow> env_ \<Rightarrow> bool"          (     "_\<Colon>\<preceq>_"   [71,71]      70)
    "xs\<Colon>\<preceq>E \<equiv> let (G, L) = E; s = snd xs; l = locals s in
-      (\<forall>r. \<forall>obj\<in>globs s r:           G,s\<turnstile>obj   \<Colon>\<preceq>\<surd>r) \<and>
-                  \<spacespace>                   G,s\<turnstile>l    [\<Colon>\<preceq>]L\<spacespace> \<and>
-      (\<forall>a. fst xs=Some(Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable))"
+    (\<forall>r. \<forall>obj\<in>globs s r:           G,s\<turnstile>obj   \<Colon>\<preceq>\<surd>r) \<and>
+                \<spacespace>                   G,s\<turnstile>l    [\<sim>\<Colon>\<preceq>]L\<spacespace> \<and>
+    (\<forall>a. fst xs=Some(Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)) \<and>
+         (fst xs=Some(Jump Ret) \<longrightarrow> l Result \<noteq> None)"
 
 section "conforms"
 
@@ -331,13 +402,17 @@
 "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); globs s r = Some obj\<rbrakk> \<Longrightarrow> G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r"
 by (auto simp: conforms_def Let_def)
 
-lemma conforms_localD: "(x, s)\<Colon>\<preceq>(G, L) \<Longrightarrow> G,s\<turnstile>locals s[\<Colon>\<preceq>]L"
+lemma conforms_localD: "(x, s)\<Colon>\<preceq>(G, L) \<Longrightarrow> G,s\<turnstile>locals s[\<sim>\<Colon>\<preceq>]L"
 by (auto simp: conforms_def Let_def)
 
 lemma conforms_XcptLocD: "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); x = Some (Xcpt (Loc a))\<rbrakk> \<Longrightarrow>  
 	  G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)"
 by (auto simp: conforms_def Let_def)
 
+lemma conforms_RetD: "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); x = Some (Jump Ret)\<rbrakk> \<Longrightarrow>  
+	  (locals s) Result \<noteq> None"
+by (auto simp: conforms_def Let_def)
+
 lemma conforms_RefTD: 
  "\<lbrakk>G,s\<turnstile>a'\<Colon>\<preceq>RefT t; a' \<noteq> Null; (x,s) \<Colon>\<preceq>(G, L)\<rbrakk> \<Longrightarrow>  
    \<exists>a obj. a' = Addr a \<and> globs s (Inl a) = Some obj \<and>  
@@ -349,8 +424,9 @@
 done
 
 lemma conforms_Jump [iff]:
-  "((Some (Jump j), s)\<Colon>\<preceq>(G, L)) = (Norm s\<Colon>\<preceq>(G, L))"
-by (auto simp: conforms_def)
+  "j=Ret \<longrightarrow> locals s Result \<noteq> None 
+   \<Longrightarrow> ((Some (Jump j), s)\<Colon>\<preceq>(G, L)) = (Norm s\<Colon>\<preceq>(G, L))"
+by (auto simp: conforms_def Let_def)
 
 lemma conforms_StdXcpt [iff]: 
   "((Some (Xcpt (Std xn)), s)\<Colon>\<preceq>(G, L)) = (Norm s\<Colon>\<preceq>(G, L))"
@@ -382,45 +458,61 @@
 done
 
 lemma conformsI: "\<lbrakk>\<forall>r. \<forall>obj\<in>globs s r: G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r;  
-     G,s\<turnstile>locals s[\<Colon>\<preceq>]L;  
-     \<forall>a. x = Some (Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)\<rbrakk> \<Longrightarrow> 
+     G,s\<turnstile>locals s[\<sim>\<Colon>\<preceq>]L;  
+     \<forall>a. x = Some (Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable);
+     x = Some (Jump Ret)\<longrightarrow> locals s Result \<noteq> None\<rbrakk> \<Longrightarrow> 
   (x, s)\<Colon>\<preceq>(G, L)"
 by (auto simp: conforms_def Let_def)
 
 lemma conforms_xconf: "\<lbrakk>(x, s)\<Colon>\<preceq>(G,L);   
- \<forall>a. x' = Some (Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)\<rbrakk> \<Longrightarrow> 
+ \<forall>a. x' = Some (Xcpt (Loc a)) \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable);
+     x' = Some (Jump Ret) \<longrightarrow> locals s Result \<noteq> None\<rbrakk> \<Longrightarrow> 
  (x',s)\<Colon>\<preceq>(G,L)"
 by (fast intro: conformsI elim: conforms_globsD conforms_localD)
 
 lemma conforms_lupd: 
  "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); L vn = Some T; G,s\<turnstile>v\<Colon>\<preceq>T\<rbrakk> \<Longrightarrow> (x, lupd(vn\<mapsto>v)s)\<Colon>\<preceq>(G, L)"
-by (force intro: conformsI lconf_upd dest: conforms_globsD conforms_localD 
-                                           conforms_XcptLocD simp: oconf_def)
+by (force intro: conformsI wlconf_upd dest: conforms_globsD conforms_localD 
+                                           conforms_XcptLocD conforms_RetD 
+          simp: oconf_def)
 
 
-lemmas conforms_allocL_aux = conforms_localD [THEN lconf_ext]
+lemmas conforms_allocL_aux = conforms_localD [THEN wlconf_ext]
 
 lemma conforms_allocL: 
   "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); G,s\<turnstile>v\<Colon>\<preceq>T\<rbrakk> \<Longrightarrow> (x, lupd(vn\<mapsto>v)s)\<Colon>\<preceq>(G, L(vn\<mapsto>T))"
-by (force intro: conformsI dest: conforms_globsD 
-          elim: conforms_XcptLocD conforms_allocL_aux simp: oconf_def)
+by (force intro: conformsI dest: conforms_globsD conforms_RetD 
+          elim: conforms_XcptLocD  conforms_allocL_aux 
+          simp: oconf_def)
 
-lemmas conforms_deallocL_aux = conforms_localD [THEN lconf_deallocL]
+lemmas conforms_deallocL_aux = conforms_localD [THEN wlconf_deallocL]
 
 lemma conforms_deallocL: "\<And>s.\<lbrakk>s\<Colon>\<preceq>(G, L(vn\<mapsto>T)); L vn = None\<rbrakk> \<Longrightarrow> s\<Colon>\<preceq>(G,L)"
-by (fast intro: conformsI dest: conforms_globsD 
+by (fast intro: conformsI dest: conforms_globsD conforms_RetD
          elim: conforms_XcptLocD conforms_deallocL_aux)
 
 lemma conforms_gext: "\<lbrakk>(x, s)\<Colon>\<preceq>(G,L); s\<le>|s';  
   \<forall>r. \<forall>obj\<in>globs s' r: G,s'\<turnstile>obj\<Colon>\<preceq>\<surd>r;  
    locals s'=locals s\<rbrakk> \<Longrightarrow> (x,s')\<Colon>\<preceq>(G,L)"
-by (force intro!: conformsI dest: conforms_localD conforms_XcptLocD)
+apply (rule conformsI)
+apply     assumption
+apply    (drule conforms_localD) apply force
+apply   (intro strip)
+apply  (drule (1) conforms_XcptLocD) apply force 
+apply (intro strip)
+apply (drule (1) conforms_RetD) apply force
+done
+
 
 
 lemma conforms_xgext: 
-  "\<lbrakk>(x ,s)\<Colon>\<preceq>(G,L); (x', s')\<Colon>\<preceq>(G, L); s'\<le>|s\<rbrakk> \<Longrightarrow> (x',s)\<Colon>\<preceq>(G,L)"
+  "\<lbrakk>(x ,s)\<Colon>\<preceq>(G,L); (x', s')\<Colon>\<preceq>(G, L); s'\<le>|s;dom (locals s') \<subseteq> dom (locals s)\<rbrakk> 
+   \<Longrightarrow> (x',s)\<Colon>\<preceq>(G,L)"
 apply (erule_tac conforms_xconf)
-apply (fast dest: conforms_XcptLocD)
+apply  (fast dest: conforms_XcptLocD)
+apply (intro strip)
+apply (drule (1) conforms_RetD) 
+apply (auto dest: domI)
 done
 
 lemma conforms_gupd: "\<And>obj. \<lbrakk>(x, s)\<Colon>\<preceq>(G, L); G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r; s\<le>|gupd(r\<mapsto>obj)s\<rbrakk> 
@@ -445,17 +537,29 @@
 done
 
 lemma conforms_set_locals: 
-  "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L'); G,s\<turnstile>l[\<Colon>\<preceq>]L\<rbrakk> \<Longrightarrow> (x,set_locals l s)\<Colon>\<preceq>(G,L)"
-apply (auto intro!: conformsI dest: conforms_globsD 
-            elim!: conforms_XcptLocD simp add: oconf_def)
+  "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L'); G,s\<turnstile>l[\<sim>\<Colon>\<preceq>]L; x=Some (Jump Ret) \<longrightarrow> l Result \<noteq> None\<rbrakk> 
+   \<Longrightarrow> (x,set_locals l s)\<Colon>\<preceq>(G,L)"
+apply (rule conformsI)
+apply     (intro strip)
+apply     simp
+apply     (drule (2) conforms_globsD)
+apply    simp
+apply   (intro strip)
+apply   (drule (1) conforms_XcptLocD)
+apply   simp
+apply (intro strip)
+apply (drule (1) conforms_RetD)
+apply simp
 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)
+lemma conforms_locals: 
+  "\<lbrakk>(a,b)\<Colon>\<preceq>(G, L); L x = Some T;locals b x \<noteq>None\<rbrakk>
+   \<Longrightarrow> G,b\<turnstile>the (locals b x)\<Colon>\<preceq>T"
+apply (force simp: conforms_def Let_def wlconf_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>  
+lemma conforms_return: 
+"\<And>s'. \<lbrakk>(x,s)\<Colon>\<preceq>(G, L); (x',s')\<Colon>\<preceq>(G, L'); s\<le>|s';x'\<noteq>Some (Jump Ret)\<rbrakk> \<Longrightarrow>  
   (x',set_locals (locals s) s')\<Colon>\<preceq>(G, L)"
 apply (rule conforms_xconf)
 prefer 2 apply (force dest: conforms_XcptLocD)
--- a/src/HOL/Bali/Decl.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Decl.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -291,13 +291,13 @@
 subsection {* Interface *}
 
 
-record  ibody = decl +  (* interface body *)
-          imethods :: "(sig \<times> mhead) list" (* method heads *)
+record  ibody = decl +  --{* interface body *}
+          imethods :: "(sig \<times> mhead) list" --{* method heads *}
 
-record  iface = ibody + (* interface *)
-         isuperIfs:: "qtname list" (* superinterface list *)
+record  iface = ibody + --{* interface *}
+         isuperIfs:: "qtname list" --{* superinterface list *}
 types	
-	idecl           (* interface declaration, cf. 9.1 *)
+	idecl           --{* interface declaration, cf. 9.1 *}
 	= "qtname \<times> iface"
 
 translations
@@ -320,16 +320,16 @@
 by (simp add: ibody_def)
 
 subsection  {* Class *}
-record cbody = decl +          (* class body *)
+record cbody = decl +          --{* class body *}
          cfields:: "fdecl list" 
          methods:: "mdecl list"
-         init   :: "stmt"       (* initializer *)
+         init   :: "stmt"       --{* initializer *}
 
-record class = cbody +           (* class *)
-        super   :: "qtname"      (* superclass *)
-        superIfs:: "qtname list" (* implemented interfaces *)
+record class = cbody +           --{* class *}
+        super   :: "qtname"      --{* superclass *}
+        superIfs:: "qtname list" --{* implemented interfaces *}
 types	
-	cdecl           (* class declaration, cf. 8.1 *)
+	cdecl           --{* class declaration, cf. 8.1 *}
 	= "qtname \<times> class"
 
 translations
@@ -366,10 +366,10 @@
 
 consts
 
-  Object_mdecls  ::  "mdecl list" (* methods of Object *)
-  SXcpt_mdecls   ::  "mdecl list" (* methods of SXcpts *)
-  ObjectC ::         "cdecl"      (* declaration  of root      class   *)
-  SXcptC  ::"xname \<Rightarrow> cdecl"      (* declarations of throwable classes *)
+  Object_mdecls  ::  "mdecl list" --{* methods of Object *}
+  SXcpt_mdecls   ::  "mdecl list" --{* methods of SXcpts *}
+  ObjectC ::         "cdecl"      --{* declaration  of root      class   *}
+  SXcptC  ::"xname \<Rightarrow> cdecl"      --{* declarations of throwable classes *}
 
 defs 
 
@@ -442,8 +442,8 @@
 section "subinterface and subclass relation, in anticipation of TypeRel.thy"
 
 consts 
-  subint1  :: "prog \<Rightarrow> (qtname \<times> qtname) set"
-  subcls1  :: "prog \<Rightarrow> (qtname \<times> qtname) set"
+  subint1  :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{* direct subinterface *}
+  subcls1  :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{* direct subclass     *}
 
 defs
   subint1_def: "subint1 G \<equiv> {(I,J). \<exists>i\<in>iface G I: J\<in>set (isuperIfs i)}"
@@ -795,65 +795,10 @@
 apply (rule class_rec.simps [THEN trans [THEN fun_cong [THEN fun_cong]]])
 apply simp
 done
-(*
-lemma bar:
- "[| P;  !!x.  P ==> Q x  |] ==> Q x"
-by simp
-
-lemma metaMP: "[| A ==> B; A |] ==> B"
-by blast
-
-lemma True
-proof- 
-  presume t: "C  ==> E"
-  thm metaMP [OF t]
-
-  presume r1: "\<And> B. P \<Longrightarrow> B"
-  presume r2: "\<And> C. C \<Longrightarrow> P"
-  thm r1 [OF r2]
-
-  thm metaMP [OF t]
-
-lemma ws_subcls1_induct4: "\<lbrakk>is_class G C; ws_prog G;  
-  \<And>C c. \<lbrakk>C \<noteq> Object\<longrightarrow> P (super c)\<rbrakk> \<Longrightarrow> P C
- \<rbrakk> \<Longrightarrow> P C"
-proof -
-  assume cls_C: "is_class G C"
-  and       ws: "ws_prog G"
-  and      hyp: "\<And>C c. \<lbrakk>C \<noteq> Object\<longrightarrow> P (super c)\<rbrakk> \<Longrightarrow> P C"
-  thm ws_subcls1_induct [OF cls_C ws hyp]
-
-show
-(\<And>C c. class G C = Some c \<and>
-       (C \<noteq> Object \<longrightarrow> G\<turnstile>C\<prec>\<^sub>C\<^sub>1super c \<and> ?P (super c) \<and> is_class G (super c)) \<Longrightarrow>
-       ?P C) \<Longrightarrow>
-?P C
-  show ?thesis
-    thm "thm ws_subcls1_induct [OF cls_C ws hyp]"
-    apply (rule ws_subcls1_induct)
-  proof (rule ws_subcls1_induct)
-    fix C c
-    assume "class G C = Some c \<and>
-            (C \<noteq> Object \<longrightarrow>
-              G\<turnstile>C\<prec>\<^sub>C\<^sub>1super c \<and> P (super c) \<and> is_class G (super c))"
-    show "C \<noteq> Object \<longrightarrow> P (super  (?c C c))" 
-apply (erule ws_subcls1_induct)
-apply assumption
-apply (erule conjE)
-apply (case_tac "C=Object")
-apply blast
-apply (erule impE)
-apply assumption
-apply (erule conjE)+
-apply (rotate_tac 2)
-sorry
-
-*)
-
 
 constdefs
 imethds:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
-  (* methods of an interface, with overriding and inheritance, cf. 9.2 *)
+  --{* methods of an interface, with overriding and inheritance, cf. 9.2 *}
 "imethds G I 
   \<equiv> iface_rec (G,I)  
               (\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus> 
--- a/src/HOL/Bali/DeclConcepts.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/DeclConcepts.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -153,10 +153,6 @@
 axclass has_static < "type"
 consts is_static :: "'a::has_static \<Rightarrow> bool"
 
-(*
-consts is_static :: "'a \<Rightarrow> bool"
-*)
-
 instance access_field_type :: ("type","has_static") has_static ..
 
 defs (overloaded)
@@ -205,48 +201,31 @@
 lemma mhead_static_simp [simp]: "is_static (mhead m) = is_static m"
 by (cases m) (simp add: mhead_def member_is_static_simp)
 
-constdefs  (* some mnemotic selectors for (qtname \<times> ('a::more) decl_scheme) 
-            * the first component is a class or interface name
-            * the second component is a method, field or method head *)
-(* "declclass":: "(qtname \<times> ('a::more) decl_scheme) \<Rightarrow> qtname"*)
-(* "declclass \<equiv> fst" *)          (* get the class component *)
-
+constdefs  --{* some mnemotic selectors for various pairs *} 
+           
  "decliface":: "(qtname \<times> ('a::type) decl_scheme) \<Rightarrow> qtname"
- "decliface \<equiv> fst"          (* get the interface component *)
+ "decliface \<equiv> fst"          --{* get the interface component *}
 
-(*
- "member"::   "(qtname \<times> ('a::type) decl_scheme) \<Rightarrow> ('a::type) decl_scheme"
-*)
  "mbr"::   "(qtname \<times> memberdecl) \<Rightarrow> memberdecl"
- "mbr \<equiv> snd"            (* get the memberdecl component *)
+ "mbr \<equiv> snd"            --{* get the memberdecl component *}
 
  "mthd"::   "('b \<times> 'a) \<Rightarrow> 'a"
-                           (* also used for mdecl,mhead *)
- "mthd \<equiv> snd"              (* get the method component *)
+                           --{* also used for mdecl, mhead *}
+ "mthd \<equiv> snd"              --{* get the method component *}
 
  "fld"::   "('b \<times> ('a::type) decl_scheme) \<Rightarrow> ('a::type) decl_scheme"
-              (* also used for ((vname \<times> qtname)\<times> field) *)
- "fld \<equiv> snd"               (* get the field component *)
+              --{* also used for @{text "((vname \<times> qtname)\<times> field)"} *}
+ "fld \<equiv> snd"               --{* get the field component *}
 
-(* "accmodi" :: "('b \<times> ('a::type) decl_scheme) \<Rightarrow> acc_modi"*)
-                           (* also used for mdecl *)
-(* "accmodi \<equiv> access \<circ> snd"*)  (* get the access modifier *) 
-(*
- "is_static" ::"('b \<times> ('a::type) member_scheme) \<Rightarrow> bool" *)
-                            (* also defined for emhead cf. WellType *)
- (*"is_static \<equiv> static \<circ> snd"*) (* get the static modifier *)
 
-constdefs (* some mnemotic selectors for (vname \<times> qtname) *)
- fname:: "(vname \<times> 'a) \<Rightarrow> vname" (* also used for fdecl *)
+constdefs --{* some mnemotic selectors for @{text "(vname \<times> qtname)"} *}
+ fname:: "(vname \<times> 'a) \<Rightarrow> vname" --{* also used for fdecl *}
  "fname \<equiv> fst"
   
   declclassf:: "(vname \<times> qtname) \<Rightarrow> qtname"
  "declclassf \<equiv> snd"
 
-(*
-lemma declclass_simp[simp]: "declclass (C,m) = C"
-by (simp add: declclass_def)
-*)
+
 
 lemma decliface_simp[simp]: "decliface (I,m) = I"
 by (simp add: decliface_def) 
@@ -272,11 +251,6 @@
 lemma access_fld_simp [simp]: "(access (fld f)) = accmodi f"
 by (cases f) (simp add:  fld_def) 
 
-(*
-lemma is_static_simp[simp]: "is_static (C,m) = static m"
-by (simp add: is_static_def)
-*)
-
 lemma static_mthd_simp[simp]: "static (mthd m) = is_static m"
 by (cases m) (simp add:  mthd_def member_is_static_simp)
 
@@ -301,7 +275,7 @@
 lemma declclassf_simp[simp]:"declclassf (n,c) = c"
 by (simp add: declclassf_def)
 
-constdefs  (* some mnemotic selectors for (vname \<times> qtname) *)
+constdefs  --{* some mnemotic selectors for @{text "(vname \<times> qtname)"} *}
   "fldname"  :: "(vname \<times> qtname) \<Rightarrow> vname" 
   "fldname \<equiv> fst"
 
@@ -1265,6 +1239,7 @@
   qed
 qed
 *)
+
 lemma accessible_fieldD: 
  "\<lbrakk>G\<turnstile>membr of C accessible_from accC; is_field membr\<rbrakk>
  \<Longrightarrow> G\<turnstile>membr member_of C \<and>
@@ -1272,34 +1247,7 @@
      G\<turnstile>membr in C permits_acc_to accC"
 by (induct rule: accessible_fromR.induct) (auto dest: is_fieldD)
       
-(* lemmata:
- Wegen  G\<turnstile>Super accessible_in (pid C) folgt:
-  G\<turnstile>m declared_in C; G\<turnstile>m member_of D; accmodi m = Package (G\<turnstile>D \<preceq>\<^sub>C C)
-  \<Longrightarrow> pid C = pid D 
 
-  C package
-  m public in C
-  für alle anderen D: G\<turnstile>m undeclared_in C
-  m wird in alle subklassen vererbt, auch aus dem Package heraus!
-
-  G\<turnstile>m member_of C \<Longrightarrow> \<exists> D. G\<turnstile>C \<preceq>\<^sub>C D \<and> G\<turnstile>m declared_in D
-*)
-
-(* Begriff (C,m) overrides (D,m)
-    3 Fälle: Direkt,
-             Indirekt über eine Zwischenklasse (ohne weiteres override)
-             Indirekt über override
-*)
-   
-(*
-"G\<turnstile>m member_of C \<equiv> 
-constdefs declares_method:: "prog \<Rightarrow> sig \<Rightarrow> qtname \<Rightarrow> methd \<Rightarrow> bool"
-                                 ("_,_\<turnstile> _ declares'_method _" [61,61,61,61] 60)
-"G,sig\<turnstile>C declares_method m \<equiv> cdeclaredmethd G C sig = Some m" 
-
-constdefs is_declared:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool"
-"is_declared G sig em \<equiv> G,sig\<turnstile>declclass em declares_method mthd em"
-*)
 
 lemma member_of_Private:
 "\<lbrakk>G\<turnstile>m member_of C; accmodi m = Private\<rbrakk> \<Longrightarrow> declclass m = C"
@@ -1541,23 +1489,7 @@
                       ) sig
               )
         else None)"
-(*
-"dynmethd G statC dynC  
-  \<equiv> \<lambda> sig. 
-     (if G\<turnstile>dynC \<preceq>\<^sub>C statC
-        then (case methd G statC sig of
-                None \<Rightarrow> None
-              | Some statM 
-                    \<Rightarrow> (class_rec (G,statC) empty
-                         (\<lambda>C c subcls_mthds. 
-                            subcls_mthds
-                            ++
-                            (filter_tab 
-                              (\<lambda> _ dynM. G,sig\<turnstile>dynM overrides statM)  
-                              (table_of (map (\<lambda>(s,m). (s,C,m)) (methods c)))))
-                        ) sig
-              )
-        else None)"*)
+
 text {* @{term "dynmethd G statC dynC"}: dynamic method lookup of a reference 
         with dynamic class @{term dynC} and static class @{term statC} *}
 text {* Note some kind of duality between @{term methd} and @{term dynmethd} 
@@ -2014,8 +1946,6 @@
 lemma dynmethdSomeD: 
  "\<lbrakk>dynmethd G statC dynC sig = Some dynM; is_class G dynC; ws_prog G\<rbrakk> 
   \<Longrightarrow> G\<turnstile>dynC \<preceq>\<^sub>C statC \<and> (\<exists> statM. methd G statC sig = Some statM)"
-apply clarify
-apply rotate_tac
 by (auto simp add: dynmethd_rec)
  
 lemma dynmethd_Some_cases [consumes 3, case_names Static Overrides]:
@@ -2243,31 +2173,6 @@
   qed
 qed
 
-(*
-lemma dom_dynmethd: 
-  "dom (dynmethd G statC dynC) \<subseteq> dom (methd G statC) \<union> dom (methd G dynC)"
-by (auto simp add: dynmethd_def dom_def)
-
-lemma finite_dom_dynmethd:
- "\<lbrakk>ws_prog G; is_class G statC; is_class G dynC\<rbrakk> 
-  \<Longrightarrow> finite (dom (dynmethd G statC dynC))"
-apply (rule_tac B="dom (methd G statC) \<union> dom (methd G dynC)" in finite_subset)
-apply (rule dom_dynmethd)
-apply (rule finite_UnI)
-apply (drule (2) finite_dom_methd)+
-done
-*)
-(*
-lemma dynmethd_SomeD: 
-"\<lbrakk>ws_prog G; is_class G statC; is_class G dynC;
- methd G statC sig = Some sm; dynmethd G statC dynC sig = Some dm; sm \<noteq> dm
- \<rbrakk> \<Longrightarrow> G\<turnstile>dynC \<prec>\<^sub>C statC \<and> 
-       (declclass dm \<noteq> dynC \<longrightarrow> G \<turnstile> dm accessible_through_inheritance_in dynC)"
-by (auto simp add: dynmethd_def 
-         dest: methd_inheritedD methd_diff_cls
-         intro: rtrancl_into_trancl3)
-*)
-
 subsection "dynlookup"
 
 lemma dynlookup_cases [consumes 1, case_names NullT IfaceT ClassT ArrayT]:
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/DefiniteAssignment.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -0,0 +1,1853 @@
+header {* Definite Assignment *}
+
+theory DefiniteAssignment = WellType: 
+
+text {* Definite Assignment Analysis (cf. 16)
+
+The definite assignment analysis approximates the sets of local 
+variables that will be assigned at a certain point of evaluation, and ensures
+that we will only read variables which previously were assigned.
+It should conform to the following idea:
+ If the evaluation of a term completes normally (no abruption (exception, 
+break, continue, return) appeared) , the set of local variables calculated 
+by the analysis is a subset of the
+variables that were actually assigned during evaluation.
+
+To get more precise information about the sets of assigned variables the 
+analysis includes the following optimisations:
+\begin{itemize}
+  \item Inside of a while loop we also take care of the variables assigned
+        before break statements, since the break causes the while loop to
+        continue normally.
+  \item For conditional statements we take care of constant conditions to 
+        statically determine the path of evaluation.
+  \item Inside a distinct path of a conditional statements we know to which
+        boolean value the condition has evaluated to, and so can retrieve more
+        information about the variables assigned during evaluation of the
+        boolean condition.
+\end{itemize}
+
+Since in our model of Java the return values of methods are stored in a local
+variable we also ensure that every path of (normal) evaluation will assign the
+result variable, or in the sense of real Java every path ends up in and 
+return instruction. 
+
+Not covered yet:
+\begin{itemize} 
+  \item analysis of definite unassigned
+  \item special treatment of final fields
+\end{itemize}
+*}
+
+section {* Correct nesting of jump statements *}
+
+text {* For definite assignment it becomes crucial, that jumps (break, 
+continue, return) are nested correctly i.e. a continue jump is nested in a
+matching while statement, a break jump is nested in a proper label statement,
+a class initialiser does not terminate abruptly with a return. With this we 
+can for example ensure that evaluation of an expression will never end up 
+with a jump, since no breaks, continues or returns are allowed in an 
+expression. *}
+
+consts jumpNestingOkS :: "jump set \<Rightarrow> stmt \<Rightarrow> bool"
+primrec
+"jumpNestingOkS jmps (Skip)   = True"
+"jumpNestingOkS jmps (Expr e) = True"
+"jumpNestingOkS jmps (j\<bullet> s) = jumpNestingOkS ({j} \<union> jmps) s"
+"jumpNestingOkS jmps (c1;;c2) = (jumpNestingOkS jmps c1 \<and> 
+                                 jumpNestingOkS jmps c2)"
+"jumpNestingOkS jmps (If(e) c1 Else c2) = (jumpNestingOkS jmps c1 \<and>  
+                                           jumpNestingOkS jmps c2)"
+"jumpNestingOkS jmps (l\<bullet> While(e) c) = jumpNestingOkS ({Cont l} \<union> jmps) c"
+--{* The label of the while loop only handles continue jumps. Breaks are only
+     handled by @{term Lab} *}
+"jumpNestingOkS jmps (Jmp j) = (j \<in> jmps)"
+"jumpNestingOkS jmps (Throw e) = True"
+"jumpNestingOkS jmps (Try c1 Catch(C vn) c2) = (jumpNestingOkS jmps c1 \<and> 
+                                                jumpNestingOkS jmps c2)"
+"jumpNestingOkS jmps (c1 Finally c2) = (jumpNestingOkS jmps c1 \<and> 
+                                        jumpNestingOkS jmps c2)"
+"jumpNestingOkS jmps (Init C) = True" 
+ --{* wellformedness of the program must enshure that for all initializers 
+      jumpNestingOkS {} holds *} 
+--{* Dummy analysis for intermediate smallstep term @{term  FinA} *}
+"jumpNestingOkS jmps (FinA a c) = False"
+
+
+constdefs jumpNestingOk :: "jump set \<Rightarrow> term \<Rightarrow> bool"
+"jumpNestingOk jmps t \<equiv> (case t of
+                      In1 se \<Rightarrow> (case se of
+                                   Inl e \<Rightarrow> True
+                                 | Inr s \<Rightarrow> jumpNestingOkS jmps s)
+                    | In2  v \<Rightarrow> True
+                    | In3  es \<Rightarrow> True)"
+
+lemma jumpNestingOk_expr_simp [simp]: "jumpNestingOk jmps (In1l e) = True"
+by (simp add: jumpNestingOk_def)
+
+lemma jumpNestingOk_expr_simp1 [simp]: "jumpNestingOk jmps \<langle>e::expr\<rangle> = True"
+by (simp add: inj_term_simps)
+
+lemma jumpNestingOk_stmt_simp [simp]: 
+  "jumpNestingOk jmps (In1r s) = jumpNestingOkS jmps s"
+by (simp add: jumpNestingOk_def)
+
+lemma jumpNestingOk_stmt_simp1 [simp]: 
+   "jumpNestingOk jmps \<langle>s::stmt\<rangle> = jumpNestingOkS jmps s"
+by (simp add: inj_term_simps)
+
+lemma jumpNestingOk_var_simp [simp]: "jumpNestingOk jmps (In2 v) = True"
+by (simp add: jumpNestingOk_def)
+
+lemma jumpNestingOk_var_simp1 [simp]: "jumpNestingOk jmps \<langle>v::var\<rangle> = True"
+by (simp add: inj_term_simps)
+
+lemma jumpNestingOk_expr_list_simp [simp]: "jumpNestingOk jmps (In3 es) = True"
+by (simp add: jumpNestingOk_def)
+
+lemma jumpNestingOk_expr_list_simp1 [simp]: 
+  "jumpNestingOk jmps \<langle>es::expr list\<rangle> = True"
+by (simp add: inj_term_simps)
+
+
+
+section {* Calculation of assigned variables for boolean expressions*}
+
+
+subsection {* Very restricted calculation fallback calculation *}
+
+consts the_LVar_name:: "var \<Rightarrow> lname"
+primrec 
+"the_LVar_name (LVar n) = n"
+
+consts assignsE :: "expr      \<Rightarrow> lname set" 
+       assignsV :: "var       \<Rightarrow> lname set"
+       assignsEs:: "expr list \<Rightarrow> lname set"
+text {* *}
+primrec
+"assignsE (NewC c)            = {}" 
+"assignsE (NewA t e)          = assignsE e"
+"assignsE (Cast t e)          = assignsE e"
+"assignsE (e InstOf r)        = assignsE e"
+"assignsE (Lit val)           = {}"
+"assignsE (UnOp unop e)       = assignsE e"
+"assignsE (BinOp binop e1 e2) = (if binop=CondAnd \<or> binop=CondOr
+                                     then (assignsE e1)
+                                     else (assignsE e1) \<union> (assignsE e2))" 
+"assignsE (Super)             = {}"
+"assignsE (Acc v)             = assignsV v"
+"assignsE (v:=e)              = (assignsV v) \<union> (assignsE e) \<union> 
+                                 (if \<exists> n. v=(LVar n) then {the_LVar_name v} 
+                                                     else {})"
+"assignsE (b? e1 : e2) = (assignsE b) \<union> ((assignsE e1) \<inter> (assignsE e2))"
+"assignsE ({accC,statT,mode}objRef\<cdot>mn({pTs}args)) 
+                          = (assignsE objRef) \<union> (assignsEs args)"
+-- {* Only dummy analysis for intermediate expressions  
+      @{term Methd}, @{term Body}, @{term InsInitE} and @{term Callee} *}
+"assignsE (Methd C sig)   = {}" 
+"assignsE (Body  C s)     = {}"   
+"assignsE (InsInitE s e)  = {}"  
+"assignsE (Callee l e)    = {}" 
+
+"assignsV (LVar n)       = {}"
+"assignsV ({accC,statDeclC,stat}objRef..fn) = assignsE objRef"
+"assignsV (e1.[e2])      = assignsE e1 \<union> assignsE e2"
+
+"assignsEs     [] = {}"
+"assignsEs (e#es) = assignsE e \<union> assignsEs es"
+
+constdefs assigns:: "term \<Rightarrow> lname set"
+"assigns t \<equiv> (case t of
+                In1 se \<Rightarrow> (case se of
+                             Inl e \<Rightarrow> assignsE e
+                           | Inr s \<Rightarrow> {})
+              | In2  v \<Rightarrow> assignsV v
+              | In3  es \<Rightarrow> assignsEs es)"
+
+lemma assigns_expr_simp [simp]: "assigns (In1l e) = assignsE e"
+by (simp add: assigns_def)
+
+lemma assigns_expr_simp1 [simp]: "assigns (\<langle>e\<rangle>) = assignsE e"
+by (simp add: inj_term_simps)
+
+lemma assigns_stmt_simp [simp]: "assigns (In1r s) = {}"
+by (simp add: assigns_def)
+
+lemma assigns_stmt_simp1 [simp]: "assigns (\<langle>s::stmt\<rangle>) = {}"
+by (simp add: inj_term_simps)
+
+lemma assigns_var_simp [simp]: "assigns (In2 v) = assignsV v"
+by (simp add: assigns_def)
+
+lemma assigns_var_simp1 [simp]: "assigns (\<langle>v\<rangle>) = assignsV v"
+by (simp add: inj_term_simps)
+
+lemma assigns_expr_list_simp [simp]: "assigns (In3 es) = assignsEs es"
+by (simp add: assigns_def)
+
+lemma assigns_expr_list_simp1 [simp]: "assigns (\<langle>es\<rangle>) = assignsEs es"
+by (simp add: inj_term_simps)
+
+subsection "Analysis of constant expressions"
+
+consts constVal :: "expr \<Rightarrow> val option"
+primrec 
+"constVal (NewC c)      = None"
+"constVal (NewA t e)    = None"
+"constVal (Cast t e)    = None"
+"constVal (Inst e r)    = None"
+"constVal (Lit val)     = Some val"
+"constVal (UnOp unop e) = (case (constVal e) of
+                             None   \<Rightarrow> None
+                           | Some v \<Rightarrow> Some (eval_unop unop v))" 
+"constVal (BinOp binop e1 e2) = (case (constVal e1) of
+                                   None    \<Rightarrow> None
+                                 | Some v1 \<Rightarrow> (case (constVal e2) of 
+                                                None    \<Rightarrow> None
+                                              | Some v2 \<Rightarrow> Some (eval_binop 
+                                                                 binop v1 v2)))"
+"constVal (Super)         = None"
+"constVal (Acc v)         = None"
+"constVal (Ass v e)       = None"
+"constVal (Cond b e1 e2)  = (case (constVal b) of
+                               None   \<Rightarrow> None
+                             | Some bv\<Rightarrow> (case the_Bool bv of
+                                            True \<Rightarrow> (case (constVal e2) of
+                                                       None   \<Rightarrow> None
+                                                     | Some v \<Rightarrow> constVal e1)
+                                          | False\<Rightarrow> (case (constVal e1) of
+                                                       None   \<Rightarrow> None
+                                                     | Some v \<Rightarrow> constVal e2)))"
+--{* Note that @{text "constVal (Cond b e1 e2)"} is stricter as it could be.
+     It requires that all tree expressions are constant even if we can decide
+     which branch to choose, provided the constant value of @{term b} *}
+"constVal (Call accC statT mode objRef mn pTs args) = None"
+"constVal (Methd C sig)   = None" 
+"constVal (Body  C s)     = None"   
+"constVal (InsInitE s e)  = None"  
+"constVal (Callee l e)    = None" 
+
+lemma constVal_Some_induct [consumes 1, case_names Lit UnOp BinOp CondL CondR]: 
+  assumes const: "constVal e = Some v" and
+        hyp_Lit: "\<And> v. P (Lit v)" and
+       hyp_UnOp: "\<And> unop e'. P e' \<Longrightarrow> P (UnOp unop e')" and
+      hyp_BinOp: "\<And> binop e1 e2. \<lbrakk>P e1; P e2\<rbrakk> \<Longrightarrow> P (BinOp binop e1 e2)" and
+      hyp_CondL: "\<And> b bv e1 e2. \<lbrakk>constVal b = Some bv; the_Bool bv; P b; P e1\<rbrakk> 
+                              \<Longrightarrow> P (b? e1 : e2)" and
+      hyp_CondR: "\<And> b bv e1 e2. \<lbrakk>constVal b = Some bv; \<not>the_Bool bv; P b; P e2\<rbrakk>
+                              \<Longrightarrow> P (b? e1 : e2)"
+  shows "P e"
+proof -
+  have "True" and "\<And> v. constVal e = Some v \<Longrightarrow> P e" and "True" and "True"
+  proof (induct "x::var" and e and "s::stmt" and "es::expr list")
+    case Lit
+    show ?case by (rule hyp_Lit)
+  next
+    case UnOp
+    thus ?case
+      by (auto intro: hyp_UnOp)
+  next
+    case BinOp
+    thus ?case
+      by (auto intro: hyp_BinOp)
+  next
+    case (Cond b e1 e2)
+    then obtain v where   v: "constVal (b ? e1 : e2) = Some v"
+      by blast
+    then obtain bv where bv: "constVal b = Some bv"
+      by simp
+    show ?case
+    proof (cases "the_Bool bv")
+      case True
+      with Cond show ?thesis using v bv
+	by (auto intro: hyp_CondL)
+    next
+      case False
+      with Cond show ?thesis using v bv
+	by (auto intro: hyp_CondR)
+    qed
+  qed (simp_all)
+  with const 
+  show ?thesis
+    by blast  
+qed
+
+lemma assignsE_const_simp: "constVal e = Some v \<Longrightarrow> assignsE e = {}"
+  by (induct rule: constVal_Some_induct) simp_all
+
+
+subsection {* Main analysis for boolean expressions *}
+
+text {* Assigned local variables after evaluating the expression if it evaluates
+to a specific boolean value. If the expression cannot evaluate to a 
+@{term Boolean} value UNIV is returned. If we expect true/false the opposite 
+constant false/true will also lead to UNIV. *}
+consts assigns_if:: "bool \<Rightarrow> expr \<Rightarrow> lname set" 
+primrec
+"assigns_if b (NewC c)            = UNIV" --{*can never evaluate to Boolean*} 
+"assigns_if b (NewA t e)          = UNIV" --{*can never evaluate to Boolean*}
+"assigns_if b (Cast t e)          = assigns_if b e" 
+"assigns_if b (Inst e r)          = assignsE e" --{*Inst has type Boolean but
+                                                     e is a reference type*}
+"assigns_if b (Lit val)           = (if val=Bool b then {} else UNIV)"  
+"assigns_if b (UnOp unop e)       = (case constVal (UnOp unop e) of
+                                         None   \<Rightarrow> (if unop = UNot 
+                                                       then assigns_if (\<not>b) e
+                                                       else UNIV)
+                                       | Some v \<Rightarrow> (if v=Bool b 
+                                                       then {} 
+                                                       else UNIV))"
+"assigns_if b (BinOp binop e1 e2) 
+  = (case constVal (BinOp binop e1 e2) of
+       None \<Rightarrow> (if binop=CondAnd then
+                   (case b of 
+                       True  \<Rightarrow> assigns_if True  e1 \<union> assigns_if True e2
+                    |  False \<Rightarrow> assigns_if False e1 \<inter> 
+                                (assigns_if True e1 \<union> assigns_if False e2))
+                else
+               (if binop=CondOr then
+                   (case b of 
+                       True  \<Rightarrow> assigns_if True e1 \<inter> 
+                                (assigns_if False e1 \<union> assigns_if True e2)
+                    |  False \<Rightarrow> assigns_if False  e1 \<union> assigns_if False e2)
+                else assignsE e1 \<union> assignsE e2))
+     | Some v \<Rightarrow> (if v=Bool b then {} else UNIV))"
+
+"assigns_if b (Super)      = UNIV" --{*can never evaluate to Boolean*}
+"assigns_if b (Acc v)      = (assignsV v)"
+"assigns_if b (v := e)     = (assignsE (Ass v e))"
+"assigns_if b (c? e1 : e2) = (assignsE c) \<union>
+                               (case (constVal c) of
+                                  None    \<Rightarrow> (assigns_if b e1) \<inter> 
+                                             (assigns_if b e2)
+                                | Some bv \<Rightarrow> (case the_Bool bv of
+                                                True  \<Rightarrow> assigns_if b e1
+                                              | False \<Rightarrow> assigns_if b e2))"
+"assigns_if b ({accC,statT,mode}objRef\<cdot>mn({pTs}args))  
+          = assignsE ({accC,statT,mode}objRef\<cdot>mn({pTs}args)) "
+-- {* Only dummy analysis for intermediate expressions  
+      @{term Methd}, @{term Body}, @{term InsInitE} and @{term Callee} *}
+"assigns_if b (Methd C sig)   = {}" 
+"assigns_if b (Body  C s)     = {}"   
+"assigns_if b (InsInitE s e)  = {}"  
+"assigns_if b (Callee l e)    = {}" 
+
+lemma assigns_if_const_b_simp:
+  assumes boolConst: "constVal e = Some (Bool b)" (is "?Const b e")
+  shows   "assigns_if b e = {}" (is "?Ass b e")
+proof -
+  have "True" and "\<And> b. ?Const b e \<Longrightarrow> ?Ass b e" and "True" and "True"
+  proof (induct _ and e and _ and _ rule: var_expr_stmt.induct) 
+    case Lit
+    thus ?case by simp
+  next
+    case UnOp 
+    thus ?case by simp
+  next 
+    case (BinOp binop)
+    thus ?case
+      by (cases binop) (simp_all)
+  next
+    case (Cond c e1 e2 b)
+    have hyp_c:  "\<And> b. ?Const b c \<Longrightarrow> ?Ass b c" .
+    have hyp_e1: "\<And> b. ?Const b e1 \<Longrightarrow> ?Ass b e1" .
+    have hyp_e2: "\<And> b. ?Const b e2 \<Longrightarrow> ?Ass b e2" .
+    have const: "constVal (c ? e1 : e2) = Some (Bool b)" .
+    then obtain bv where bv: "constVal c = Some bv"
+      by simp
+    hence emptyC: "assignsE c = {}" by (rule assignsE_const_simp)
+    show ?case
+    proof (cases "the_Bool bv")
+      case True
+      with const bv  
+      have "?Const b e1" by simp
+      hence "?Ass b e1" by (rule hyp_e1)
+      with emptyC bv True
+      show ?thesis
+	by simp
+    next
+      case False
+      with const bv  
+      have "?Const b e2" by simp
+      hence "?Ass b e2" by (rule hyp_e2)
+      with emptyC bv False
+      show ?thesis
+	by simp
+    qed
+  qed (simp_all)
+  with boolConst
+  show ?thesis
+    by blast
+qed
+
+lemma assigns_if_const_not_b_simp:
+  assumes boolConst: "constVal e = Some (Bool b)"        (is "?Const b e")  
+    shows "assigns_if (\<not>b) e = UNIV"                    (is "?Ass b e")
+proof -
+  have True and "\<And> b. ?Const b e \<Longrightarrow> ?Ass b e" and True and True
+  proof (induct _ and e and _ and _ rule: var_expr_stmt.induct) 
+    case Lit
+    thus ?case by simp
+  next
+    case UnOp 
+    thus ?case by simp
+  next 
+    case (BinOp binop)
+    thus ?case
+      by (cases binop) (simp_all)
+  next
+    case (Cond c e1 e2 b)
+    have hyp_c:  "\<And> b. ?Const b c \<Longrightarrow> ?Ass b c" .
+    have hyp_e1: "\<And> b. ?Const b e1 \<Longrightarrow> ?Ass b e1" .
+    have hyp_e2: "\<And> b. ?Const b e2 \<Longrightarrow> ?Ass b e2" .
+    have const: "constVal (c ? e1 : e2) = Some (Bool b)" .
+    then obtain bv where bv: "constVal c = Some bv"
+      by simp
+    show ?case
+    proof (cases "the_Bool bv")
+      case True
+      with const bv  
+      have "?Const b e1" by simp
+      hence "?Ass b e1" by (rule hyp_e1)
+      with bv True
+      show ?thesis
+	by simp
+    next
+      case False
+      with const bv  
+      have "?Const b e2" by simp
+      hence "?Ass b e2" by (rule hyp_e2)
+      with bv False 
+      show ?thesis
+	by simp
+    qed
+  qed (simp_all)
+  with boolConst
+  show ?thesis
+    by blast
+qed
+
+subsection {* Lifting set operations to range of tables (map to a set) *}
+
+constdefs 
+ union_ts:: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
+                    ("_ \<Rightarrow>\<union> _" [67,67] 65)
+ "A \<Rightarrow>\<union> B \<equiv> \<lambda> k. A k \<union> B k"
+
+constdefs
+ intersect_ts:: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
+                    ("_ \<Rightarrow>\<inter>  _" [72,72] 71)
+ "A \<Rightarrow>\<inter>  B \<equiv> \<lambda> k. A k \<inter> B k"
+
+constdefs
+ all_union_ts:: "('a,'b) tables \<Rightarrow> 'b set \<Rightarrow> ('a,'b) tables" 
+                                                     (infixl "\<Rightarrow>\<union>\<^sub>\<forall>" 40)
+"A \<Rightarrow>\<union>\<^sub>\<forall> B \<equiv> \<lambda> k. A k \<union> B"
+  
+subsubsection {* Binary union of tables *}
+
+lemma union_ts_iff [simp]: "(c \<in> (A \<Rightarrow>\<union> B) k) = (c \<in> A k \<or>  c \<in> B k)"
+  by (unfold union_ts_def) blast
+
+lemma union_tsI1 [elim?]: "c \<in> A k \<Longrightarrow> c \<in> (A \<Rightarrow>\<union> B) k"
+  by simp
+
+lemma union_tsI2 [elim?]: "c \<in> B k \<Longrightarrow> c \<in> (A \<Rightarrow>\<union> B) k"
+  by simp
+
+lemma union_tsCI [intro!]: "(c \<notin> B k \<Longrightarrow> c \<in> A k) \<Longrightarrow> c \<in> (A \<Rightarrow>\<union> B) k"
+  by auto
+
+lemma union_tsE [elim!]: 
+ "\<lbrakk>c \<in> (A \<Rightarrow>\<union> B) k; (c \<in> A k \<Longrightarrow> P); (c \<in> B k \<Longrightarrow> P)\<rbrakk> \<Longrightarrow> P"
+  by (unfold union_ts_def) blast
+
+subsubsection {* Binary intersection of tables *}
+
+lemma intersect_ts_iff [simp]: "c \<in> (A \<Rightarrow>\<inter> B) k = (c \<in> A k \<and> c \<in> B k)"
+  by (unfold intersect_ts_def) blast
+
+lemma intersect_tsI [intro!]: "\<lbrakk>c \<in> A k; c \<in> B k\<rbrakk> \<Longrightarrow> c \<in>  (A \<Rightarrow>\<inter> B) k"
+  by simp
+
+lemma intersect_tsD1: "c \<in> (A \<Rightarrow>\<inter> B) k \<Longrightarrow> c \<in> A k"
+  by simp
+
+lemma intersect_tsD2: "c \<in> (A \<Rightarrow>\<inter> B) k \<Longrightarrow> c \<in> B k"
+  by simp
+
+lemma intersect_tsE [elim!]: 
+   "\<lbrakk>c \<in> (A \<Rightarrow>\<inter> B) k; \<lbrakk>c \<in> A k; c \<in> B k\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+  by simp
+
+
+subsubsection {* All-Union of tables and set *}
+
+lemma all_union_ts_iff [simp]: "(c \<in> (A \<Rightarrow>\<union>\<^sub>\<forall> B) k) = (c \<in> A k \<or>  c \<in> B)"
+  by (unfold all_union_ts_def) blast
+
+lemma all_union_tsI1 [elim?]: "c \<in> A k \<Longrightarrow> c \<in> (A \<Rightarrow>\<union>\<^sub>\<forall> B) k"
+  by simp
+
+lemma all_union_tsI2 [elim?]: "c \<in> B \<Longrightarrow> c \<in> (A \<Rightarrow>\<union>\<^sub>\<forall> B) k"
+  by simp
+
+lemma all_union_tsCI [intro!]: "(c \<notin> B \<Longrightarrow> c \<in> A k) \<Longrightarrow> c \<in> (A \<Rightarrow>\<union>\<^sub>\<forall> B) k"
+  by auto
+
+lemma all_union_tsE [elim!]: 
+ "\<lbrakk>c \<in> (A \<Rightarrow>\<union>\<^sub>\<forall> B) k; (c \<in> A k \<Longrightarrow> P); (c \<in> B \<Longrightarrow> P)\<rbrakk> \<Longrightarrow> P"
+  by (unfold all_union_ts_def) blast
+
+
+section "The rules of definite assignment"
+
+ 
+types breakass = "(label, lname) tables" 
+--{* Mapping from a break label, to the set of variables that will be assigned 
+     if the evaluation terminates with this break *}
+    
+record assigned = 
+         nrm :: "lname set" --{* Definetly assigned variables 
+                                 for normal completion*}
+         brk :: "breakass" --{* Definetly assigned variables for 
+                                abnormal completion with a break *}
+
+consts da :: "(env \<times> lname set \<times> term \<times> assigned) set"  
+text {* The environment @{term env} is only needed for the 
+        conditional @{text "_ ? _ : _"}.
+        The definite assignment rules refer to the typing rules here to
+        distinguish boolean and other expressions.
+      *}
+
+syntax
+da :: "env \<Rightarrow> lname set \<Rightarrow> term \<Rightarrow> assigned \<Rightarrow> bool" 
+                           ("_\<turnstile> _ \<guillemotright>_\<guillemotright> _" [65,65,65,65] 71)
+
+translations 
+  "E\<turnstile> B \<guillemotright>t\<guillemotright> A" == "(E,B,t,A) \<in> da"
+
+text {* @{text B}: the ''assigned'' variables before evaluating term @{text t};
+        @{text A}: the ''assigned'' variables after evaluating term @{text t}
+*}
+
+
+constdefs rmlab :: "'a \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
+"rmlab k A \<equiv> \<lambda> x. if x=k then UNIV else A x"
+ 
+(*
+constdefs setbrk :: "breakass \<Rightarrow> assigned \<Rightarrow> breakass set"
+"setbrk b A \<equiv> {b} \<union> {a| a. a\<in> brk A \<and> lab a \<noteq> lab b}"
+*)
+
+constdefs range_inter_ts :: "('a,'b) tables \<Rightarrow> 'b set" ("\<Rightarrow>\<Inter>_" 80)
+ "\<Rightarrow>\<Inter>A \<equiv> {x |x. \<forall> k. x \<in> A k}"
+
+inductive "da" intros
+
+ Skip: "Env\<turnstile> B \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>"
+
+ Expr: "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> A 
+        \<Longrightarrow>  
+        Env\<turnstile> B \<guillemotright>\<langle>Expr e\<rangle>\<guillemotright> A"
+ Lab:  "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>c\<rangle>\<guillemotright> C; nrm A = nrm C \<inter> (brk C) l; brk A = rmlab l (brk C)\<rbrakk>
+        \<Longrightarrow> 
+        Env\<turnstile> B \<guillemotright>\<langle>Break l\<bullet> c\<rangle>\<guillemotright> A" 
+
+ Comp: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1; Env\<turnstile> nrm C1 \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2; 
+        nrm A = nrm C2; brk A = (brk C1) \<Rightarrow>\<inter> (brk C2)\<rbrakk> 
+        \<Longrightarrow>  
+        Env\<turnstile> B \<guillemotright>\<langle>c1;; c2\<rangle>\<guillemotright> A"
+
+ If:   "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E;
+         Env\<turnstile> (B \<union> assigns_if True  e) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1;
+         Env\<turnstile> (B \<union> assigns_if False e) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2;
+         nrm A = nrm C1 \<inter> nrm C2;
+         brk A = brk C1 \<Rightarrow>\<inter> brk C2 \<rbrakk>  
+         \<Longrightarrow>
+         Env\<turnstile> B \<guillemotright>\<langle>If(e) c1 Else c2\<rangle>\<guillemotright> A"
+
+--{* Note that @{term E} is not further used, because we take the specialized
+     sets that also consider if the expression evaluates to true or false. 
+     Inside of @{term e} there is no {\tt break} or {\tt finally}, so the break
+     map of @{term E} will be the trivial one. So 
+     @{term "Env\<turnstile>B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E"} is just used to enshure the definite assignment in
+     expression @{term e}.
+     Notice the implicit analysis of a constant boolean expression @{term e}
+     in this rule. For example, if @{term e} is constantly @{term True} then 
+     @{term "assigns_if False e = UNIV"} and therefor @{term "nrm C2=UNIV"}.
+     So finally @{term "nrm A = nrm C1"}. For the break maps this trick 
+     workd too, because the trival break map will map all labels to 
+     @{term UNIV}. In the example, if no break occurs in @{term c2} the break
+     maps will trivially map to @{term UNIV} and if a break occurs it will map
+     to @{term UNIV} too, because @{term "assigns_if False e = UNIV"}. So
+     in the intersection of the break maps the path @{term c2} will have no
+     contribution.
+  *}
+
+ Loop: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E; 
+         Env\<turnstile> (B \<union> assigns_if True e) \<guillemotright>\<langle>c\<rangle>\<guillemotright> C;
+         nrm A = nrm C \<inter> (B \<union> assigns_if False e);
+         brk A = brk C\<rbrakk>  
+         \<Longrightarrow>
+         Env\<turnstile> B \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A"
+--{* The @{text Loop} rule resembles some of the ideas of the @{text If} rule.
+     For the @{term "nrm A"} the set @{term "B \<union> assigns_if False e"} 
+     will be @{term UNIV} if the condition is constantly true. To normally exit
+     the while loop, we must consider the body @{term c} to be completed 
+     normally (@{term "nrm C"}) or with a break. But in this model, 
+     the label @{term l} of the loop
+     only handles continue labels, not break labels. The break label will be
+     handled by an enclosing @{term Lab} statement. So we don't have to
+     handle the breaks specially. 
+  *}
+
+ Jmp: "\<lbrakk>jump=Ret \<longrightarrow> Result \<in> B;
+        nrm A = UNIV;
+        brk A = (case jump of
+                   Break l \<Rightarrow> \<lambda> k. if k=l then B else UNIV     
+                 | Cont l  \<Rightarrow> \<lambda> k. UNIV
+                 | Ret     \<Rightarrow> \<lambda> k. UNIV)\<rbrakk> 
+       \<Longrightarrow> 
+       Env\<turnstile> B \<guillemotright>\<langle>Jmp jump\<rangle>\<guillemotright> A"
+--{* In case of a break to label @{term l} the corresponding break set is all
+     variables assigned before the break. The assigned variables for normal
+     completion of the @{term Jmp} is @{term UNIV}, because the statement will
+     never complete normally. For continue and return the break map is the 
+     trivial one. In case of a return we enshure that the result value is
+     assigned.
+  *}
+
+ Throw: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E; nrm A = UNIV; brk A = (\<lambda> l. UNIV)\<rbrakk> 
+        \<Longrightarrow> Env\<turnstile> B \<guillemotright>\<langle>Throw e\<rangle>\<guillemotright> A"
+
+ Try:  "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1; 
+         Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>\<turnstile> (B \<union> {VName vn}) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2;  
+         nrm A = nrm C1 \<inter> nrm C2;
+         brk A = brk C1 \<Rightarrow>\<inter> brk C2\<rbrakk> 
+        \<Longrightarrow> Env\<turnstile> B \<guillemotright>\<langle>Try c1 Catch(C vn) c2\<rangle>\<guillemotright> A"
+
+ Fin:  "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1;
+         Env\<turnstile> B \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2;
+         nrm A = nrm C1 \<union> nrm C2;
+         brk A = ((brk C1) \<Rightarrow>\<union>\<^sub>\<forall> (nrm C2)) \<Rightarrow>\<inter> (brk C2)\<rbrakk>  
+         \<Longrightarrow>
+         Env\<turnstile> B \<guillemotright>\<langle>c1 Finally c2\<rangle>\<guillemotright> A" 
+--{* The set of assigned variables before execution @{term c2} are the same
+     as before execution @{term c1}, because @{term c1} could throw an exception
+     and so we can't guarantee that any variable will be assigned in @{term c1}.
+     The @{text Finally} statement completes
+     normally if both @{term c1} and @{term c2} complete normally. If @{term c1}
+     completes abnormally with a break, then @{term c2} also will be executed 
+     and may terminate normally or with a break. The overall break map then is
+     the intersection of the maps of both paths. If @{term c2} terminates 
+     normally we have to extend all break sets in @{term "brk C1"} with 
+     @{term "nrm C2"} (@{text "\<Rightarrow>\<union>\<^sub>\<forall>"}). If @{term c2} exits with a break this
+     break will appear in the overall result state. We don't know if 
+     @{term c1} completed normally or abruptly (maybe with an exception not only
+     a break) so @{term c1} has no contribution to the break map following this
+     path.
+  *}
+
+--{* Evaluation of expressions and the break sets of definite assignment:
+     Thinking of a Java expression we assume that we can never have
+     a break statement inside of a expression. So for all expressions the
+     break sets could be set to the trivial one: @{term "\<lambda> l. UNIV"}. 
+     But we can't
+     trivially proof, that evaluating an expression will never result in a 
+     break, allthough Java expressions allready syntactically don't allow
+     nested stetements in them. The reason are the nested class initialzation 
+     statements which are inserted by the evaluation rules. So to proof the
+     absence of a break we need to ensure, that the initialization statements
+     will never end up in a break. In a wellfromed initialization statement, 
+     of course, were breaks are nested correctly inside of @{term Lab} 
+     or @{term Loop} statements evaluation of the whole initialization 
+     statement will never result in a break, because this break will be 
+     handled inside of the statement. But for simplicity we haven't added
+     the analysis of the correct nesting of breaks in the typing judgments 
+     right now. So we have decided to adjust the rules of definite assignment
+     to fit to these circumstances. If an initialization is involved during
+     evaluation of the expression (evaluation rules @{text FVar}, @{text NewC} 
+     and @{text NewA}
+*}
+
+ Init: "Env\<turnstile> B \<guillemotright>\<langle>Init C\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>"
+--{* Wellformedness of a program will ensure, that every static initialiser 
+     is definetly assigned and the jumps are nested correctly. The case here
+     for @{term Init} is just for convenience, to get a proper precondition 
+     for the induction hypothesis in various proofs, so that we don't have to
+     expand the initialisation on every point where it is triggerred by the
+     evaluation rules.
+  *}   
+ NewC: "Env\<turnstile> B \<guillemotright>\<langle>NewC C\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>" 
+
+ NewA: "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> A 
+        \<Longrightarrow>
+        Env\<turnstile> B \<guillemotright>\<langle>New T[e]\<rangle>\<guillemotright> A"
+
+ Cast: "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> A
+        \<Longrightarrow>
+        Env\<turnstile> B \<guillemotright>\<langle>Cast T e\<rangle>\<guillemotright> A"
+
+ Inst: "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> A 
+        \<Longrightarrow> 
+        Env\<turnstile> B \<guillemotright>\<langle>e InstOf T\<rangle>\<guillemotright> A"
+
+ Lit:  "Env\<turnstile> B \<guillemotright>\<langle>Lit v\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>"
+
+ UnOp: "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> A
+        \<Longrightarrow> 
+        Env\<turnstile> B \<guillemotright>\<langle>UnOp unop e\<rangle>\<guillemotright> A"
+
+ CondAnd: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1; Env\<turnstile> (B \<union> assigns_if True e1) \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2; 
+            nrm A = B \<union> (assigns_if True (BinOp CondAnd e1 e2) \<inter> 
+                            assigns_if False (BinOp CondAnd e1 e2));
+            brk A = (\<lambda> l. UNIV) \<rbrakk>
+           \<Longrightarrow>
+           Env\<turnstile> B \<guillemotright>\<langle>BinOp CondAnd e1 e2\<rangle>\<guillemotright> A"
+
+ CondOr: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1; Env\<turnstile> (B \<union> assigns_if False e1) \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2; 
+           nrm A = B \<union> (assigns_if True (BinOp CondOr e1 e2) \<inter> 
+                             assigns_if False (BinOp CondOr e1 e2));
+           brk A = (\<lambda> l. UNIV) \<rbrakk>
+           \<Longrightarrow>
+           Env\<turnstile> B \<guillemotright>\<langle>BinOp CondOr e1 e2\<rangle>\<guillemotright> A"
+
+ BinOp: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1; Env\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A; 
+          binop \<noteq> CondAnd; binop \<noteq> CondOr\<rbrakk>
+         \<Longrightarrow>
+         Env\<turnstile> B \<guillemotright>\<langle>BinOp binop e1 e2\<rangle>\<guillemotright> A"
+
+ Super: "This \<in> B 
+         \<Longrightarrow> 
+         Env\<turnstile> B \<guillemotright>\<langle>Super\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>"
+
+ AccLVar: "\<lbrakk>vn \<in> B;
+            nrm A = B; brk A = (\<lambda> k. UNIV)\<rbrakk> 
+            \<Longrightarrow> 
+            Env\<turnstile> B \<guillemotright>\<langle>Acc (LVar vn)\<rangle>\<guillemotright> A"
+--{* To properly access a local variable we have to test the definite 
+     assignment here. The variable must occur in the set @{term B} 
+  *}
+
+ Acc: "\<lbrakk>\<forall> vn. v \<noteq> LVar vn;
+        Env\<turnstile> B \<guillemotright>\<langle>v\<rangle>\<guillemotright> A\<rbrakk>
+        \<Longrightarrow>
+        Env\<turnstile> B \<guillemotright>\<langle>Acc v\<rangle>\<guillemotright> A"
+
+ AssLVar: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E; nrm A = nrm E \<union> {vn}; brk A = brk E\<rbrakk> 
+           \<Longrightarrow> 
+           Env\<turnstile> B \<guillemotright>\<langle>(LVar vn) := e\<rangle>\<guillemotright> A"
+
+ Ass: "\<lbrakk>\<forall> vn. v \<noteq> LVar vn; Env\<turnstile> B \<guillemotright>\<langle>v\<rangle>\<guillemotright> V; Env\<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<guillemotright> A\<rbrakk>
+        \<Longrightarrow>
+        Env\<turnstile> B \<guillemotright>\<langle>v := e\<rangle>\<guillemotright> A"
+
+ CondBool: "\<lbrakk>Env\<turnstile>(c ? e1 : e2)\<Colon>-(PrimT Boolean);
+             Env\<turnstile> B \<guillemotright>\<langle>c\<rangle>\<guillemotright> C;
+             Env\<turnstile> (B \<union> assigns_if True  c) \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1;
+             Env\<turnstile> (B \<union> assigns_if False c) \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2;
+             nrm A = B \<union> (assigns_if True  (c ? e1 : e2) \<inter> 
+                              assigns_if False (c ? e1 : e2));
+             brk A = (\<lambda> l. UNIV)\<rbrakk>
+             \<Longrightarrow> 
+             Env\<turnstile> B \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A" 
+
+ Cond: "\<lbrakk>\<not> Env\<turnstile>(c ? e1 : e2)\<Colon>-(PrimT Boolean);
+         Env\<turnstile> B \<guillemotright>\<langle>c\<rangle>\<guillemotright> C;
+         Env\<turnstile> (B \<union> assigns_if True  c) \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1;
+         Env\<turnstile> (B \<union> assigns_if False c) \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2;
+        nrm A = nrm E1 \<inter> nrm E2; brk A = (\<lambda> l. UNIV)\<rbrakk>
+        \<Longrightarrow> 
+        Env\<turnstile> B \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A" 
+
+ Call: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E; Env\<turnstile> nrm E \<guillemotright>\<langle>args\<rangle>\<guillemotright> A\<rbrakk> 
+        \<Longrightarrow>  
+        Env\<turnstile> B \<guillemotright>\<langle>{accC,statT,mode}e\<cdot>mn({pTs}args)\<rangle>\<guillemotright> A"
+
+-- {* The interplay of @{term Call}, @{term Methd} and @{term Body}:
+      Why rules for @{term Methd} and @{term Body} at all? Note that a
+      Java source program will not include bare  @{term Methd} or @{term Body}
+      terms. These terms are just introduced during evaluation. So definite
+      assignment of @{term Call} does not consider @{term Methd} or 
+      @{term Body} at all. So for definite assignment alone we could omit the
+      rules for @{term Methd} and @{term Body}. 
+      But since evaluation of the method invocation is
+      split up into three rules we must ensure that we have enough information
+      about the call even in the @{term Body} term to make sure that we can
+      proof type safety. Also we must be able transport this information 
+      from @{term Call} to @{term Methd} and then further to @{term Body} 
+      during evaluation to establish the definite assignment of @{term Methd}
+      during evaluation of @{term Call}, and of @{term Body} during evaluation
+      of @{term Methd}. This is necessary since definite assignment will be
+      a precondition for each induction hypothesis coming out of the evaluation
+      rules, and therefor we have to establish the definite assignment of the
+      sub-evaluation during the type-safety proof. Note that well-typedness is
+      also a precondition for type-safety and so we can omit some assertion 
+      that are already ensured by well-typedness. 
+   *}
+ Methd: "\<lbrakk>methd (prg Env) D sig = Some m;
+          Env\<turnstile> B \<guillemotright>\<langle>Body (declclass m) (stmt (mbody (mthd m)))\<rangle>\<guillemotright> A
+         \<rbrakk>
+         \<Longrightarrow>
+         Env\<turnstile> B \<guillemotright>\<langle>Methd D sig\<rangle>\<guillemotright> A" 
+
+ Body: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>c\<rangle>\<guillemotright> C; jumpNestingOkS {Ret} c; Result \<in> nrm C;
+         nrm A = B; brk A = (\<lambda> l. UNIV)\<rbrakk>
+        \<Longrightarrow>
+        Env\<turnstile> B \<guillemotright>\<langle>Body D c\<rangle>\<guillemotright> A"
+-- {* Note that @{term A} is not correlated to  @{term C}. If the body
+      statement returns abruptly with return, evaluation of  @{term Body}
+      will absorb this return and complete normally. So we cannot trivially
+      get the assigned variables of the body statement since it has not 
+      completed normally or with a break.
+      If the body completes normally we guarantee that the result variable
+      is set with this rule. But if the body completes abruptly with a return
+      we can't guarantee that the result variable is set here, since 
+      definite assignment only talks about normal completion and breaks. So
+      for a return the @{term Jump} rule ensures that the result variable is
+      set and then this information must be carried over to the @{term Body}
+      rule by the conformance predicate of the state.
+   *}
+ LVar: "Env\<turnstile> B \<guillemotright>\<langle>LVar vn\<rangle>\<guillemotright> \<lparr>nrm=B, brk=\<lambda> l. UNIV\<rparr>" 
+
+ FVar: "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> A 
+        \<Longrightarrow> 
+        Env\<turnstile> B \<guillemotright>\<langle>{accC,statDeclC,stat}e..fn\<rangle>\<guillemotright> A" 
+
+ AVar: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1; Env\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A\<rbrakk>
+         \<Longrightarrow> 
+         Env\<turnstile> B \<guillemotright>\<langle>e1.[e2]\<rangle>\<guillemotright> A" 
+
+ Nil: "Env\<turnstile> B \<guillemotright>\<langle>[]::expr list\<rangle>\<guillemotright> \<lparr>nrm=B, brk=\<lambda> l. UNIV\<rparr>" 
+
+ Cons: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e::expr\<rangle>\<guillemotright> E; Env\<turnstile> nrm E \<guillemotright>\<langle>es\<rangle>\<guillemotright> A\<rbrakk>
+        \<Longrightarrow> 
+        Env\<turnstile> B \<guillemotright>\<langle>e#es\<rangle>\<guillemotright> A" 
+
+
+declare inj_term_sym_simps [simp]
+declare assigns_if.simps [simp del]
+declare split_paired_All [simp del] split_paired_Ex [simp del]
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac"
+*}
+inductive_cases da_elim_cases [cases set]:
+  "Env\<turnstile> B \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>In1r Skip\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>\<langle>Expr e\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1r (Expr e)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>l\<bullet> c\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1r (l\<bullet> c)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>c1;; c2\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1r (c1;; c2)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>If(e) c1 Else c2\<rangle>\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>In1r (If(e) c1 Else c2)\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1r (l\<bullet> While(e) c)\<guillemotright> A"  
+  "Env\<turnstile> B \<guillemotright>\<langle>Jmp jump\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1r (Jmp jump)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>Throw e\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1r (Throw e)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>Try c1 Catch(C vn) c2\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1r (Try c1 Catch(C vn) c2)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>c1 Finally c2\<rangle>\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>In1r (c1 Finally c2)\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>\<langle>Init C\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1r (Init C)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>NewC C\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (NewC C)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>New T[e]\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (New T[e])\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>Cast T e\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (Cast T e)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>e InstOf T\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (e InstOf T)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>Lit v\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (Lit v)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>UnOp unop e\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (UnOp unop e)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>BinOp binop e1 e2\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (BinOp binop e1 e2)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>Super\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (Super)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>Acc v\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (Acc v)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>v := e\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l (v := e)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>In1l (c ? e1 : e2)\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>\<langle>{accC,statT,mode}e\<cdot>mn({pTs}args)\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In1l ({accC,statT,mode}e\<cdot>mn({pTs}args))\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>Methd C sig\<rangle>\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>In1l (Methd C sig)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>Body D c\<rangle>\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>In1l (Body D c)\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>\<langle>LVar vn\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In2 (LVar vn)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>{accC,statDeclC,stat}e..fn\<rangle>\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>In2 ({accC,statDeclC,stat}e..fn)\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>\<langle>e1.[e2]\<rangle>\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>In2 (e1.[e2])\<guillemotright> A" 
+  "Env\<turnstile> B \<guillemotright>\<langle>[]::expr list\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In3 ([]::expr list)\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>\<langle>e#es\<rangle>\<guillemotright> A"
+  "Env\<turnstile> B \<guillemotright>In3 (e#es)\<guillemotright> A"
+declare inj_term_sym_simps [simp del]
+declare assigns_if.simps [simp]
+declare split_paired_All [simp] split_paired_Ex [simp]
+ML_setup {*
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
+*}
+(* To be able to eliminate both the versions with the overloaded brackets: 
+   (B \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> A) and with the explicit constructor (B \<guillemotright>In1r Skip\<guillemotright> A), 
+   every rule appears in both versions
+ *)
+
+lemma da_Skip: "A = \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr> \<Longrightarrow> Env\<turnstile> B \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> A"
+  by (auto intro: da.Skip)
+
+lemma da_NewC: "A = \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr> \<Longrightarrow> Env\<turnstile> B \<guillemotright>\<langle>NewC C\<rangle>\<guillemotright> A"
+  by (auto intro: da.NewC)
+ 
+lemma da_Lit:  "A = \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr> \<Longrightarrow> Env\<turnstile> B \<guillemotright>\<langle>Lit v\<rangle>\<guillemotright> A"
+  by (auto intro: da.Lit)
+
+lemma da_Super: "\<lbrakk>This \<in> B;A = \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>\<rbrakk> \<Longrightarrow> Env\<turnstile> B \<guillemotright>\<langle>Super\<rangle>\<guillemotright> A"
+  by (auto intro: da.Super)
+
+lemma da_Init: "A = \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr> \<Longrightarrow> Env\<turnstile> B \<guillemotright>\<langle>Init C\<rangle>\<guillemotright> A"
+  by (auto intro: da.Init)
+
+
+(*
+For boolean expressions:
+
+The following holds: "assignsE e \<subseteq> assigns_if True e \<inter> assigns_if False e"
+but not vice versa:
+ "assigns_if True e \<inter> assigns_if False e \<subseteq> assignsE e"
+
+Example: 
+ e = ((x < 5) || (y = true)) && (y = true)
+
+   =  (   a    ||    b     ) &&    c
+
+assigns_if True  a = {}
+assigns_if False a = {}
+
+assigns_if True  b = {y}
+assigns_if False b = {y}
+
+assigns_if True  c = {y}
+assigns_if False c = {y}
+
+assigns_if True (a || b) = assigns_if True a \<inter> 
+                                (assigns_if False a \<union> assigns_if True b)
+                           = {} \<inter> ({} \<union> {y}) = {}
+assigns_if False (a || b) = assigns_if False a \<union> assigns_if False b
+                            = {} \<union> {y} = {y}
+
+
+
+assigns_ifE True e =  assigns_if True (a || b) \<union> assigns_if True c
+                    = {} \<union> {y} = {y}
+assigns_ifE False e = assigns_if False (a || b) \<inter> 
+                       (assigns_if True (a || b) \<union> assigns_if False c)
+                     = {y} \<inter> ({} \<union> {y}) = {y}
+
+assignsE e = {}
+*)  
+
+lemma assignsE_subseteq_assigns_ifs:
+ assumes boolEx: "E\<turnstile>e\<Colon>-PrimT Boolean" (is "?Boolean e")
+   shows "assignsE e \<subseteq> assigns_if True e \<inter> assigns_if False e" (is "?Incl e")
+proof -
+  have True and "?Boolean e \<Longrightarrow> ?Incl e" and True and True
+  proof (induct _ and e and _ and _ rule: var_expr_stmt.induct)
+    case (Cast T e)
+    have "E\<turnstile>e\<Colon>- (PrimT Boolean)"
+    proof -
+      have "E\<turnstile>(Cast T e)\<Colon>- (PrimT Boolean)" .
+      then obtain Te where "E\<turnstile>e\<Colon>-Te"
+                           "prg E\<turnstile>Te\<preceq>? PrimT Boolean"
+	by cases simp
+      thus ?thesis
+	by - (drule cast_Boolean2,simp)
+    qed
+    with Cast.hyps
+    show ?case
+      by simp
+  next	
+    case (Lit val) 
+    thus ?case
+      by - (erule wt_elim_cases, cases "val", auto simp add: empty_dt_def)
+  next
+    case (UnOp unop e) 
+    thus ?case
+      by - (erule wt_elim_cases,cases unop,
+            (fastsimp simp add: assignsE_const_simp)+)
+  next
+    case (BinOp binop e1 e2)
+    from BinOp.prems obtain e1T e2T
+      where "E\<turnstile>e1\<Colon>-e1T" and "E\<turnstile>e2\<Colon>-e2T" and "wt_binop (prg E) binop e1T e2T"
+            and "(binop_type binop) = Boolean"
+      by (elim wt_elim_cases) simp
+    with BinOp.hyps
+    show ?case
+      by - (cases binop, auto simp add: assignsE_const_simp)
+  next
+    case (Cond c e1 e2)
+    have hyp_c: "?Boolean c \<Longrightarrow> ?Incl c" .
+    have hyp_e1: "?Boolean e1 \<Longrightarrow> ?Incl e1" .
+    have hyp_e2: "?Boolean e2 \<Longrightarrow> ?Incl e2" .
+    have wt: "E\<turnstile>(c ? e1 : e2)\<Colon>-PrimT Boolean" .
+    then obtain
+      boolean_c:  "E\<turnstile>c\<Colon>-PrimT Boolean" and
+      boolean_e1: "E\<turnstile>e1\<Colon>-PrimT Boolean" and
+      boolean_e2: "E\<turnstile>e2\<Colon>-PrimT Boolean"
+      by (elim wt_elim_cases) (auto dest: widen_Boolean2)
+    show ?case
+    proof (cases "constVal c")
+      case None
+      with boolean_e1 boolean_e2
+      show ?thesis
+	using hyp_e1 hyp_e2 
+	by (auto)
+    next
+      case (Some bv)
+      show ?thesis
+      proof (cases "the_Bool bv")
+	case True
+	with Some show ?thesis using hyp_e1 boolean_e1 by auto
+      next
+	case False
+	with Some show ?thesis using hyp_e2 boolean_e2 by auto
+      qed
+    qed
+  qed simp_all
+  with boolEx 
+  show ?thesis
+    by blast
+qed
+  
+
+(* Trick:
+   If you have a rule with the abstract term injections:
+   e.g:  da.Skip "B \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> A"
+   and the current goal state as an concrete injection:
+   e.g: "B \<guillemotright>In1r Skip\<guillemotright> A"
+   you can convert the rule by: da.Skip [simplified]
+   if inj_term_simps is in the simpset
+
+*)
+
+lemma rmlab_same_label [simp]: "(rmlab l A) l = UNIV"
+  by (simp add: rmlab_def)
+
+lemma rmlab_same_label1 [simp]: "l=l' \<Longrightarrow> (rmlab l A) l' = UNIV"
+  by (simp add: rmlab_def)
+
+lemma rmlab_other_label [simp]: "l\<noteq>l'\<Longrightarrow> (rmlab l A) l' = A l'"
+  by (auto simp add: rmlab_def)
+
+lemma range_inter_ts_subseteq [intro]: "\<forall> k. A k  \<subseteq> B k \<Longrightarrow>  \<Rightarrow>\<Inter>A \<subseteq> \<Rightarrow>\<Inter>B"
+  by (auto simp add: range_inter_ts_def)
+
+lemma range_inter_ts_subseteq': 
+  "\<lbrakk>\<forall> k. A k  \<subseteq> B k; x \<in> \<Rightarrow>\<Inter>A\<rbrakk> \<Longrightarrow> x \<in> \<Rightarrow>\<Inter>B"
+  by (auto simp add: range_inter_ts_def)
+
+lemma da_monotone: 
+      assumes      da: "Env\<turnstile> B  \<guillemotright>t\<guillemotright> A"   and
+        subseteq_B_B': "B \<subseteq> B'"          and
+                  da': "Env\<turnstile> B' \<guillemotright>t\<guillemotright> A'" 
+  shows "(nrm A \<subseteq> nrm A') \<and> (\<forall> l. (brk A l \<subseteq> brk A' l))"
+proof -
+  from da
+  show "\<And> B' A'. \<lbrakk>Env\<turnstile> B' \<guillemotright>t\<guillemotright> A'; B \<subseteq> B'\<rbrakk> 
+                  \<Longrightarrow> (nrm A \<subseteq> nrm A') \<and> (\<forall> l. (brk A l \<subseteq> brk A' l))"
+       (is "PROP ?Hyp Env B t A")  
+  proof (induct)
+    case Skip 
+    from Skip.prems Skip.hyps 
+    show ?case by cases simp
+  next
+    case Expr 
+    from Expr.prems Expr.hyps 
+    show ?case by cases simp
+  next
+    case (Lab A B C Env c l B' A')
+    have A: "nrm A = nrm C \<inter> brk C l" "brk A = rmlab l (brk C)" .
+    have "PROP ?Hyp Env B \<langle>c\<rangle> C" .
+    moreover
+    have "B \<subseteq> B'" .
+    moreover
+    obtain C'
+      where "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
+        and A': "nrm A' = nrm C' \<inter> brk C' l" "brk A' = rmlab l (brk C')"
+      using Lab.prems
+      by - (erule da_elim_cases,simp)
+    ultimately
+    have "nrm C \<subseteq> nrm C'" and hyp_brk: "(\<forall>l. brk C l \<subseteq> brk C' l)" by auto
+    then 
+    have "nrm C \<inter> brk C l \<subseteq> nrm C' \<inter> brk C' l" by auto
+    moreover
+    {
+      fix l'
+      from hyp_brk
+      have "rmlab l (brk C) l'  \<subseteq> rmlab l (brk C') l'"
+	by  (cases "l=l'") simp_all
+    }
+    moreover note A A'
+    ultimately show ?case
+      by simp
+  next
+    case (Comp A B C1 C2 Env c1 c2 B' A')
+    have A: "nrm A = nrm C2" "brk A = brk C1 \<Rightarrow>\<inter>  brk C2" .
+    have "Env\<turnstile> B' \<guillemotright>\<langle>c1;; c2\<rangle>\<guillemotright> A'" .
+    then obtain  C1' C2'
+      where da_c1: "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
+            da_c2: "Env\<turnstile> nrm C1' \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'"  and
+            A': "nrm A' = nrm C2'" "brk A' = brk C1' \<Rightarrow>\<inter>  brk C2'"
+      by (rule da_elim_cases) auto
+    have "PROP ?Hyp Env B \<langle>c1\<rangle> C1" .
+    moreover have "B \<subseteq> B'" .
+    moreover note da_c1
+    ultimately have C1': "nrm C1 \<subseteq> nrm C1'" "(\<forall>l. brk C1 l \<subseteq> brk C1' l)"
+      by (auto)
+    have "PROP ?Hyp Env (nrm C1) \<langle>c2\<rangle> C2" .
+    with da_c2 C1' 
+    have C2': "nrm C2 \<subseteq> nrm C2'" "(\<forall>l. brk C2 l \<subseteq> brk C2' l)"
+      by (auto)
+    with A A' C1'
+    show ?case
+      by auto
+  next
+    case (If A B C1 C2 E Env c1 c2 e B' A')
+    have A: "nrm A = nrm C1 \<inter> nrm C2" "brk A = brk C1 \<Rightarrow>\<inter>  brk C2" .
+    have "Env\<turnstile> B' \<guillemotright>\<langle>If(e) c1 Else c2\<rangle>\<guillemotright> A'" .
+    then obtain C1' C2'
+      where da_c1: "Env\<turnstile> B' \<union> assigns_if True e \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
+            da_c2: "Env\<turnstile> B' \<union> assigns_if False e \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
+               A': "nrm A' = nrm C1' \<inter> nrm C2'" "brk A' = brk C1' \<Rightarrow>\<inter>  brk C2'"
+      by (rule da_elim_cases) auto
+    have "PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c1\<rangle> C1" .
+    moreover have B': "B \<subseteq> B'" .
+    moreover note da_c1 
+    ultimately obtain C1': "nrm C1 \<subseteq> nrm C1'" "(\<forall>l. brk C1 l \<subseteq> brk C1' l)"
+      by blast
+    have "PROP ?Hyp Env (B \<union> assigns_if False e) \<langle>c2\<rangle> C2" .
+    with da_c2 B'
+    obtain C2': "nrm C2 \<subseteq> nrm C2'" "(\<forall>l. brk C2 l \<subseteq> brk C2' l)"
+      by blast
+    with A A' C1'
+    show ?case
+      by auto
+  next
+    case (Loop A B C E Env c e l B' A')
+    have A: "nrm A = nrm C \<inter> (B \<union> assigns_if False e)"
+            "brk A = brk C" .
+    have "Env\<turnstile> B' \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A'" .
+    then obtain C'
+      where 
+       da_c': "Env\<turnstile> B' \<union> assigns_if True e \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'" and
+          A': "nrm A' = nrm C' \<inter> (B' \<union> assigns_if False e)"
+              "brk A' = brk C'" 
+      by (rule da_elim_cases) auto
+    have "PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c\<rangle> C" .
+    moreover have B': "B \<subseteq> B'" .
+    moreover note da_c'
+    ultimately obtain C': "nrm C \<subseteq> nrm C'" "(\<forall>l. brk C l \<subseteq> brk C' l)"
+      by blast
+    with A A' B'
+    have "nrm A \<subseteq> nrm A'"
+      by blast
+    moreover
+    { fix l'
+      have  "brk A l' \<subseteq> brk A' l'"
+      proof (cases "constVal e")
+	case None
+	with A A' C' 
+	show ?thesis
+	   by (cases "l=l'") auto
+      next
+	case (Some bv)
+	with A A' C'
+	show ?thesis
+	  by (cases "the_Bool bv", cases "l=l'") auto
+      qed
+    }
+    ultimately show ?case
+      by auto
+  next
+    case (Jmp A B Env jump B' A')
+    thus ?case by (elim da_elim_cases) (auto split: jump.splits)
+  next
+    case Throw thus ?case by -  (erule da_elim_cases, auto)
+  next
+    case (Try A B C C1 C2 Env c1 c2 vn B' A')
+    have A: "nrm A = nrm C1 \<inter> nrm C2"
+            "brk A = brk C1 \<Rightarrow>\<inter>  brk C2" .
+    have "Env\<turnstile> B' \<guillemotright>\<langle>Try c1 Catch(C vn) c2\<rangle>\<guillemotright> A'" .
+    then obtain C1' C2'
+      where da_c1': "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
+            da_c2': "Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>\<turnstile> B' \<union> {VName vn} 
+                      \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
+            A': "nrm A' = nrm C1' \<inter> nrm C2'"
+                "brk A' = brk C1' \<Rightarrow>\<inter>  brk C2'" 
+      by (rule da_elim_cases) auto
+    have "PROP ?Hyp Env B \<langle>c1\<rangle> C1" .
+    moreover have B': "B \<subseteq> B'" .
+    moreover note da_c1'
+    ultimately obtain C1': "nrm C1 \<subseteq> nrm C1'" "(\<forall>l. brk C1 l \<subseteq> brk C1' l)"
+      by blast
+    have "PROP ?Hyp (Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>)
+                    (B \<union> {VName vn}) \<langle>c2\<rangle> C2" .
+    with B' da_c2'
+    obtain "nrm C2 \<subseteq> nrm C2'" "(\<forall>l. brk C2 l \<subseteq> brk C2' l)"
+      by blast
+    with C1' A A'
+    show ?case
+      by auto
+  next
+    case (Fin A B C1 C2 Env c1 c2 B' A')
+    have A: "nrm A = nrm C1 \<union> nrm C2"
+            "brk A = (brk C1 \<Rightarrow>\<union>\<^sub>\<forall> nrm C2) \<Rightarrow>\<inter> (brk C2)" .
+    have "Env\<turnstile> B' \<guillemotright>\<langle>c1 Finally c2\<rangle>\<guillemotright> A'" .
+    then obtain C1' C2'
+      where  da_c1': "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
+             da_c2': "Env\<turnstile> B' \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
+             A':  "nrm A' = nrm C1' \<union> nrm C2'"
+                  "brk A' = (brk C1' \<Rightarrow>\<union>\<^sub>\<forall> nrm C2') \<Rightarrow>\<inter> (brk C2')"
+      by (rule da_elim_cases) auto
+    have "PROP ?Hyp Env B \<langle>c1\<rangle> C1" .
+    moreover have B': "B \<subseteq> B'" .
+    moreover note da_c1'
+    ultimately obtain C1': "nrm C1 \<subseteq> nrm C1'" "(\<forall>l. brk C1 l \<subseteq> brk C1' l)"
+      by blast
+    have hyp_c2: "PROP ?Hyp Env B \<langle>c2\<rangle> C2" .
+    from da_c2' B' 
+     obtain "nrm C2 \<subseteq> nrm C2'" "(\<forall>l. brk C2 l \<subseteq> brk C2' l)"
+       by - (drule hyp_c2,auto)
+     with A A' C1'
+     show ?case
+       by auto
+   next
+     case Init thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case NewC thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case NewA thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case Cast thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case Inst thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case Lit thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case UnOp thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case (CondAnd A B E1 E2 Env e1 e2 B' A')
+     have A: "nrm A = B \<union>
+                       assigns_if True (BinOp CondAnd e1 e2) \<inter>
+                       assigns_if False (BinOp CondAnd e1 e2)"
+             "brk A = (\<lambda>l. UNIV)" .
+     have "Env\<turnstile> B' \<guillemotright>\<langle>BinOp CondAnd e1 e2\<rangle>\<guillemotright> A'" .
+     then obtain  A': "nrm A' = B' \<union>
+                                 assigns_if True (BinOp CondAnd e1 e2) \<inter>
+                                 assigns_if False (BinOp CondAnd e1 e2)"
+                      "brk A' = (\<lambda>l. UNIV)" 
+       by (rule da_elim_cases) auto
+     have B': "B \<subseteq> B'" .
+     with A A' show ?case 
+       by auto 
+   next
+     case CondOr thus ?case by - (erule da_elim_cases, auto)
+   next
+     case BinOp thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case Super thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case AccLVar thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case Acc thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case AssLVar thus ?case by - (erule da_elim_cases, auto)
+   next
+     case Ass thus ?case by -  (erule da_elim_cases, auto)
+   next
+     case (CondBool A B C E1 E2 Env c e1 e2 B' A')
+     have A: "nrm A = B \<union> 
+                        assigns_if True (c ? e1 : e2) \<inter> 
+                        assigns_if False (c ? e1 : e2)"
+             "brk A = (\<lambda>l. UNIV)" .
+     have "Env\<turnstile> (c ? e1 : e2)\<Colon>- (PrimT Boolean)" .
+     moreover
+     have "Env\<turnstile> B' \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A'" .
+     ultimately
+     obtain A': "nrm A' = B' \<union> 
+                                  assigns_if True (c ? e1 : e2) \<inter> 
+                                  assigns_if False (c ? e1 : e2)"
+                     "brk A' = (\<lambda>l. UNIV)"
+       by - (erule da_elim_cases,auto simp add: inj_term_simps) 
+       (* inj_term_simps needed to handle wt (defined without \<langle>\<rangle>) *)
+     have B': "B \<subseteq> B'" .
+     with A A' show ?case 
+       by auto 
+   next
+     case (Cond A B C E1 E2 Env c e1 e2 B' A')  
+     have A: "nrm A = nrm E1 \<inter> nrm E2"
+             "brk A = (\<lambda>l. UNIV)" .
+     have not_bool: "\<not> Env\<turnstile> (c ? e1 : e2)\<Colon>- (PrimT Boolean)" .
+     have "Env\<turnstile> B' \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A'" .
+     then obtain E1' E2'
+       where da_e1': "Env\<turnstile> B' \<union> assigns_if True c \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'" and
+             da_e2': "Env\<turnstile> B' \<union> assigns_if False c \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2'" and
+                 A': "nrm A' = nrm E1' \<inter> nrm E2'"
+                     "brk A' = (\<lambda>l. UNIV)"
+       using not_bool
+       by  - (erule da_elim_cases, auto simp add: inj_term_simps)
+       (* inj_term_simps needed to handle wt (defined without \<langle>\<rangle>) *)
+     have "PROP ?Hyp Env (B \<union> assigns_if True c) \<langle>e1\<rangle> E1" .
+     moreover have B': "B \<subseteq> B'" .
+     moreover note da_e1'
+     ultimately obtain E1': "nrm E1 \<subseteq> nrm E1'" "(\<forall>l. brk E1 l \<subseteq> brk E1' l)"
+      by blast
+     have "PROP ?Hyp Env (B \<union> assigns_if False c) \<langle>e2\<rangle> E2" .
+     with B' da_e2'
+     obtain "nrm E2 \<subseteq> nrm E2'" "(\<forall>l. brk E2 l \<subseteq> brk E2' l)"
+       by blast
+    with E1' A A'
+    show ?case
+      by auto
+  next
+    case Call
+    from Call.prems and Call.hyps
+    show ?case by cases auto
+  next
+    case Methd thus ?case by -  (erule da_elim_cases, auto)
+  next
+    case Body thus ?case by -  (erule da_elim_cases, auto)
+  next
+    case LVar thus ?case by -  (erule da_elim_cases, auto)
+  next
+    case FVar thus ?case by -  (erule da_elim_cases, auto)
+  next
+    case AVar thus ?case by -  (erule da_elim_cases, auto)
+  next
+    case Nil thus ?case by -  (erule da_elim_cases, auto)
+  next
+    case Cons thus ?case by -  (erule da_elim_cases, auto)
+  qed
+qed
+  
+lemma da_weaken:     
+      assumes            da: "Env\<turnstile> B \<guillemotright>t\<guillemotright> A" and
+              subseteq_B_B': "B \<subseteq> B'" 
+        shows "\<exists> A'. Env \<turnstile> B' \<guillemotright>t\<guillemotright> A'"
+proof -
+  note assigned.select_convs [CPure.intro]
+  from da  
+  show "\<And> B'. B \<subseteq> B' \<Longrightarrow> \<exists> A'. Env\<turnstile> B' \<guillemotright>t\<guillemotright> A'" (is "PROP ?Hyp Env B t")
+  proof (induct) 
+    case Skip thus ?case by (rules intro: da.Skip)
+  next
+    case Expr thus ?case by (rules intro: da.Expr)
+  next
+    case (Lab A B C Env c l B')  
+    have "PROP ?Hyp Env B \<langle>c\<rangle>" .
+    moreover
+    have B': "B \<subseteq> B'" .
+    ultimately obtain C' where "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
+      by rules
+    then obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>Break l\<bullet> c\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Lab)
+    thus ?case ..
+  next
+    case (Comp A B C1 C2 Env c1 c2 B')
+    have da_c1: "Env\<turnstile> B \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1" .
+    have "PROP ?Hyp Env B \<langle>c1\<rangle>" .
+    moreover
+    have B': "B \<subseteq> B'" .
+    ultimately obtain C1' where da_c1': "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'"
+      by rules
+    with da_c1 B'
+    have
+      "nrm C1 \<subseteq> nrm C1'"
+      by (rule da_monotone [elim_format]) simp
+    moreover
+    have "PROP ?Hyp Env (nrm C1) \<langle>c2\<rangle>" .
+    ultimately obtain C2' where "Env\<turnstile> nrm C1' \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'"
+      by rules
+    with da_c1' obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>c1;; c2\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Comp)
+    thus ?case ..
+  next
+    case (If A B C1 C2 E Env c1 c2 e B')
+    have B': "B \<subseteq> B'" .
+    obtain  E' where "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>e\<rangle>" by (rule If.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain C1' where "Env\<turnstile> (B' \<union> assigns_if True e) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'"
+    proof -
+      from B'
+      have "(B \<union> assigns_if True e) \<subseteq> (B' \<union> assigns_if True e)"
+	by blast
+      moreover
+      have "PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c1\<rangle>" by (rule If.hyps)
+      ultimately 
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain C2' where "Env\<turnstile> (B' \<union> assigns_if False e) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'"
+    proof - 
+      from B' have "(B \<union> assigns_if False e) \<subseteq> (B' \<union> assigns_if False e)"
+	by blast
+      moreover
+      have "PROP ?Hyp Env (B \<union> assigns_if False e) \<langle>c2\<rangle>" by (rule If.hyps)
+      ultimately
+      show ?thesis using that by rules
+    qed
+    ultimately
+    obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>If(e) c1 Else c2\<rangle>\<guillemotright> A'"
+      by (rules intro: da.If)
+    thus ?case ..
+  next  
+    case (Loop A B C E Env c e l B')
+    have B': "B \<subseteq> B'" .
+    obtain  E' where "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>e\<rangle>" by (rule Loop.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain C' where "Env\<turnstile> (B' \<union> assigns_if True e) \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
+    proof -
+      from B'
+      have "(B \<union> assigns_if True e) \<subseteq> (B' \<union> assigns_if True e)"
+	by blast
+      moreover
+      have "PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c\<rangle>" by (rule Loop.hyps)
+      ultimately 
+      show ?thesis using that by rules
+    qed
+    ultimately
+    obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Loop )
+    thus ?case ..
+  next
+    case (Jmp A B Env jump B') 
+    have B': "B \<subseteq> B'" .
+    with Jmp.hyps have "jump = Ret \<longrightarrow> Result \<in> B' "
+      by auto
+    moreover
+    obtain A'::assigned 
+              where  "nrm A' = UNIV"
+                     "brk A' = (case jump of 
+                                  Break l \<Rightarrow> \<lambda>k. if k = l then B' else UNIV 
+                                | Cont l \<Rightarrow> \<lambda>k. UNIV
+                                | Ret \<Rightarrow> \<lambda>k. UNIV)"
+                     
+      by  rules
+    ultimately have "Env\<turnstile> B' \<guillemotright>\<langle>Jmp jump\<rangle>\<guillemotright> A'"
+      by (rule da.Jmp)
+    thus ?case ..
+  next
+    case Throw thus ?case by (rules intro: da.Throw )
+  next
+    case (Try A B C C1 C2 Env c1 c2 vn B')
+    have B': "B \<subseteq> B'" .
+    obtain C1' where "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>c1\<rangle>" by (rule Try.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain C2' where 
+      "Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>\<turnstile> B' \<union> {VName vn} \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'"
+    proof -
+      from B' have "B \<union> {VName vn} \<subseteq> B' \<union> {VName vn}" by blast
+      moreover
+      have "PROP ?Hyp (Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>) 
+                      (B \<union> {VName vn}) \<langle>c2\<rangle>" 
+	by (rule Try.hyps)
+      ultimately
+      show ?thesis using that by rules
+    qed
+    ultimately
+    obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>Try c1 Catch(C vn) c2\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Try )
+    thus ?case ..
+  next
+    case (Fin A B C1 C2 Env c1 c2 B')
+    have B': "B \<subseteq> B'" .
+    obtain C1' where C1': "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>c1\<rangle>" by (rule Fin.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain C2' where "Env\<turnstile> B' \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>c2\<rangle>" by (rule Fin.hyps)
+      with B'
+      show ?thesis using that by rules
+    qed
+    ultimately
+    obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>c1 Finally c2\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Fin )
+    thus ?case ..
+  next
+    case Init thus ?case by (rules intro: da.Init)
+  next
+    case NewC thus ?case by (rules intro: da.NewC)
+  next
+    case NewA thus ?case by (rules intro: da.NewA)
+  next
+    case Cast thus ?case by (rules intro: da.Cast)
+  next
+    case Inst thus ?case by (rules intro: da.Inst)
+  next
+    case Lit thus ?case by (rules intro: da.Lit)
+  next
+    case UnOp thus ?case by (rules intro: da.UnOp)
+  next
+    case (CondAnd A B E1 E2 Env e1 e2 B')
+    have B': "B \<subseteq> B'" .
+    obtain E1' where "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>e1\<rangle>" by (rule CondAnd.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain E2' where "Env\<turnstile> B' \<union>  assigns_if True e1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2'"
+    proof -
+      from B' have "B \<union> assigns_if True e1 \<subseteq> B' \<union>  assigns_if True e1"
+	by blast
+      moreover
+      have "PROP ?Hyp Env (B \<union> assigns_if True e1) \<langle>e2\<rangle>" by (rule CondAnd.hyps)
+      ultimately show ?thesis using that by rules
+    qed
+    ultimately
+    obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>BinOp CondAnd e1 e2\<rangle>\<guillemotright> A'"
+      by (rules intro: da.CondAnd)
+    thus ?case ..
+  next
+    case (CondOr A B E1 E2 Env e1 e2 B')
+    have B': "B \<subseteq> B'" .
+    obtain E1' where "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>e1\<rangle>" by (rule CondOr.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain E2' where "Env\<turnstile> B' \<union> assigns_if False e1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2'"
+    proof -
+      from B' have "B \<union> assigns_if False e1 \<subseteq> B' \<union>  assigns_if False e1"
+	by blast
+      moreover
+      have "PROP ?Hyp Env (B \<union> assigns_if False e1) \<langle>e2\<rangle>" by (rule CondOr.hyps)
+      ultimately show ?thesis using that by rules
+    qed
+    ultimately
+    obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>BinOp CondOr e1 e2\<rangle>\<guillemotright> A'"
+      by (rules intro: da.CondOr)
+    thus ?case ..
+  next
+    case (BinOp A B E1 Env binop e1 e2 B')
+    have B': "B \<subseteq> B'" .
+    obtain E1' where E1': "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>e1\<rangle>" by (rule BinOp.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain A' where "Env\<turnstile> nrm E1' \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A'"
+    proof -
+      have "Env\<turnstile> B \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1" by (rule BinOp.hyps)
+      from this B' E1'
+      have "nrm E1 \<subseteq> nrm E1'"
+	by (rule da_monotone [THEN conjE])
+      moreover 
+      have "PROP ?Hyp Env (nrm E1) \<langle>e2\<rangle>" by (rule BinOp.hyps)
+      ultimately show ?thesis using that by rules
+    qed
+    ultimately
+    have "Env\<turnstile> B' \<guillemotright>\<langle>BinOp binop e1 e2\<rangle>\<guillemotright> A'"
+      using BinOp.hyps by (rules intro: da.BinOp)
+    thus ?case ..
+  next
+    case (Super B Env B')
+    have B': "B \<subseteq> B'" .
+    with Super.hyps have "This \<in> B' "
+      by auto
+    thus ?case by (rules intro: da.Super)
+  next
+    case (AccLVar A B Env vn B')
+    have "vn \<in> B" .
+    moreover
+    have "B \<subseteq> B'" .
+    ultimately have "vn \<in> B'" by auto
+    thus ?case by (rules intro: da.AccLVar)
+  next
+    case Acc thus ?case by (rules intro: da.Acc)
+  next 
+    case (AssLVar A B E Env e vn B')
+    have B': "B \<subseteq> B'" .
+    then obtain E' where "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
+      by (rule AssLVar.hyps [elim_format]) rules
+    then obtain A' where  
+      "Env\<turnstile> B' \<guillemotright>\<langle>LVar vn:=e\<rangle>\<guillemotright> A'"
+      by (rules intro: da.AssLVar)
+    thus ?case ..
+  next
+    case (Ass A B Env V e v B') 
+    have B': "B \<subseteq> B'" .
+    have "\<forall>vn. v \<noteq> LVar vn".
+    moreover
+    obtain V' where V': "Env\<turnstile> B' \<guillemotright>\<langle>v\<rangle>\<guillemotright> V'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>v\<rangle>" by (rule Ass.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain A' where "Env\<turnstile> nrm V' \<guillemotright>\<langle>e\<rangle>\<guillemotright> A'"
+    proof -
+      have "Env\<turnstile> B \<guillemotright>\<langle>v\<rangle>\<guillemotright> V" by (rule Ass.hyps)
+      from this B' V'
+      have "nrm V \<subseteq> nrm V'"
+	by (rule da_monotone [THEN conjE])
+      moreover 
+      have "PROP ?Hyp Env (nrm V) \<langle>e\<rangle>" by (rule Ass.hyps)
+      ultimately show ?thesis using that by rules
+    qed  
+    ultimately
+    have "Env\<turnstile> B' \<guillemotright>\<langle>v := e\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Ass)
+    thus ?case ..
+  next
+    case (CondBool A B C E1 E2 Env c e1 e2 B')
+    have B': "B \<subseteq> B'" .
+    have "Env\<turnstile>(c ? e1 : e2)\<Colon>-(PrimT Boolean)" .
+    moreover obtain C' where C': "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>c\<rangle>" by (rule CondBool.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover 
+    obtain E1' where "Env\<turnstile> B' \<union> assigns_if True c \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
+    proof -
+      from B'
+      have "(B \<union> assigns_if True c) \<subseteq> (B' \<union> assigns_if True c)"
+	by blast
+      moreover
+      have "PROP ?Hyp Env (B \<union> assigns_if True c) \<langle>e1\<rangle>" by (rule CondBool.hyps)
+      ultimately 
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain E2' where "Env\<turnstile> B' \<union> assigns_if False c \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2'"
+    proof -
+      from B'
+      have "(B \<union> assigns_if False c) \<subseteq> (B' \<union> assigns_if False c)"
+	by blast
+      moreover
+      have "PROP ?Hyp Env (B \<union> assigns_if False c) \<langle>e2\<rangle>" by(rule CondBool.hyps)
+      ultimately 
+      show ?thesis using that by rules
+    qed
+    ultimately 
+    obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A'"
+      by (rules intro: da.CondBool)
+    thus ?case ..
+  next
+    case (Cond A B C E1 E2 Env c e1 e2 B')
+    have B': "B \<subseteq> B'" .
+    have "\<not> Env\<turnstile>(c ? e1 : e2)\<Colon>-(PrimT Boolean)" .
+    moreover obtain C' where C': "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>c\<rangle>" by (rule Cond.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover 
+    obtain E1' where "Env\<turnstile> B' \<union> assigns_if True c \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
+    proof -
+      from B'
+      have "(B \<union> assigns_if True c) \<subseteq> (B' \<union> assigns_if True c)"
+	by blast
+      moreover
+      have "PROP ?Hyp Env (B \<union> assigns_if True c) \<langle>e1\<rangle>" by (rule Cond.hyps)
+      ultimately 
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain E2' where "Env\<turnstile> B' \<union> assigns_if False c \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2'"
+    proof -
+      from B'
+      have "(B \<union> assigns_if False c) \<subseteq> (B' \<union> assigns_if False c)"
+	by blast
+      moreover
+      have "PROP ?Hyp Env (B \<union> assigns_if False c) \<langle>e2\<rangle>" by (rule Cond.hyps)
+      ultimately 
+      show ?thesis using that by rules
+    qed
+    ultimately 
+    obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Cond)
+    thus ?case ..
+  next
+    case (Call A B E Env accC args e mn mode pTs statT B')
+    have B': "B \<subseteq> B'" .
+    obtain E' where E': "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>e\<rangle>" by (rule Call.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain A' where "Env\<turnstile> nrm E' \<guillemotright>\<langle>args\<rangle>\<guillemotright> A'"
+    proof -
+      have "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E" by (rule Call.hyps)
+      from this B' E'
+      have "nrm E \<subseteq> nrm E'"
+	by (rule da_monotone [THEN conjE])
+      moreover 
+      have "PROP ?Hyp Env (nrm E) \<langle>args\<rangle>" by (rule Call.hyps)
+      ultimately show ?thesis using that by rules
+    qed  
+    ultimately
+    have "Env\<turnstile> B' \<guillemotright>\<langle>{accC,statT,mode}e\<cdot>mn( {pTs}args)\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Call)
+    thus ?case ..
+  next
+    case Methd thus ?case by (rules intro: da.Methd)
+  next
+    case (Body A B C D Env c B')  
+    have B': "B \<subseteq> B'" .
+    obtain C' where C': "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'" and nrm_C': "nrm C \<subseteq> nrm C'"
+    proof -
+      have "Env\<turnstile> B \<guillemotright>\<langle>c\<rangle>\<guillemotright> C" by (rule Body.hyps)
+      moreover note B'
+      moreover
+      from B' obtain C' where da_c: "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
+	by (rule Body.hyps [elim_format]) blast
+      ultimately
+      have "nrm C \<subseteq> nrm C'"
+	by (rule da_monotone [THEN conjE])
+      with da_c that show ?thesis by rules
+    qed
+    moreover 
+    have "Result \<in> nrm C" .
+    with nrm_C' have "Result \<in> nrm C'"
+      by blast
+    moreover have "jumpNestingOkS {Ret} c" .
+    ultimately obtain A' where
+      "Env\<turnstile> B' \<guillemotright>\<langle>Body D c\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Body)
+    thus ?case ..
+  next
+    case LVar thus ?case by (rules intro: da.LVar)
+  next
+    case FVar thus ?case by (rules intro: da.FVar)
+  next
+    case (AVar A B E1 Env e1 e2 B')
+    have B': "B \<subseteq> B'" .
+    obtain E1' where E1': "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>e1\<rangle>" by (rule AVar.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain A' where "Env\<turnstile> nrm E1' \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A'"
+    proof -
+      have "Env\<turnstile> B \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1" by (rule AVar.hyps)
+      from this B' E1'
+      have "nrm E1 \<subseteq> nrm E1'"
+	by (rule da_monotone [THEN conjE])
+      moreover 
+      have "PROP ?Hyp Env (nrm E1) \<langle>e2\<rangle>" by (rule AVar.hyps)
+      ultimately show ?thesis using that by rules
+    qed  
+    ultimately
+    have "Env\<turnstile> B' \<guillemotright>\<langle>e1.[e2]\<rangle>\<guillemotright> A'"
+      by (rules intro: da.AVar)
+    thus ?case ..
+  next
+    case Nil thus ?case by (rules intro: da.Nil)
+  next
+    case (Cons A B E Env e es B')
+    have B': "B \<subseteq> B'" .
+    obtain E' where E': "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
+    proof -
+      have "PROP ?Hyp Env B \<langle>e\<rangle>" by (rule Cons.hyps)
+      with B'  
+      show ?thesis using that by rules
+    qed
+    moreover
+    obtain A' where "Env\<turnstile> nrm E' \<guillemotright>\<langle>es\<rangle>\<guillemotright> A'"
+    proof -
+      have "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E" by (rule Cons.hyps)
+      from this B' E'
+      have "nrm E \<subseteq> nrm E'"
+	by (rule da_monotone [THEN conjE])
+      moreover 
+      have "PROP ?Hyp Env (nrm E) \<langle>es\<rangle>" by (rule Cons.hyps)
+      ultimately show ?thesis using that by rules
+    qed  
+    ultimately
+    have "Env\<turnstile> B' \<guillemotright>\<langle>e # es\<rangle>\<guillemotright> A'"
+      by (rules intro: da.Cons)
+    thus ?case ..
+  qed
+qed
+
+(* Remarks about the proof style:
+
+   "by (rule <Case>.hyps)" vs "."
+   --------------------------
+  
+   with <Case>.hyps you state more precise were the rule comes from
+
+   . takes all assumptions into account, but looks more "light"
+   and is more resistent for cut and paste proof in different 
+   cases.
+
+  "intro: da.intros" vs "da.<Case>"
+  ---------------------------------
+  The first ist more convinient for cut and paste between cases,
+  the second is more informativ for the reader
+*)
+
+corollary da_weakenE [consumes 2]:
+  assumes          da: "Env\<turnstile> B  \<guillemotright>t\<guillemotright> A"   and
+                   B': "B \<subseteq> B'"          and
+              ex_mono: "\<And> A'.  \<lbrakk>Env\<turnstile> B' \<guillemotright>t\<guillemotright> A'; nrm A \<subseteq> nrm A'; 
+                               \<And> l. brk A l \<subseteq> brk A' l\<rbrakk> \<Longrightarrow> P" 
+  shows "P"
+proof -
+  from da B'
+  obtain A' where A': "Env\<turnstile> B' \<guillemotright>t\<guillemotright> A'"
+    by (rule da_weaken [elim_format]) rules
+  with da B'
+  have "nrm A \<subseteq> nrm A' \<and> (\<forall> l. brk A l \<subseteq> brk A' l)"
+    by (rule da_monotone)
+  with A' ex_mono
+  show ?thesis
+    by rules
+qed
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Bali/DefiniteAssignmentCorrect.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -0,0 +1,4517 @@
+header {* Correctness of Definite Assignment *}
+
+theory DefiniteAssignmentCorrect = WellForm + Eval:
+
+ML {*
+Delsimprocs [wt_expr_proc,wt_var_proc,wt_exprs_proc,wt_stmt_proc]
+*}
+
+lemma sxalloc_no_jump:
+  assumes sxalloc: "G\<turnstile>s0 \<midarrow>sxalloc\<rightarrow> s1" and
+           no_jmp: "abrupt s0 \<noteq> Some (Jump j)"
+   shows "abrupt s1 \<noteq> Some (Jump j)"
+using sxalloc no_jmp
+by cases simp_all
+
+lemma sxalloc_no_jump':
+  assumes sxalloc: "G\<turnstile>s0 \<midarrow>sxalloc\<rightarrow> s1" and
+          jump:  "abrupt s1 = Some (Jump j)"
+ shows "abrupt s0 = Some (Jump j)"
+using sxalloc jump
+by cases simp_all
+
+lemma halloc_no_jump:
+  assumes halloc: "G\<turnstile>s0 \<midarrow>halloc oi\<succ>a\<rightarrow> s1" and
+           no_jmp: "abrupt s0 \<noteq> Some (Jump j)"
+   shows "abrupt s1 \<noteq> Some (Jump j)"
+using halloc no_jmp
+by cases simp_all
+
+lemma halloc_no_jump':
+  assumes halloc: "G\<turnstile>s0 \<midarrow>halloc oi\<succ>a\<rightarrow> s1" and
+          jump:  "abrupt s1 = Some (Jump j)"
+  shows "abrupt s0 = Some (Jump j)"
+using halloc jump
+by cases simp_all
+
+lemma Body_no_jump: 
+   assumes eval: "G\<turnstile>s0 \<midarrow>Body D c-\<succ>v\<rightarrow>s1" and
+           jump: "abrupt s0 \<noteq> Some (Jump j)"
+   shows "abrupt s1 \<noteq> Some (Jump j)"
+proof (cases "normal s0")
+  case True
+  with eval obtain s0' where eval': "G\<turnstile>Norm s0' \<midarrow>Body D c-\<succ>v\<rightarrow>s1" and
+                             s0: "s0 = Norm s0'"
+    by (cases s0) simp
+  from eval' obtain s2 where
+     s1: "s1 = abupd (absorb Ret)
+             (if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or>
+                     abrupt s2 = Some (Jump (Cont l))
+              then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)"
+    by cases simp
+  show ?thesis
+  proof (cases "\<exists>l. abrupt s2 = Some (Jump (Break l)) \<or> 
+                   abrupt s2 = Some (Jump (Cont l))")
+    case True
+    with s1 have "abrupt s1 = Some (Error CrossMethodJump)"
+      by (cases s2) simp
+    thus ?thesis by simp
+  next
+    case False
+    with s1 have "s1=abupd (absorb Ret) s2"
+      by simp
+    with False show ?thesis
+      by (cases s2,cases j) (auto simp add: absorb_def)
+  qed
+next
+  case False
+  with eval obtain s0' abr where "G\<turnstile>(Some abr,s0') \<midarrow>Body D c-\<succ>v\<rightarrow>s1"
+                                 "s0 = (Some abr, s0')"
+    by (cases s0) fastsimp
+  with this jump
+  show ?thesis
+    by (cases) (simp)
+qed
+
+lemma Methd_no_jump: 
+   assumes eval: "G\<turnstile>s0 \<midarrow>Methd D sig-\<succ>v\<rightarrow> s1" and
+           jump: "abrupt s0 \<noteq> Some (Jump j)"
+   shows "abrupt s1 \<noteq> Some (Jump j)"
+proof (cases "normal s0")
+  case True
+   with eval obtain s0' where "G\<turnstile>Norm s0' \<midarrow>Methd D sig-\<succ>v\<rightarrow>s1" 
+                              "s0 = Norm s0'"
+    by (cases s0) simp
+  then obtain D' body where "G\<turnstile>s0 \<midarrow>Body D' body-\<succ>v\<rightarrow>s1"
+    by (cases) (simp add: body_def2)
+  from this jump
+  show ?thesis
+    by (rule Body_no_jump)
+next
+  case False
+  with eval obtain s0' abr where "G\<turnstile>(Some abr,s0') \<midarrow>Methd D sig-\<succ>v\<rightarrow>s1"
+                                 "s0 = (Some abr, s0')"
+    by (cases s0) fastsimp
+  with this jump
+  show ?thesis
+    by (cases) (simp)
+qed
+
+lemma jumpNestingOkS_mono: 
+  assumes jumpNestingOk_l': "jumpNestingOkS jmps' c" 
+      and      subset: "jmps' \<subseteq> jmps"
+ shows "jumpNestingOkS jmps c"
+proof -
+    have True and True and 
+       "\<And> jmps' jmps. \<lbrakk>jumpNestingOkS jmps' c; jmps' \<subseteq> jmps\<rbrakk> \<Longrightarrow> jumpNestingOkS jmps c" 
+       and True
+  proof (induct rule: var_expr_stmt.induct)
+    case (Lab j c jmps' jmps)
+    have jmpOk: "jumpNestingOkS jmps' (j\<bullet> c)" .
+    have jmps: "jmps' \<subseteq> jmps" .
+    with jmpOk have "jumpNestingOkS ({j} \<union> jmps') c" by simp
+    moreover from jmps have "({j} \<union> jmps') \<subseteq> ({j} \<union> jmps)" by auto
+    ultimately
+    have "jumpNestingOkS ({j} \<union> jmps) c"
+      by (rule Lab.hyps)
+    thus ?case 
+      by simp
+  next
+    case (Jmp j jmps' jmps) 
+    thus ?case 
+      by (cases j) auto
+  next
+    case (Comp c1 c2 jmps' jmps)
+    from Comp.prems 
+    have "jumpNestingOkS jmps c1" by - (rule Comp.hyps,auto)
+    moreover from Comp.prems
+    have "jumpNestingOkS jmps c2" by - (rule Comp.hyps,auto)
+    ultimately show ?case
+      by simp
+  next
+    case (If_ e c1 c2 jmps' jmps)
+    from If_.prems 
+    have "jumpNestingOkS jmps c1" by - (rule If_.hyps,auto)
+    moreover from If_.prems 
+    have "jumpNestingOkS jmps c2" by - (rule If_.hyps,auto)
+    ultimately show ?case
+      by simp
+  next
+    case (Loop l e c jmps' jmps)
+    have "jumpNestingOkS jmps' (l\<bullet> While(e) c)" .
+    hence "jumpNestingOkS ({Cont l} \<union> jmps') c" by simp
+    moreover have "jmps' \<subseteq> jmps" .
+    hence "{Cont l} \<union> jmps'  \<subseteq> {Cont l} \<union> jmps" by auto
+    ultimately
+    have "jumpNestingOkS ({Cont l} \<union> jmps) c"
+      by (rule Loop.hyps)
+    thus ?case by simp
+  next
+    case (TryC c1 C vn c2 jmps' jmps)
+    from TryC.prems 
+    have "jumpNestingOkS jmps c1" by - (rule TryC.hyps,auto)
+    moreover from TryC.prems 
+    have "jumpNestingOkS jmps c2" by - (rule TryC.hyps,auto)
+    ultimately show ?case
+      by simp
+  next
+    case (Fin c1 c2 jmps' jmps)
+    from Fin.prems 
+    have "jumpNestingOkS jmps c1" by - (rule Fin.hyps,auto)
+    moreover from Fin.prems 
+    have "jumpNestingOkS jmps c2" by - (rule Fin.hyps,auto)
+    ultimately show ?case
+      by simp
+  qed (simp_all)
+  with jumpNestingOk_l' subset
+  show ?thesis
+    by rules
+qed
+   
+corollary jumpNestingOk_mono: 
+  assumes jmpOk: "jumpNestingOk jmps' t" 
+     and subset: "jmps' \<subseteq> jmps"
+  shows  "jumpNestingOk jmps t"
+proof (cases t)
+  case (In1 expr_stmt)
+  show ?thesis
+  proof (cases expr_stmt)
+    case (Inl e)
+    with In1 show ?thesis by simp
+  next
+    case (Inr s)
+    with In1 jmpOk subset show ?thesis by (auto intro: jumpNestingOkS_mono)
+  qed
+qed (simp_all)
+
+lemma assign_abrupt_propagation: 
+ assumes f_ok: "abrupt (f n s) \<noteq> x"
+   and    ass: "abrupt (assign f n s) = x"
+  shows "abrupt s = x"
+proof (cases x)
+  case None
+  with ass show ?thesis
+    by (cases s) (simp add: assign_def Let_def)
+next
+  case (Some xcpt)
+  from f_ok
+  obtain xf sf where "f n s = (xf,sf)"
+    by (cases "f n s")
+  with Some ass f_ok show ?thesis
+    by (cases s) (simp add: assign_def Let_def)
+qed
+ 
+lemma wt_init_comp_ty': 
+"is_acc_type (prg Env) (pid (cls Env)) T \<Longrightarrow> Env\<turnstile>init_comp_ty T\<Colon>\<surd>"
+apply (unfold init_comp_ty_def)
+apply (clarsimp simp add: accessible_in_RefT_simp 
+                          is_acc_type_def is_acc_class_def)
+done
+
+lemma fvar_upd_no_jump: 
+      assumes upd: "upd = snd (fst (fvar statDeclC stat fn a s'))"
+        and noJmp: "abrupt s \<noteq> Some (Jump j)"
+        shows "abrupt (upd val s) \<noteq> Some (Jump j)"
+proof (cases "stat")
+  case True
+  with noJmp upd
+  show ?thesis
+    by (cases s) (simp add: fvar_def2)
+next
+  case False
+  with noJmp upd
+  show ?thesis
+    by (cases s) (simp add: fvar_def2)
+qed
+
+
+lemma avar_state_no_jump: 
+  assumes jmp: "abrupt (snd (avar G i a s)) = Some (Jump j)"
+  shows "abrupt s = Some (Jump j)"
+proof (cases "normal s")
+  case True with jmp show ?thesis by (auto simp add: avar_def2 abrupt_if_def)
+next
+  case False with jmp show ?thesis by (auto simp add: avar_def2 abrupt_if_def)
+qed
+
+lemma avar_upd_no_jump: 
+      assumes upd: "upd = snd (fst (avar G i a s'))"
+        and noJmp: "abrupt s \<noteq> Some (Jump j)"
+        shows "abrupt (upd val s) \<noteq> Some (Jump j)"
+using upd noJmp
+by (cases s) (simp add: avar_def2 abrupt_if_def)
+
+
+text {* 
+The next theorem expresses: If jumps (breaks, continues, returns) are nested
+correctly, we won't find an unexpected jump in the result state of the 
+evaluation. For exeample, a break can't leave its enclosing loop, an return
+cant leave its enclosing method.
+To proove this, the method call is critical. Allthough the
+wellformedness of the whole program guarantees that the jumps (breaks,continues
+and returns) are nested
+correctly in all method bodies, the call rule alone does not guarantee that I
+will call a method or even a class that is part of the program due to dynamic
+binding! To be able to enshure this we need a kind of conformance of the
+state, like in the typesafety proof. But then we will redo the typesafty
+proof here. It would be nice if we could find an easy precondition that will
+guarantee that all calls will actually call classes and methods of the current
+program, which can be instantiated in the typesafty proof later on.
+To fix this problem, I have instrumented the semantic definition of a call
+to filter out any breaks in the state and to throw an error instead. 
+
+To get an induction hypothesis which is strong enough to perform the
+proof, we can't just 
+assume @{term jumpNestingOk} for the empty set and conlcude, that no jump at 
+all will be in the resulting state,  because the set is altered by 
+the statements @{term Lab} and @{term While}. 
+
+The wellformedness of the program is used to enshure that for all
+classinitialisations and methods the nesting of jumps is wellformed, too.
+*}  
+theorem jumpNestingOk_eval:
+  assumes eval: "G\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
+     and jmpOk: "jumpNestingOk jmps t" 
+     and wt: "Env\<turnstile>t\<Colon>T" 
+     and wf: "wf_prog G"
+     and  G: "prg Env = G"
+     and no_jmp: "\<forall> j. abrupt s0 = Some (Jump j) \<longrightarrow> j \<in> jmps"
+                    (is "?Jmp jmps s0")
+  shows  "?Jmp jmps s1 \<and> 
+                 (normal s1 \<longrightarrow>
+                  (\<forall> w upd. v=In2 (w,upd)  
+                   \<longrightarrow>   (\<forall> s j val. 
+                          abrupt s \<noteq> Some (Jump j) \<longrightarrow>
+                          abrupt (upd val s) \<noteq> Some (Jump j))))"
+        (is "?Jmp jmps s1 \<and> ?Upd v s1")  
+proof -
+  let ?HypObj = "\<lambda> t s0 s1 v.
+       (\<forall> jmps T Env. 
+          ?Jmp jmps s0 \<longrightarrow> jumpNestingOk jmps t \<longrightarrow> Env\<turnstile>t\<Colon>T \<longrightarrow> prg Env=G\<longrightarrow>
+          ?Jmp jmps s1 \<and> ?Upd v s1)"
+  -- {* Variable @{text ?HypObj} is the following goal spelled in terms of
+        the object logic, instead of the meta logic. It is needed in some
+        cases of the induction were, the atomize-rulify process of induct 
+        does not work fine, because the eval rules mix up object and meta
+        logic. See for example the case for the loop. *} 
+  from eval 
+  have "\<And> jmps T Env. \<lbrakk>?Jmp jmps s0; jumpNestingOk jmps t; Env\<turnstile>t\<Colon>T;prg Env=G\<rbrakk>
+            \<Longrightarrow> ?Jmp jmps s1 \<and> ?Upd v s1" 
+        (is "PROP ?Hyp t s0 s1 v")
+  -- {* We need to abstract over @{term jmps} since @{term jmps} are extended
+        during analysis of @{term Lab}. Also we need to abstract over 
+        @{term T} and @{term Env} since they are altered in various
+        typing judgements. *}    
+  proof (induct)   
+    case Abrupt thus ?case by simp 
+  next
+    case Skip thus ?case by simp
+  next
+    case Expr thus ?case by (elim wt_elim_cases) simp
+  next
+    case (Lab c jmp s0 s1 jmps T Env) 
+    have jmpOK: "jumpNestingOk jmps (In1r (jmp\<bullet> c))" .
+    have G: "prg Env = G" .
+    have wt_c: "Env\<turnstile>c\<Colon>\<surd>" 
+      using Lab.prems by (elim wt_elim_cases)
+    { 
+      fix j
+      assume ab_s1: "abrupt (abupd (absorb jmp) s1) = Some (Jump j)"
+      have "j\<in>jmps"          
+      proof -
+        from ab_s1 have jmp_s1: "abrupt s1 = Some (Jump j)"
+          by (cases s1) (simp add: absorb_def)
+        have hyp_c: "PROP ?Hyp (In1r c) (Norm s0) s1 \<diamondsuit>" .
+        from ab_s1 have "j \<noteq> jmp" 
+          by (cases s1) (simp add: absorb_def)
+        moreover have "j \<in> {jmp} \<union> jmps"
+        proof -
+          from jmpOK 
+          have "jumpNestingOk ({jmp} \<union> jmps) (In1r c)" by simp
+          with wt_c jmp_s1 G hyp_c
+          show ?thesis
+            by - (rule hyp_c [THEN conjunct1,rule_format],simp)
+        qed
+        ultimately show ?thesis
+          by simp
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Comp c1 c2 s0 s1 s2 jmps T Env)
+    have jmpOk: "jumpNestingOk jmps (In1r (c1;; c2))" .
+    have G: "prg Env = G" .
+    from Comp.prems obtain 
+      wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and wt_c2: "Env\<turnstile>c2\<Colon>\<surd>"
+      by (elim wt_elim_cases)
+    {
+      fix j
+      assume abr_s2: "abrupt s2 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have jmp: "?Jmp jmps s1"
+        proof -
+          have hyp_c1: "PROP ?Hyp (In1r c1) (Norm s0) s1 \<diamondsuit>" .
+          with wt_c1 jmpOk G 
+          show ?thesis by simp
+        qed
+        moreover have hyp_c2: "PROP ?Hyp (In1r c2) s1 s2 (\<diamondsuit>::vals)" .
+        have jmpOk': "jumpNestingOk jmps (In1r c2)" using jmpOk by simp
+        moreover note wt_c2 G abr_s2
+        ultimately show "j \<in> jmps"
+          by (rule hyp_c2 [THEN conjunct1,rule_format (no_asm)])
+      qed
+    } thus ?case by simp
+  next
+    case (If b c1 c2 e s0 s1 s2 jmps T Env)
+    have jmpOk: "jumpNestingOk jmps (In1r (If(e) c1 Else c2))" .
+    have G: "prg Env = G" .
+    from If.prems obtain 
+              wt_e: "Env\<turnstile>e\<Colon>-PrimT Boolean" and 
+      wt_then_else: "Env\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
+      by (elim wt_elim_cases) simp
+    { 
+      fix j
+      assume jmp: "abrupt s2 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have "PROP ?Hyp (In1l e) (Norm s0) s1 (In1 b)" .
+        with wt_e G have "?Jmp jmps s1" 
+          by simp
+        moreover have hyp_then_else: 
+          "PROP ?Hyp (In1r (if the_Bool b then c1 else c2)) s1 s2 \<diamondsuit>" .
+        have "jumpNestingOk jmps (In1r (if the_Bool b then c1 else c2))"
+          using jmpOk by (cases "the_Bool b") simp_all
+        moreover note wt_then_else G jmp
+        ultimately show "j\<in> jmps" 
+          by (rule hyp_then_else [THEN conjunct1,rule_format (no_asm)])
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Loop b c e l s0 s1 s2 s3 jmps T Env)
+    have jmpOk: "jumpNestingOk jmps (In1r (l\<bullet> While(e) c))" .
+    have G: "prg Env = G" .
+    have wt: "Env\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T" .
+    then obtain 
+              wt_e: "Env\<turnstile>e\<Colon>-PrimT Boolean" and 
+              wt_c: "Env\<turnstile>c\<Colon>\<surd>"
+      by (elim wt_elim_cases)
+    {
+      fix j
+      assume jmp: "abrupt s3 = Some (Jump j)" 
+      have "j\<in>jmps"
+      proof -
+        have "PROP ?Hyp (In1l e) (Norm s0) s1 (In1 b)" .
+        with wt_e G have jmp_s1: "?Jmp jmps s1" 
+          by simp
+        show ?thesis
+	proof (cases "the_Bool b")
+          case False
+          from Loop.hyps
+          have "s3=s1"
+            by (simp (no_asm_use) only: if_False False) 
+          with jmp_s1 jmp have "j \<in> jmps" by simp
+          thus ?thesis by simp
+        next
+          case True
+          from Loop.hyps 
+            (* Because of the mixture of object and pure logic in the rule 
+               eval.If the atomization-rulification of the induct method 
+               leaves the hypothesis in object logic *)
+          have "?HypObj (In1r c) s1 s2 (\<diamondsuit>::vals)"
+            apply (simp (no_asm_use) only: True if_True )
+            apply (erule conjE)+
+            apply assumption
+            done
+          note hyp_c = this [rule_format (no_asm)]
+          moreover from jmpOk have "jumpNestingOk ({Cont l} \<union> jmps) (In1r c)"
+            by simp
+          moreover from jmp_s1 have "?Jmp ({Cont l} \<union> jmps) s1" by simp
+          ultimately have jmp_s2: "?Jmp ({Cont l} \<union> jmps) s2" 
+            using wt_c G by rules
+          have "?Jmp jmps (abupd (absorb (Cont l)) s2)"
+          proof -
+            {
+      	      fix j'
+      	      assume abs: "abrupt (abupd (absorb (Cont l)) s2)=Some (Jump j')"
+      	      have "j' \<in> jmps"
+      	      proof (cases "j' = Cont l")
+      		case True
+      		with abs show ?thesis
+      		  by (cases s2) (simp add: absorb_def)
+      	      next
+      		case False 
+      		with abs have "abrupt s2 = Some (Jump j')"
+      		  by (cases s2) (simp add: absorb_def) 
+      		with jmp_s2 False show ?thesis
+      		  by simp
+      	      qed
+            }
+            thus ?thesis by simp
+          qed
+          moreover
+          from Loop.hyps
+          have "?HypObj (In1r (l\<bullet> While(e) c)) 
+                        (abupd (absorb (Cont l)) s2) s3 (\<diamondsuit>::vals)"
+            apply (simp (no_asm_use) only: True if_True)
+            apply (erule conjE)+
+            apply assumption
+            done
+          note hyp_w = this [rule_format (no_asm)]
+          note jmpOk wt G jmp 
+          ultimately show "j\<in> jmps" 
+            by (rule hyp_w [THEN conjunct1,rule_format (no_asm)])
+        qed
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Jmp j s jmps T Env) thus ?case by simp
+  next
+    case (Throw a e s0 s1 jmps T Env)
+    have jmpOk: "jumpNestingOk jmps (In1r (Throw e))" .
+    have G: "prg Env = G" .
+    from Throw.prems obtain Te where 
+      wt_e: "Env\<turnstile>e\<Colon>-Te" 
+      by (elim wt_elim_cases)
+    {
+      fix j
+      assume jmp: "abrupt (abupd (throw a) s1) = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have "PROP ?Hyp (In1l e) (Norm s0) s1 (In1 a)" .
+        hence "?Jmp jmps s1" using wt_e G by simp
+        moreover
+        from jmp 
+        have "abrupt s1 = Some (Jump j)"
+          by (cases s1) (simp add: throw_def abrupt_if_def)
+        ultimately show "j \<in> jmps" by simp
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Try C c1 c2 s0 s1 s2 s3 vn jmps T Env)
+    have jmpOk: "jumpNestingOk jmps (In1r (Try c1 Catch(C vn) c2))" .
+    have G: "prg Env = G" .
+    from Try.prems obtain 
+      wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and  
+      wt_c2: "Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>\<turnstile>c2\<Colon>\<surd>"
+      by (elim wt_elim_cases)
+    {
+      fix j
+      assume jmp: "abrupt s3 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have "PROP ?Hyp (In1r c1) (Norm s0) s1 (\<diamondsuit>::vals)" .
+        with jmpOk wt_c1 G
+        have jmp_s1: "?Jmp jmps s1" by simp
+        have s2: "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" .
+        show "j \<in> jmps"
+        proof (cases "G,s2\<turnstile>catch C")
+          case False
+          from Try.hyps have "s3=s2" 
+	    by (simp (no_asm_use) only: False if_False)
+          with jmp have "abrupt s1 = Some (Jump j)"
+            using sxalloc_no_jump' [OF s2] by simp
+          with jmp_s1 
+          show ?thesis by simp
+        next
+          case True
+          with Try.hyps 
+          have "?HypObj (In1r c2) (new_xcpt_var vn s2) s3 (\<diamondsuit>::vals)"
+            apply (simp (no_asm_use) only: True if_True simp_thms)
+            apply (erule conjE)+
+            apply assumption
+            done
+          note hyp_c2 = this [rule_format (no_asm)]
+          from jmp_s1 sxalloc_no_jump' [OF s2] 
+          have "?Jmp jmps s2"
+            by simp
+          hence "?Jmp jmps (new_xcpt_var vn s2)"
+            by (cases s2) simp
+          moreover have "jumpNestingOk jmps (In1r c2)" using jmpOk by simp
+          moreover note wt_c2
+          moreover from G 
+          have "prg (Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>) = G"
+            by simp
+          moreover note jmp
+          ultimately show ?thesis 
+            by (rule hyp_c2 [THEN conjunct1,rule_format (no_asm)])
+        qed
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Fin c1 c2 s0 s1 s2 s3 x1 jmps T Env)
+    have jmpOk: " jumpNestingOk jmps (In1r (c1 Finally c2))" .
+    have G: "prg Env = G" .
+    from Fin.prems obtain 
+      wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and wt_c2: "Env\<turnstile>c2\<Colon>\<surd>"
+      by (elim wt_elim_cases)
+    {
+      fix j
+      assume jmp: "abrupt s3 = Some (Jump j)" 
+      have "j \<in> jmps"
+      proof (cases "x1=Some (Jump j)")
+        case True
+        have hyp_c1: "PROP ?Hyp (In1r c1) (Norm s0) (x1,s1) \<diamondsuit>" .
+        with True jmpOk wt_c1 G show ?thesis 
+          by - (rule hyp_c1 [THEN conjunct1,rule_format (no_asm)],simp_all)
+      next
+        case False
+        have hyp_c2:  "PROP ?Hyp (In1r c2) (Norm s1) s2 \<diamondsuit>" .
+        have "s3 = (if \<exists>err. x1 = Some (Error err) then (x1, s1)
+                             else abupd (abrupt_if (x1 \<noteq> None) x1) s2)" .
+        with False jmp have "abrupt s2 = Some (Jump j)"
+          by (cases s2, simp add: abrupt_if_def)
+        with jmpOk wt_c2 G show ?thesis 
+          by - (rule hyp_c2 [THEN conjunct1,rule_format (no_asm)],simp_all)
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Init C c s0 s1 s2 s3 jmps T Env)
+    have "jumpNestingOk jmps (In1r (Init C))".
+    have G: "prg Env = G" .
+    have "the (class G C) = c" .
+    with Init.prems have c: "class G C = Some c"
+      by (elim wt_elim_cases) auto
+    {
+      fix j
+      assume jmp: "abrupt s3 = (Some (Jump j))" 
+      have "j\<in>jmps"
+      proof (cases "inited C (globs s0)") 
+        case True
+        with Init.hyps have "s3=Norm s0"
+          by simp
+        with jmp
+        have "False" by simp thus ?thesis ..
+      next
+        case False
+        from wf c G
+        have wf_cdecl: "wf_cdecl G (C,c)"
+          by (simp add: wf_prog_cdecl)
+        from Init.hyps
+        have "?HypObj (In1r (if C = Object then Skip else Init (super c))) 
+                            (Norm ((init_class_obj G C) s0)) s1 (\<diamondsuit>::vals)"
+          apply (simp (no_asm_use) only: False if_False simp_thms)
+          apply (erule conjE)+
+          apply assumption
+          done
+        note hyp_s1 = this [rule_format (no_asm)]
+        from wf_cdecl G have
+          wt_super: "Env\<turnstile>(if C = Object then Skip else Init (super c))\<Colon>\<surd>"
+          by (cases "C=Object")
+             (auto dest: wf_cdecl_supD is_acc_classD) 
+        from hyp_s1 [OF _ _ wt_super G]
+        have "?Jmp jmps s1" 
+          by simp
+        hence jmp_s1: "?Jmp jmps ((set_lvars empty) s1)" by (cases s1) simp
+        from False Init.hyps
+        have "?HypObj (In1r (init c)) ((set_lvars empty) s1) s2 (\<diamondsuit>::vals)" 
+          apply (simp (no_asm_use) only: False if_False simp_thms)
+          apply (erule conjE)+
+          apply assumption
+          done
+        note hyp_init_c = this [rule_format (no_asm)] 
+        from wf_cdecl 
+        have wt_init_c: "\<lparr>prg = G, cls = C, lcl = empty\<rparr>\<turnstile>init c\<Colon>\<surd>"
+          by (rule wf_cdecl_wt_init)
+        from wf_cdecl have "jumpNestingOkS {} (init c)" 
+          by (cases rule: wf_cdeclE)
+        hence "jumpNestingOkS jmps (init c)"
+          by (rule jumpNestingOkS_mono) simp
+        moreover 
+        have "abrupt s2 = Some (Jump j)"
+        proof -
+          from False Init.hyps 
+          have "s3 = (set_lvars (locals (store s1))) s2"  by simp
+          with jmp show ?thesis by (cases s2) simp
+        qed
+        ultimately show ?thesis
+          using hyp_init_c [OF jmp_s1 _ wt_init_c]
+          by simp
+      qed
+    }
+    thus ?case by simp
+  next
+    case (NewC C a s0 s1 s2 jmps T Env)
+    {
+      fix j
+      assume jmp: "abrupt s2 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof - 
+        have "prg Env = G" .
+        moreover have hyp_init: "PROP ?Hyp (In1r (Init C)) (Norm s0) s1 \<diamondsuit>" .
+        moreover from wf NewC.prems 
+        have "Env\<turnstile>(Init C)\<Colon>\<surd>"
+          by (elim wt_elim_cases) (drule is_acc_classD,simp)
+        moreover 
+        have "abrupt s1 = Some (Jump j)"
+        proof -
+          have "G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2" .
+          from this jmp show ?thesis
+            by (rule halloc_no_jump')
+        qed
+        ultimately show "j \<in> jmps" 
+          by - (rule hyp_init [THEN conjunct1,rule_format (no_asm)],auto)
+      qed
+    }
+    thus ?case by simp
+  next
+    case (NewA elT a e i s0 s1 s2 s3 jmps T Env)
+    {
+      fix j
+      assume jmp: "abrupt s3 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have G: "prg Env = G" .
+        from NewA.prems 
+        obtain wt_init: "Env\<turnstile>init_comp_ty elT\<Colon>\<surd>" and 
+               wt_size: "Env\<turnstile>e\<Colon>-PrimT Integer"
+          by (elim wt_elim_cases) (auto dest:  wt_init_comp_ty')
+        have "PROP ?Hyp (In1r (init_comp_ty elT)) (Norm s0) s1 \<diamondsuit>" .
+        with wt_init G 
+        have "?Jmp jmps s1" 
+          by (simp add: init_comp_ty_def)
+        moreover
+        have hyp_e: "PROP ?Hyp (In1l e) s1 s2 (In1 i)" .
+        have "abrupt s2 = Some (Jump j)"
+        proof -
+          have "G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr elT (the_Intg i)\<succ>a\<rightarrow> s3".
+          moreover note jmp
+          ultimately 
+          have "abrupt (abupd (check_neg i) s2) = Some (Jump j)"
+            by (rule halloc_no_jump')
+          thus ?thesis by (cases s2) auto
+        qed
+        ultimately show "j\<in>jmps" using wt_size G 
+          by - (rule hyp_e [THEN conjunct1,rule_format (no_asm)],simp_all)
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Cast cT e s0 s1 s2 v jmps T Env)
+    {
+      fix j
+      assume jmp: "abrupt s2 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have hyp_e: "PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)" .
+        have "prg Env = G" .
+        moreover from Cast.prems
+        obtain eT where "Env\<turnstile>e\<Colon>-eT" by (elim wt_elim_cases)
+        moreover 
+        have "abrupt s1 = Some (Jump j)"
+        proof -
+          have "s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits cT) ClassCast) s1" .
+          moreover note jmp
+          ultimately show ?thesis by (cases s1) (simp add: abrupt_if_def)
+        qed
+        ultimately show ?thesis 
+          by - (rule hyp_e [THEN conjunct1,rule_format (no_asm)], simp_all)
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Inst eT b e s0 s1 v jmps T Env)
+    {
+      fix j
+      assume jmp: "abrupt s1 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have hyp_e: "PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)" .
+        have "prg Env = G" .
+        moreover from Inst.prems
+        obtain eT where "Env\<turnstile>e\<Colon>-eT" by (elim wt_elim_cases)
+        moreover note jmp
+        ultimately show "j\<in>jmps" 
+          by - (rule hyp_e [THEN conjunct1,rule_format (no_asm)], simp_all)
+      qed
+    }
+    thus ?case by simp
+  next
+    case Lit thus ?case by simp
+  next
+    case (UnOp e s0 s1 unop v jmps T Env)
+    {
+      fix j
+      assume jmp: "abrupt s1 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have hyp_e: "PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)" .
+        have "prg Env = G" .
+        moreover from UnOp.prems
+        obtain eT where "Env\<turnstile>e\<Colon>-eT" by (elim wt_elim_cases)
+        moreover note jmp
+        ultimately show  "j\<in>jmps" 
+          by - (rule hyp_e [THEN conjunct1,rule_format (no_asm)], simp_all) 
+      qed
+    }
+    thus ?case by simp
+  next
+    case (BinOp binop e1 e2 s0 s1 s2 v1 v2 jmps T Env)
+    {
+      fix j
+      assume jmp: "abrupt s2 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have G: "prg Env = G" .
+        from BinOp.prems
+        obtain e1T e2T where 
+          wt_e1: "Env\<turnstile>e1\<Colon>-e1T" and
+          wt_e2: "Env\<turnstile>e2\<Colon>-e2T" 
+          by (elim wt_elim_cases)
+        have "PROP ?Hyp (In1l e1) (Norm s0) s1 (In1 v1)" .
+        with G wt_e1 have jmp_s1: "?Jmp jmps s1" by simp
+        have hyp_e2: 
+          "PROP ?Hyp (if need_second_arg binop v1 then In1l e2 else In1r Skip)
+                     s1 s2 (In1 v2)" .
+        show "j\<in>jmps"
+        proof (cases "need_second_arg binop v1")
+          case True with jmp_s1 wt_e2 jmp G
+          show ?thesis 
+            by - (rule hyp_e2 [THEN conjunct1,rule_format (no_asm)],simp_all)
+        next
+          case False with jmp_s1 jmp G
+          show ?thesis 
+            by - (rule hyp_e2 [THEN conjunct1,rule_format (no_asm)],auto)
+        qed
+      qed
+    }
+    thus ?case by simp
+  next
+    case Super thus ?case by simp
+  next
+    case (Acc f s0 s1 v va jmps T Env)
+    {
+      fix j
+      assume jmp: "abrupt s1 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have hyp_va: "PROP ?Hyp (In2 va) (Norm s0) s1 (In2 (v,f))" .
+        have "prg Env = G" .
+        moreover from Acc.prems
+        obtain vT where "Env\<turnstile>va\<Colon>=vT" by (elim wt_elim_cases)
+        moreover note jmp
+        ultimately show "j\<in>jmps" 
+          by - (rule hyp_va [THEN conjunct1,rule_format (no_asm)], simp_all)
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Ass e f s0 s1 s2 v va w jmps T Env)
+    have G: "prg Env = G" .
+    from Ass.prems
+    obtain vT eT where
+      wt_va: "Env\<turnstile>va\<Colon>=vT" and
+       wt_e: "Env\<turnstile>e\<Colon>-eT"
+      by (elim wt_elim_cases)
+    have hyp_v: "PROP ?Hyp (In2 va) (Norm s0) s1 (In2 (w,f))" .
+    have hyp_e: "PROP ?Hyp (In1l e) s1 s2 (In1 v)" . 
+    {
+      fix j
+      assume jmp: "abrupt (assign f v s2) = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have "abrupt s2 = Some (Jump j)"
+        proof (cases "normal s2")
+          case True
+          have "G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2" .
+          from this True have nrm_s1: "normal s1" 
+            by (rule eval_no_abrupt_lemma [rule_format]) 
+          with nrm_s1 wt_va G True
+          have "abrupt (f v s2) \<noteq> Some (Jump j)"
+            using hyp_v [THEN conjunct2,rule_format (no_asm)]
+            by simp
+          from this jmp
+          show ?thesis
+            by (rule assign_abrupt_propagation)
+        next
+          case False with jmp 
+          show ?thesis by (cases s2) (simp add: assign_def Let_def)
+        qed
+        moreover from wt_va G
+        have "?Jmp jmps s1"
+          by - (rule hyp_v [THEN conjunct1],simp_all)
+        ultimately show ?thesis using G 
+          by - (rule hyp_e [THEN conjunct1,rule_format (no_asm)],simp_all)
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Cond b e0 e1 e2 s0 s1 s2 v jmps T Env)
+    have G: "prg Env = G" .
+    have hyp_e0: "PROP ?Hyp (In1l e0) (Norm s0) s1 (In1 b)" .
+    have hyp_e1_e2: "PROP ?Hyp (In1l (if the_Bool b then e1 else e2)) 
+                               s1 s2 (In1 v)" .
+    from Cond.prems
+    obtain e1T e2T
+      where wt_e0: "Env\<turnstile>e0\<Colon>-PrimT Boolean"
+       and  wt_e1: "Env\<turnstile>e1\<Colon>-e1T"
+       and  wt_e2: "Env\<turnstile>e2\<Colon>-e2T"
+      by (elim wt_elim_cases)
+    {
+      fix j
+      assume jmp: "abrupt s2 = Some (Jump j)"
+      have "j\<in>jmps" 
+      proof -
+        from wt_e0 G 
+        have jmp_s1: "?Jmp jmps s1"
+          by - (rule hyp_e0 [THEN conjunct1],simp_all)
+        show ?thesis
+        proof (cases "the_Bool b")
+          case True
+          with jmp_s1 wt_e1 G jmp
+          show ?thesis
+            by-(rule hyp_e1_e2 [THEN conjunct1,rule_format (no_asm)],simp_all)
+        next
+          case False
+          with jmp_s1 wt_e2 G jmp
+          show ?thesis
+            by-(rule hyp_e1_e2 [THEN conjunct1,rule_format (no_asm)],simp_all)
+        qed
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Call D a accC args e mn mode pTs s0 s1 s2 s3 s3' s4 statT v vs 
+               jmps T Env)
+    have G: "prg Env = G" .
+    from Call.prems
+    obtain eT argsT 
+      where wt_e: "Env\<turnstile>e\<Colon>-eT" and wt_args: "Env\<turnstile>args\<Colon>\<doteq>argsT"
+      by (elim wt_elim_cases)
+    {
+      fix j
+      assume jmp: "abrupt ((set_lvars (locals (store s2))) s4) 
+                     = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have hyp_e: "PROP ?Hyp (In1l e) (Norm s0) s1 (In1 a)" .
+        from wt_e G 
+        have jmp_s1: "?Jmp jmps s1"
+          by - (rule hyp_e [THEN conjunct1],simp_all)
+        have hyp_args: "PROP ?Hyp (In3 args) s1 s2 (In3 vs)" .
+        have "abrupt s2 = Some (Jump j)"
+        proof -
+          have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name = mn, parTs = pTs\<rparr>-\<succ>v\<rightarrow> s4" .
+          moreover
+          from jmp have "abrupt s4 = Some (Jump j)"
+            by (cases s4) simp
+          ultimately have "abrupt s3' = Some (Jump j)"
+            by - (rule ccontr,drule (1) Methd_no_jump,simp)
+          moreover have "s3' = check_method_access G accC statT mode 
+                              \<lparr>name = mn, parTs = pTs\<rparr> a s3" .
+          ultimately have "abrupt s3 = Some (Jump j)"
+            by (cases s3) 
+               (simp add: check_method_access_def abrupt_if_def Let_def)
+          moreover 
+          have "s3 = init_lvars G D \<lparr>name=mn, parTs=pTs\<rparr> mode a vs s2" .
+          ultimately show ?thesis
+            by (cases s2) (auto simp add: init_lvars_def2)
+        qed
+        with jmp_s1 wt_args G
+        show ?thesis
+          by - (rule hyp_args [THEN conjunct1,rule_format (no_asm)], simp_all)
+      qed
+    }
+    thus ?case by simp
+  next
+    case (Methd D s0 s1 sig v jmps T Env)
+    have "G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<rightarrow> s1"
+      by (rule eval.Methd)
+    hence "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+      by (rule Methd_no_jump) simp
+    thus ?case by simp
+  next
+    case (Body D c s0 s1 s2 s3 jmps T Env)
+    have "G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
+           \<rightarrow> abupd (absorb Ret) s3"
+      by (rule eval.Body)
+    hence "\<And> j. abrupt (abupd (absorb Ret) s3) \<noteq> Some (Jump j)"
+      by (rule Body_no_jump) simp
+    thus ?case by simp
+  next
+    case LVar
+    thus ?case by (simp add: lvar_def Let_def)
+  next
+    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v jmps T Env)
+    have G: "prg Env = G" .
+    from wf FVar.prems 
+    obtain  statC f where
+      wt_e: "Env\<turnstile>e\<Colon>-Class statC" and
+      accfield: "accfield (prg Env) accC statC fn = Some (statDeclC,f)"
+      by  (elim wt_elim_cases) simp
+    have wt_init: "Env\<turnstile>Init statDeclC\<Colon>\<surd>"
+    proof -
+      from wf wt_e G 
+      have "is_class (prg Env) statC"
+        by (auto dest: ty_expr_is_type type_is_class)
+      with wf accfield G
+      have "is_class (prg Env) statDeclC"
+        by (auto dest!: accfield_fields dest: fields_declC)
+      thus ?thesis
+        by simp
+    qed
+    have fvar: "(v, s2') = fvar statDeclC stat fn a s2" .
+    {
+      fix j
+      assume jmp: "abrupt s3 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have hyp_init: "PROP ?Hyp (In1r (Init statDeclC)) (Norm s0) s1 \<diamondsuit>" .
+        from G wt_init 
+        have "?Jmp jmps s1"
+          by - (rule hyp_init [THEN conjunct1],auto)
+        moreover
+        have hyp_e: "PROP ?Hyp (In1l e) s1 s2 (In1 a)" .
+        have "abrupt s2 = Some (Jump j)"
+        proof -
+          have "s3 = check_field_access G accC statDeclC fn stat a s2'" .
+          with jmp have "abrupt s2' = Some (Jump j)"
+            by (cases s2') 
+               (simp add: check_field_access_def abrupt_if_def Let_def)
+         with fvar show "abrupt s2 =  Some (Jump j)"
+            by (cases s2) (simp add: fvar_def2 abrupt_if_def)
+        qed
+        ultimately show ?thesis
+          using G wt_e
+          by - (rule hyp_e [THEN conjunct1, rule_format (no_asm)],simp_all)
+      qed
+    }
+    moreover
+    from fvar obtain upd w
+      where upd: "upd = snd (fst (fvar statDeclC stat fn a s2))" and
+              v: "v=(w,upd)"
+      by (cases "fvar statDeclC stat fn a s2") simp
+    {
+      fix j val fix s::state
+      assume "normal s3"
+      assume no_jmp: "abrupt s \<noteq> Some (Jump j)"
+      with upd
+      have "abrupt (upd val s) \<noteq> Some (Jump j)"
+        by (rule fvar_upd_no_jump)
+    }
+    ultimately show ?case using v by simp
+  next
+    case (AVar a e1 e2 i s0 s1 s2 s2' v jmps T Env)
+    have G: "prg Env = G" .
+    from AVar.prems 
+    obtain  e1T e2T where
+      wt_e1: "Env\<turnstile>e1\<Colon>-e1T" and wt_e2: "Env\<turnstile>e2\<Colon>-e2T"
+      by  (elim wt_elim_cases) simp
+    have avar: "(v, s2') = avar G i a s2" .
+    {
+      fix j
+      assume jmp: "abrupt s2' = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have hyp_e1: "PROP ?Hyp (In1l e1) (Norm s0) s1 (In1 a)" .
+        from G wt_e1
+        have "?Jmp jmps s1"
+          by - (rule hyp_e1 [THEN conjunct1],auto)
+        moreover
+        have hyp_e2: "PROP ?Hyp (In1l e2) s1 s2 (In1 i)" .
+        have "abrupt s2 = Some (Jump j)"
+        proof -
+          from avar have "s2' = snd (avar G i a s2)"
+            by (cases "avar G i a s2") simp
+          with jmp show ?thesis by - (rule avar_state_no_jump,simp) 
+        qed  
+        ultimately show ?thesis
+          using wt_e2 G
+          by - (rule hyp_e2 [THEN conjunct1, rule_format (no_asm)],simp_all)
+      qed
+    }
+    moreover
+    from avar obtain upd w
+      where upd: "upd = snd (fst (avar G i a s2))" and
+              v: "v=(w,upd)"
+      by (cases "avar G i a s2") simp
+    {
+      fix j val fix s::state
+      assume "normal s2'"
+      assume no_jmp: "abrupt s \<noteq> Some (Jump j)"
+      with upd
+      have "abrupt (upd val s) \<noteq> Some (Jump j)"
+        by (rule avar_upd_no_jump)
+    }
+    ultimately show ?case using v by simp
+  next
+    case Nil thus ?case by simp
+  next
+    case (Cons e es s0 s1 s2 v vs jmps T Env)
+    have G: "prg Env = G" .
+    from Cons.prems obtain eT esT
+      where wt_e: "Env\<turnstile>e\<Colon>-eT" and wt_e2: "Env\<turnstile>es\<Colon>\<doteq>esT"
+      by  (elim wt_elim_cases) simp
+    {
+      fix j
+      assume jmp: "abrupt s2 = Some (Jump j)"
+      have "j\<in>jmps"
+      proof -
+        have hyp_e: "PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)" .
+        from G wt_e
+        have "?Jmp jmps s1"
+          by - (rule hyp_e [THEN conjunct1],simp_all)
+        moreover
+        have hyp_es: "PROP ?Hyp (In3 es) s1 s2 (In3 vs)" .  
+        ultimately show ?thesis
+          using wt_e2 G jmp
+          by - (rule hyp_es [THEN conjunct1, rule_format (no_asm)],
+                (assumption|simp (no_asm_simp))+)
+      qed
+    }
+    thus ?case by simp
+  qed
+  note generalized = this
+  from no_jmp jmpOk wt G
+  show ?thesis
+    by (rule generalized)
+qed
+
+lemmas jumpNestingOk_evalE = jumpNestingOk_eval [THEN conjE,rule_format]
+
+
+lemma jumpNestingOk_eval_no_jump:
+ assumes    eval: "prg Env\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" and
+           jmpOk: "jumpNestingOk {} t" and
+          no_jmp: "abrupt s0 \<noteq> Some (Jump j)" and
+              wt: "Env\<turnstile>t\<Colon>T" and 
+              wf: "wf_prog (prg Env)" 
+ shows "abrupt s1 \<noteq> Some (Jump j) \<and>
+        (normal s1 \<longrightarrow> v=In2 (w,upd)  
+         \<longrightarrow> abrupt s \<noteq> Some (Jump j')
+         \<longrightarrow> abrupt (upd val s) \<noteq> Some (Jump j'))"
+proof (cases "\<exists> j'. abrupt s0 = Some (Jump j')") 
+  case True
+  then obtain j' where jmp: "abrupt s0 = Some (Jump j')" ..
+  with no_jmp have "j'\<noteq>j" by simp
+  with eval jmp have "s1=s0" by auto
+  with no_jmp jmp show ?thesis by simp
+next
+  case False
+  obtain G where G: "prg Env = G"
+    by (cases Env) simp
+  from G eval have "G\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" by simp
+  moreover note jmpOk wt
+  moreover from G wf have "wf_prog G" by simp
+  moreover note G 
+  moreover from False have "\<And> j. abrupt s0 = Some (Jump j) \<Longrightarrow> j \<in> {}"
+    by simp
+  ultimately show ?thesis
+    apply (rule jumpNestingOk_evalE)
+    apply assumption
+    apply simp
+    apply fastsimp
+    done
+qed
+
+lemmas jumpNestingOk_eval_no_jumpE 
+       = jumpNestingOk_eval_no_jump [THEN conjE,rule_format]
+
+corollary eval_expression_no_jump:
+  assumes eval: "prg Env\<turnstile>s0 \<midarrow>e-\<succ>v\<rightarrow> s1" and
+        no_jmp: "abrupt s0 \<noteq> Some (Jump j)" and
+        wt: "Env\<turnstile>e\<Colon>-T" and 
+        wf: "wf_prog (prg Env)" 
+  shows "abrupt s1 \<noteq> Some (Jump j)"
+using eval _ no_jmp wt wf
+by (rule jumpNestingOk_eval_no_jumpE, simp_all)
+
+corollary eval_var_no_jump:
+  assumes eval: "prg Env\<turnstile>s0 \<midarrow>var=\<succ>(w,upd)\<rightarrow> s1" and
+        no_jmp: "abrupt s0 \<noteq> Some (Jump j)" and
+        wt: "Env\<turnstile>var\<Colon>=T" and 
+        wf: "wf_prog (prg Env)" 
+  shows "abrupt s1 \<noteq> Some (Jump j) \<and> 
+         (normal s1 \<longrightarrow> 
+          (abrupt s \<noteq> Some (Jump j') 
+           \<longrightarrow> abrupt (upd val s) \<noteq> Some (Jump j')))"
+apply (rule_tac upd="upd" and val="val" and s="s" and w="w" and j'=j' 
+         in jumpNestingOk_eval_no_jumpE [OF eval _ no_jmp wt wf])
+by simp_all
+
+lemmas eval_var_no_jumpE = eval_var_no_jump [THEN conjE,rule_format]
+
+corollary eval_statement_no_jump:
+  assumes eval: "prg Env\<turnstile>s0 \<midarrow>c\<rightarrow> s1" and
+         jmpOk: "jumpNestingOkS {} c" and
+        no_jmp: "abrupt s0 \<noteq> Some (Jump j)" and
+        wt: "Env\<turnstile>c\<Colon>\<surd>" and 
+        wf: "wf_prog (prg Env)" 
+  shows "abrupt s1 \<noteq> Some (Jump j)"
+using eval _ no_jmp wt wf
+by (rule jumpNestingOk_eval_no_jumpE) (simp_all add: jmpOk)
+
+corollary eval_expression_list_no_jump:
+  assumes eval: "prg Env\<turnstile>s0 \<midarrow>es\<doteq>\<succ>v\<rightarrow> s1" and
+        no_jmp: "abrupt s0 \<noteq> Some (Jump j)" and
+        wt: "Env\<turnstile>es\<Colon>\<doteq>T" and 
+        wf: "wf_prog (prg Env)" 
+  shows "abrupt s1 \<noteq> Some (Jump j)"
+using eval _ no_jmp wt wf
+by (rule jumpNestingOk_eval_no_jumpE, simp_all)
+
+(* ### FIXME: Do i need this *)
+lemma union_subseteq_elim [elim]: "\<lbrakk>A \<union> B \<subseteq> C; \<lbrakk>A \<subseteq> C; B \<subseteq> C\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+  by blast
+
+lemma dom_locals_halloc_mono:
+  assumes halloc: "G\<turnstile>s0\<midarrow>halloc oi\<succ>a\<rightarrow>s1"
+  shows "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+proof -
+  from halloc show ?thesis
+    by cases simp_all
+qed
+ 
+lemma dom_locals_sxalloc_mono:
+  assumes sxalloc: "G\<turnstile>s0\<midarrow>sxalloc\<rightarrow>s1"
+  shows "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+proof -
+  from sxalloc show ?thesis
+  proof (cases)
+    case Norm thus ?thesis by simp
+  next
+    case Jmp thus ?thesis by simp
+  next
+    case Error thus ?thesis by simp
+  next
+    case XcptL thus ?thesis by simp
+  next
+    case SXcpt thus ?thesis 
+      by - (drule dom_locals_halloc_mono,simp)
+  qed
+qed
+    
+
+lemma dom_locals_assign_mono: 
+ assumes f_ok: "dom (locals (store s)) \<subseteq> dom (locals (store (f n s)))"
+   shows  "dom (locals (store s)) \<subseteq> dom (locals (store (assign f n s)))"
+proof (cases "normal s")
+  case False thus ?thesis
+    by (cases s) (auto simp add: assign_def Let_def)
+next
+  case True
+  then obtain s' where s': "s = (None,s')"
+    by auto
+  moreover
+  obtain x1 s1 where "f n s = (x1,s1)"
+    by (cases "f n s", simp)
+  ultimately
+  show ?thesis
+    using f_ok
+    by (simp add: assign_def Let_def)
+qed
+
+(*
+lemma dom_locals_init_lvars_mono: 
+ "dom (locals (store s)) 
+   \<subseteq> dom (locals (store (init_lvars G D sig mode a vs s)))"
+proof (cases "mode = Static")
+  case True
+  thus ?thesis
+    apply (cases s)
+    apply (simp add: init_lvars_def Let_def)
+*)
+
+lemma dom_locals_lvar_mono:
+ "dom (locals (store s)) \<subseteq> dom (locals (store (snd (lvar vn s') val s)))"
+by (simp add: lvar_def) blast
+
+
+lemma dom_locals_fvar_vvar_mono:
+"dom (locals (store s)) 
+  \<subseteq> dom (locals (store (snd (fst (fvar statDeclC stat fn a s')) val s)))"
+proof (cases stat) 
+  case True
+  thus ?thesis
+    by (cases s) (simp add: fvar_def2)
+next
+  case False
+  thus ?thesis
+    by (cases s) (simp add: fvar_def2)
+qed
+
+lemma dom_locals_fvar_mono:
+"dom (locals (store s)) 
+  \<subseteq> dom (locals (store (snd (fvar statDeclC stat fn a s))))"
+proof (cases stat) 
+  case True
+  thus ?thesis
+    by (cases s) (simp add: fvar_def2)
+next
+  case False
+  thus ?thesis
+    by (cases s) (simp add: fvar_def2)
+qed
+
+
+lemma dom_locals_avar_vvar_mono:
+"dom (locals (store s)) 
+  \<subseteq> dom (locals (store (snd (fst (avar G i a s')) val s)))"
+by (cases s, simp add: avar_def2)
+
+lemma dom_locals_avar_mono:
+"dom (locals (store s)) 
+  \<subseteq> dom (locals (store (snd (avar G i a s))))"
+by (cases s, simp add: avar_def2)
+
+  text {* 
+Since assignments are modelled as functions from states to states, we
+  must take into account these functions. They  appear only in the assignment 
+  rule and as result from evaluating a variable. Thats why we need the 
+  complicated second part of the conjunction in the goal.
+ The reason for the very generic way to treat assignments was the aim
+to omit redundancy. There is only one evaluation rule for each kind of
+variable (locals, fields, arrays). These rules are used for both accessing 
+variables and updating variables. Thats why the evaluation rules for variables
+result in a pair consisting of a value and an update function. Of course we
+could also think of a pair of a value and a reference in the store, instead of
+the generic update function. But as only array updates can cause a special
+exception (if the types mismatch) and not array reads we then have to introduce
+two different rules to handle array reads and updates *} 
+lemma dom_locals_eval_mono: 
+  assumes   eval: "G\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" 
+  shows "dom (locals (store s0)) \<subseteq> dom (locals (store s1)) \<and>
+         (\<forall> vv. v=In2 vv \<and> normal s1 
+            \<longrightarrow> (\<forall> s val. dom (locals (store s)) 
+                           \<subseteq> dom (locals (store ((snd vv) val s)))))"
+proof -
+  from eval show ?thesis
+  proof (induct)
+    case Abrupt thus ?case by simp 
+  next
+    case Skip thus ?case by simp
+  next
+    case Expr thus ?case by simp
+  next
+    case Lab thus ?case by simp
+  next
+    case (Comp c1 c2 s0 s1 s2) 
+    from Comp.hyps 
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by simp
+    also
+    from Comp.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" 
+      by simp
+    finally show ?case by simp
+  next
+    case (If b c1 c2 e s0 s1 s2)
+    from If.hyps 
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by simp
+    also
+    from If.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" 
+      by simp
+    finally show ?case by simp
+  next
+    case (Loop b c e l s0 s1 s2 s3) 
+    show ?case
+    proof (cases "the_Bool b")
+      case True
+      with Loop.hyps
+      obtain
+	s0_s1: 
+	"dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))" and
+	s1_s2: "dom (locals (store s1)) \<subseteq> dom (locals (store s2))" and
+	s2_s3: "dom (locals (store s2)) \<subseteq> dom (locals (store s3))"
+	by simp
+      note s0_s1 also note s1_s2 also note s2_s3
+      finally show ?thesis
+	by simp
+    next
+      case False
+      with Loop.hyps show ?thesis
+	by simp
+    qed
+  next
+    case Jmp thus ?case by simp
+  next
+    case Throw thus ?case by simp
+  next
+    case (Try C c1 c2 s0 s1 s2 s3 vn)
+    then
+    have s0_s1: "dom (locals (store ((Norm s0)::state))) 
+                  \<subseteq> dom (locals (store s1))" by simp
+    have "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" .
+    hence s1_s2: "dom (locals (store s1)) \<subseteq> dom (locals (store s2))" 
+      by (rule dom_locals_sxalloc_mono)
+    thus ?case 
+    proof (cases "G,s2\<turnstile>catch C")
+      case True
+      note s0_s1 also note s1_s2
+      also
+      from True Try.hyps 
+      have "dom (locals (store (new_xcpt_var vn s2))) 
+              \<subseteq> dom (locals (store s3))"
+	by simp
+      hence "dom (locals (store s2)) \<subseteq> dom (locals (store s3))"
+	by (cases s2, simp )
+      finally show ?thesis by simp
+    next
+      case False
+      note s0_s1 also note s1_s2
+      finally 
+      show ?thesis 
+	using False Try.hyps by simp
+    qed
+  next
+    case (Fin c1 c2 s0 s1 s2 s3 x1) 
+    show ?case
+    proof (cases "\<exists>err. x1 = Some (Error err)")
+      case True
+      with Fin.hyps show ?thesis
+	by simp
+    next
+      case False
+      from Fin.hyps
+      have "dom (locals (store ((Norm s0)::state))) 
+             \<subseteq> dom (locals (store (x1, s1)))" 
+	by simp
+      hence "dom (locals (store ((Norm s0)::state))) 
+              \<subseteq> dom (locals (store ((Norm s1)::state)))"
+	by simp
+      also
+      from Fin.hyps
+      have "\<dots> \<subseteq> dom (locals (store s2))"
+	by simp
+      finally show ?thesis 
+	using Fin.hyps by simp
+    qed
+  next
+    case (Init C c s0 s1 s2 s3)
+    show ?case
+    proof (cases "inited C (globs s0)")
+      case True
+      with Init.hyps show ?thesis by simp 
+    next
+      case False
+      with Init.hyps 
+      obtain s0_s1: "dom (locals (store (Norm ((init_class_obj G C) s0))))
+                     \<subseteq> dom (locals (store s1))" and
+	     s3: "s3 = (set_lvars (locals (snd s1))) s2"
+	by simp
+      from s0_s1
+      have "dom (locals (store (Norm s0))) \<subseteq> dom (locals (store s1))"
+	by (cases s0) simp
+      with s3
+      have "dom (locals (store (Norm s0))) \<subseteq> dom (locals (store s3))"
+	by (cases s2) simp
+      thus ?thesis by simp
+    qed
+  next
+    case (NewC C a s0 s1 s2)
+    have halloc: "G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2" .
+    from NewC.hyps
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))" 
+      by simp
+    also
+    from halloc
+    have "\<dots>  \<subseteq> dom (locals (store s2))" by (rule dom_locals_halloc_mono)
+    finally show ?case by simp
+  next
+    case (NewA T a e i s0 s1 s2 s3)
+    have halloc: "G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3" .
+    from NewA.hyps
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))" 
+      by simp
+    also
+    from NewA.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" by simp
+    also
+    from halloc 
+    have "\<dots> \<subseteq>  dom (locals (store s3))"
+      by (rule dom_locals_halloc_mono [elim_format]) simp
+    finally show ?case by simp
+  next
+    case Cast thus ?case by simp
+  next
+    case Inst thus ?case by simp
+  next
+    case Lit thus ?case by simp
+  next
+    case UnOp thus ?case by simp
+  next
+    case (BinOp binop e1 e2 s0 s1 s2 v1 v2) 
+    from BinOp.hyps
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))" 
+      by simp
+    also
+    from BinOp.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" by simp
+    finally show ?case by simp
+  next
+    case Super thus ?case by simp
+  next
+    case Acc thus ?case by simp
+  next
+    case (Ass e f s0 s1 s2 v va w)
+    from Ass.hyps
+    have s0_s1: 
+      "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))" 
+      by simp
+    show ?case
+    proof (cases "normal s1")
+      case True
+      with Ass.hyps 
+      have ass_ok:
+        "\<And> s val. dom (locals (store s)) \<subseteq> dom (locals (store (f val s)))"
+	by simp
+      note s0_s1
+      also
+      from Ass.hyps
+      have "dom (locals (store s1)) \<subseteq> dom (locals (store s2))" 
+	by simp
+      also
+      from ass_ok
+      have "\<dots> \<subseteq> dom (locals (store (assign f v s2)))"
+	by (rule dom_locals_assign_mono)
+      finally show ?thesis by simp
+    next
+      case False
+      have "G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2" .
+      with False 
+      have "s2=s1"
+	by auto
+      with s0_s1 False
+      have "dom (locals (store ((Norm s0)::state))) 
+            \<subseteq> dom (locals (store (assign f v s2)))"
+	by simp
+      thus ?thesis
+	by simp
+    qed
+  next
+    case (Cond b e0 e1 e2 s0 s1 s2 v)
+    from Cond.hyps 
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by simp
+    also
+    from Cond.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" 
+      by simp
+    finally show ?case by simp
+  next
+    case (Call D a' accC args e mn mode pTs s0 s1 s2 s3 s3' s4 statT v vs)
+    have s3: "s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a' vs s2" .
+    from Call.hyps 
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by simp
+    also
+    from Call.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" 
+      by simp
+    also
+    have "\<dots> \<subseteq> dom (locals (store ((set_lvars (locals (store s2))) s4)))"
+      by (cases s4) simp
+    finally show ?case by simp
+  next
+    case Methd thus ?case by simp
+  next
+    case (Body D c s0 s1 s2 s3)
+    from Body.hyps 
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by simp
+    also
+    from Body.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" 
+      by simp
+    also
+    have "\<dots> \<subseteq> dom (locals (store (abupd (absorb Ret) s2)))"
+      by simp
+    also
+    have "\<dots> \<subseteq> dom (locals (store (abupd (absorb Ret) s3)))"
+    proof -
+       have "s3 =
+         (if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or> 
+                 abrupt s2 = Some (Jump (Cont l))
+             then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)".
+      thus ?thesis
+	by simp
+    qed
+    finally show ?case by simp
+  next
+    case LVar
+    thus ?case
+      using dom_locals_lvar_mono
+      by simp
+  next
+    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v)
+    from FVar.hyps
+    obtain s2': "s2' = snd (fvar statDeclC stat fn a s2)" and
+             v: "v = fst (fvar statDeclC stat fn a s2)"
+      by (cases "fvar statDeclC stat fn a s2" ) simp
+    from v 
+    have "\<forall>s val. dom (locals (store s)) 
+                          \<subseteq> dom (locals (store (snd v val s)))" (is "?V_ok")
+      by (simp add: dom_locals_fvar_vvar_mono) 
+    hence v_ok: "(\<forall>vv. In2 v = In2 vv \<and> normal s3 \<longrightarrow> ?V_ok)"
+      by - (intro strip, simp)
+    have s3: "s3 = check_field_access G accC statDeclC fn stat a s2'" .
+    from FVar.hyps 
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by simp
+    also
+    from FVar.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" 
+      by simp
+    also
+    from s2'
+    have "\<dots> \<subseteq> dom (locals (store s2'))"
+      by (simp add: dom_locals_fvar_mono)
+    also
+    from s3
+    have "\<dots> \<subseteq> dom (locals (store s3))"
+      by (simp add: check_field_access_def Let_def)
+    finally
+    show ?case
+      using v_ok
+      by simp
+  next
+    case (AVar a e1 e2 i s0 s1 s2 s2' v)
+    from AVar.hyps
+    obtain s2': "s2' = snd (avar G i a s2)" and
+             v: "v   = fst (avar G i a s2)"
+      by (cases "avar G i a s2") simp
+    from v
+    have "\<forall>s val. dom (locals (store s)) 
+                          \<subseteq> dom (locals (store (snd v val s)))" (is "?V_ok")
+      by (simp add: dom_locals_avar_vvar_mono)
+    hence v_ok: "(\<forall>vv. In2 v = In2 vv \<and> normal s2' \<longrightarrow> ?V_ok)"
+      by - (intro strip, simp)
+    from AVar.hyps 
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by simp
+    also
+    from AVar.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" 
+      by simp
+    also
+    from s2'
+    have "\<dots> \<subseteq> dom (locals (store s2'))"
+      by (simp add: dom_locals_avar_mono)
+    finally
+    show ?case using v_ok by simp
+  next
+    case Nil thus ?case by simp
+  next
+    case (Cons e es s0 s1 s2 v vs)
+    from Cons.hyps
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by simp
+    also
+    from Cons.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))" 
+      by simp
+    finally show ?case by simp
+  qed
+qed
+     
+lemma dom_locals_eval_mono_elim [consumes 1]: 
+  assumes   eval: "G\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" and
+    hyps: "\<lbrakk>dom (locals (store s0)) \<subseteq> dom (locals (store s1));
+           \<And> vv s val. \<lbrakk>v=In2 vv; normal s1\<rbrakk> 
+                        \<Longrightarrow> dom (locals (store s)) 
+                             \<subseteq> dom (locals (store ((snd vv) val s)))\<rbrakk> \<Longrightarrow> P"
+ shows "P"
+using eval
+proof (rule dom_locals_eval_mono [THEN conjE])
+qed (rule hyps,auto)
+
+lemma halloc_no_abrupt:
+  assumes halloc: "G\<turnstile>s0\<midarrow>halloc oi\<succ>a\<rightarrow>s1" and
+          normal: "normal s1"
+  shows "normal s0"
+proof -
+  from halloc normal show ?thesis
+    by cases simp_all
+qed
+ 
+lemma sxalloc_mono_no_abrupt:
+  assumes sxalloc: "G\<turnstile>s0\<midarrow>sxalloc\<rightarrow>s1" and
+           normal: "normal s1"
+  shows "normal s0"
+proof -
+  from sxalloc normal show ?thesis
+    by cases simp_all
+qed
+   
+lemma union_subseteqI: "\<lbrakk>A \<union> B \<subseteq> C; A' \<subseteq> A; B' \<subseteq> B\<rbrakk>  \<Longrightarrow>   A' \<union> B' \<subseteq> C"
+  by blast
+
+lemma union_subseteqIl: "\<lbrakk>A \<union> B \<subseteq> C; A' \<subseteq> A\<rbrakk>  \<Longrightarrow>   A' \<union> B \<subseteq> C"
+  by blast
+
+lemma union_subseteqIr: "\<lbrakk>A \<union> B \<subseteq> C; B' \<subseteq> B\<rbrakk>  \<Longrightarrow>   A \<union> B' \<subseteq> C"
+  by blast
+
+lemma subseteq_union_transl [trans]: "\<lbrakk>A \<subseteq> B; B \<union> C \<subseteq> D\<rbrakk> \<Longrightarrow> A \<union> C \<subseteq> D"
+  by blast
+
+lemma subseteq_union_transr [trans]: "\<lbrakk>A \<subseteq> B; C \<union> B \<subseteq> D\<rbrakk> \<Longrightarrow> A \<union> C \<subseteq> D"
+  by blast
+
+lemma union_subseteq_weaken: "\<lbrakk>A \<union> B \<subseteq> C; \<lbrakk>A \<subseteq> C; B \<subseteq> C\<rbrakk> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"  
+  by blast
+
+lemma assigns_good_approx: 
+  assumes
+      eval: "G\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" and
+    normal: "normal s1" 
+  shows "assigns t \<subseteq> dom (locals (store s1))"
+proof -
+  from eval normal show ?thesis
+  proof (induct)
+    case Abrupt thus ?case by simp 
+  next -- {* For statements its trivial, since then @{term "assigns t = {}"} *}
+    case Skip show ?case by simp
+  next
+    case Expr show ?case by simp 
+  next
+    case Lab show ?case by simp
+  next
+    case Comp show ?case by simp
+  next
+    case If show ?case by simp
+  next
+    case Loop show ?case by simp
+  next
+    case Jmp show ?case by simp
+  next
+    case Throw show ?case by simp
+  next
+    case Try show ?case by simp 
+  next
+    case Fin show ?case by simp
+  next
+    case Init show ?case by simp 
+  next
+    case NewC show ?case by simp
+  next
+    case (NewA T a e i s0 s1 s2 s3)
+    have halloc: "G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3" .
+    have "assigns (In1l e) \<subseteq> dom (locals (store s2))"
+    proof -
+      from NewA
+      have "normal (abupd (check_neg i) s2)"
+	by - (erule halloc_no_abrupt [rule_format])
+      hence "normal s2" by (cases s2) simp
+      with NewA.hyps
+      show ?thesis by rules
+    qed
+    also
+    from halloc 
+    have "\<dots> \<subseteq>  dom (locals (store s3))"
+      by (rule dom_locals_halloc_mono [elim_format]) simp
+    finally show ?case by simp 
+  next
+    case (Cast T e s0 s1 s2 v)
+    hence "normal s1" by (cases s1,simp)
+    with Cast.hyps
+    have "assigns (In1l e) \<subseteq> dom (locals (store s1))"
+      by simp
+    also
+    from Cast.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))"
+      by simp
+    finally 
+    show ?case
+      by simp
+  next
+    case Inst thus ?case by simp
+  next
+    case Lit thus ?case by simp
+  next
+    case UnOp thus ?case by simp
+  next
+    case (BinOp binop e1 e2 s0 s1 s2 v1 v2)
+    hence "normal s1" by - (erule eval_no_abrupt_lemma [rule_format]) 
+    with BinOp.hyps
+    have "assigns (In1l e1) \<subseteq> dom (locals (store s1))"
+      by rules
+    also
+    have "\<dots>  \<subseteq> dom (locals (store s2))"
+    proof -
+      have "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
+                      else In1r Skip)\<succ>\<rightarrow> (In1 v2, s2)" .
+      thus ?thesis
+	by (rule dom_locals_eval_mono_elim)
+    qed
+    finally have s2: "assigns (In1l e1) \<subseteq> dom (locals (store s2))" .
+    show ?case
+    proof (cases "binop=CondAnd \<or> binop=CondOr")
+      case True
+      with s2 show ?thesis by simp 
+    next
+      case False
+      with BinOp 
+      have "assigns (In1l e2) \<subseteq> dom (locals (store s2))"
+	by (simp add: need_second_arg_def)
+      with s2
+      show ?thesis using False by (simp add: Un_subset_iff)
+    qed
+  next
+    case Super thus ?case by simp
+  next
+    case Acc thus ?case by simp
+  next 
+    case (Ass e f s0 s1 s2 v va w)
+    have  nrm_ass_s2: "normal (assign f v s2)" .
+    hence nrm_s2: "normal s2"
+      by (cases s2, simp add: assign_def Let_def)
+    with Ass.hyps 
+    have nrm_s1: "normal s1"
+      by - (erule eval_no_abrupt_lemma [rule_format]) 
+    with Ass.hyps
+    have "assigns (In2 va) \<subseteq> dom (locals (store s1))" 
+      by rules
+    also
+    from Ass.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))"
+      by - (erule dom_locals_eval_mono_elim)
+    also
+    from nrm_s2 Ass.hyps
+    have "assigns (In1l e) \<subseteq> dom (locals (store s2))"
+      by rules
+    ultimately
+    have "assigns (In2 va) \<union> assigns (In1l e) \<subseteq>  dom (locals (store s2))"
+      by (rule Un_least)
+    also
+    from Ass.hyps nrm_s1
+    have "\<dots>  \<subseteq> dom (locals (store (f v s2)))"
+      by - (erule dom_locals_eval_mono_elim, cases s2,simp)
+    then
+    have "dom (locals (store s2))  \<subseteq> dom (locals (store (assign f v s2)))"
+      by (rule dom_locals_assign_mono)
+    finally
+    have va_e: " assigns (In2 va) \<union> assigns (In1l e)
+                  \<subseteq> dom (locals (snd (assign f v s2)))" .
+    show ?case
+    proof (cases "\<exists> n. va = LVar n")
+      case False
+      with va_e show ?thesis 
+	by (simp add: Un_assoc)
+    next
+      case True
+      then obtain n where va: "va = LVar n"
+	by blast
+      with Ass.hyps 
+      have "G\<turnstile>Norm s0 \<midarrow>LVar n=\<succ>(w,f)\<rightarrow> s1" 
+	by simp
+      hence "(w,f) = lvar n s0"
+	by (rule eval_elim_cases) simp
+      with nrm_ass_s2
+      have "n \<in> dom (locals (store (assign f v s2)))"
+	by (cases s2) (simp add: assign_def Let_def lvar_def)
+      with va_e True va 
+      show ?thesis by (simp add: Un_assoc)
+    qed
+  next
+    case (Cond b e0 e1 e2 s0 s1 s2 v) 
+    hence "normal s1"
+      by - (erule eval_no_abrupt_lemma [rule_format])
+    with Cond.hyps
+    have "assigns (In1l e0) \<subseteq> dom (locals (store s1))"
+      by rules
+    also from Cond.hyps
+    have "\<dots> \<subseteq> dom (locals (store s2))"
+      by - (erule dom_locals_eval_mono_elim)
+    finally have e0: "assigns (In1l e0) \<subseteq> dom (locals (store s2))" .
+    show ?case
+    proof (cases "the_Bool b")
+      case True
+      with Cond 
+      have "assigns (In1l e1) \<subseteq> dom (locals (store s2))"
+	by simp
+      hence "assigns (In1l e1) \<inter> assigns (In1l e2) \<subseteq> \<dots>" 
+	by blast
+      with e0
+      have "assigns (In1l e0) \<union> assigns (In1l e1) \<inter> assigns (In1l e2) 
+             \<subseteq> dom (locals (store s2))"
+	by (rule Un_least)
+      thus ?thesis using True by simp 
+    next
+      case False
+      with Cond 
+      have " assigns (In1l e2) \<subseteq> dom (locals (store s2))"
+	by simp
+      hence "assigns (In1l e1) \<inter> assigns (In1l e2) \<subseteq> \<dots>"
+	by blast
+      with e0
+      have "assigns (In1l e0) \<union> assigns (In1l e1) \<inter> assigns (In1l e2) 
+             \<subseteq> dom (locals (store s2))"
+	by (rule Un_least)
+      thus ?thesis using False by simp 
+    qed
+  next
+    case (Call D a' accC args e mn mode pTs s0 s1 s2 s3 s3' s4 statT v vs)
+    have nrm_s2: "normal s2"
+    proof -
+      have "normal ((set_lvars (locals (snd s2))) s4)" .
+      hence normal_s4: "normal s4" by simp
+      hence "normal s3'" using Call.hyps
+	by - (erule eval_no_abrupt_lemma [rule_format]) 
+      moreover have  
+       "s3' = check_method_access G accC statT mode \<lparr>name=mn, parTs=pTs\<rparr> a' s3".
+      ultimately have "normal s3"
+	by (cases s3) (simp add: check_method_access_def Let_def) 
+      moreover
+      have s3: "s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a' vs s2" .
+      ultimately show "normal s2"
+	by (cases s2) (simp add: init_lvars_def2)
+    qed
+    hence "normal s1" using Call.hyps
+      by - (erule eval_no_abrupt_lemma [rule_format])
+    with Call.hyps
+    have "assigns (In1l e) \<subseteq> dom (locals (store s1))"
+      by rules
+    also from Call.hyps
+    have "\<dots> \<subseteq>  dom (locals (store s2))"
+      by - (erule dom_locals_eval_mono_elim)
+    also
+    from nrm_s2 Call.hyps
+    have "assigns (In3 args) \<subseteq> dom (locals (store s2))"
+      by rules
+    ultimately have "assigns (In1l e) \<union> assigns (In3 args) \<subseteq> \<dots>"
+      by (rule Un_least)
+    also 
+    have "\<dots> \<subseteq> dom (locals (store ((set_lvars (locals (store s2))) s4)))"
+      by (cases s4) simp
+    finally show ?case
+      by simp
+  next
+    case Methd thus ?case by simp
+  next
+    case Body thus ?case by simp
+  next
+    case LVar thus ?case by simp
+  next
+    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v)
+    have s3:  "s3 = check_field_access G accC statDeclC fn stat a s2'" .
+    have avar: "(v, s2') = fvar statDeclC stat fn a s2" .
+    have nrm_s2: "normal s2"
+    proof -
+      have "normal s3" .
+      with s3 have "normal s2'"
+	by (cases s2') (simp add: check_field_access_def Let_def)
+      with avar show "normal s2"
+	by (cases s2) (simp add: fvar_def2)
+    qed
+    with FVar.hyps 
+    have "assigns (In1l e) \<subseteq> dom (locals (store s2))"
+      by rules
+    also
+    have "\<dots> \<subseteq>  dom (locals (store s2'))"
+    proof -
+      from avar
+      have "s2' = snd (fvar statDeclC stat fn a s2)"
+	by (cases "fvar statDeclC stat fn a s2") simp
+      thus ?thesis
+	by simp (rule dom_locals_fvar_mono)
+    qed
+    also from s3
+    have "\<dots> \<subseteq>  dom (locals (store s3))"
+      by (cases s2') (simp add: check_field_access_def Let_def)
+    finally show ?case
+      by simp
+  next
+    case (AVar a e1 e2 i s0 s1 s2 s2' v)
+    have avar: "(v, s2') = avar G i a s2" .
+    have nrm_s2: "normal s2"
+    proof -
+      have "normal s2'" .
+      with avar
+      show ?thesis by (cases s2) (simp add: avar_def2)
+    qed
+    with AVar.hyps 
+    have "normal s1"
+      by - (erule eval_no_abrupt_lemma [rule_format])
+    with AVar.hyps
+    have "assigns (In1l e1) \<subseteq> dom (locals (store s1))"
+      by rules
+    also from AVar.hyps
+    have "\<dots> \<subseteq>  dom (locals (store s2))"
+      by - (erule dom_locals_eval_mono_elim)
+    also
+    from AVar.hyps nrm_s2
+    have "assigns (In1l e2) \<subseteq> dom (locals (store s2))"
+      by rules
+    ultimately
+    have "assigns (In1l e1) \<union> assigns (In1l e2) \<subseteq> \<dots>"
+      by (rule Un_least)
+    also
+    have "dom (locals (store s2)) \<subseteq>  dom (locals (store s2'))"
+    proof -
+      from avar have "s2' = snd (avar G i a s2)"
+	by (cases "avar G i a s2") simp
+      thus ?thesis
+	by simp (rule dom_locals_avar_mono)
+    qed
+    finally  
+    show ?case
+      by simp
+  next
+    case Nil show ?case by simp
+  next
+    case (Cons e es s0 s1 s2 v vs)
+    have "assigns (In1l e) \<subseteq> dom (locals (store s1))"
+    proof -
+      from Cons
+      have "normal s1" by - (erule eval_no_abrupt_lemma [rule_format])
+      with Cons.hyps show ?thesis by rules
+    qed
+    also from Cons.hyps
+    have "\<dots> \<subseteq>  dom (locals (store s2))"
+      by - (erule dom_locals_eval_mono_elim)
+    also from Cons
+    have "assigns (In3 es) \<subseteq> dom (locals (store s2))"
+      by rules
+    ultimately
+    have "assigns (In1l e) \<union> assigns (In3 es) \<subseteq> dom (locals (store s2))"
+      by (rule Un_least)
+    thus ?case
+      by simp
+  qed
+qed
+
+corollary assignsE_good_approx:
+  assumes
+      eval: "prg Env\<turnstile> s0 \<midarrow>e-\<succ>v\<rightarrow> s1" and
+    normal: "normal s1"
+  shows "assignsE e \<subseteq> dom (locals (store s1))"
+proof -
+from eval normal show ?thesis
+  by (rule assigns_good_approx [elim_format]) simp
+qed
+
+corollary assignsV_good_approx:
+  assumes
+     eval: "prg Env\<turnstile> s0 \<midarrow>v=\<succ>vf\<rightarrow> s1" and
+   normal: "normal s1"
+  shows "assignsV v \<subseteq> dom (locals (store s1))"
+proof -
+from eval normal show ?thesis
+  by (rule assigns_good_approx [elim_format]) simp
+qed
+   
+corollary assignsEs_good_approx:
+  assumes
+      eval: "prg Env\<turnstile> s0 \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s1" and
+    normal: "normal s1" 
+  shows "assignsEs es \<subseteq> dom (locals (store s1))"
+proof -
+from eval normal show ?thesis
+  by (rule assigns_good_approx [elim_format]) simp
+qed
+
+lemma constVal_eval: 
+ assumes const: "constVal e = Some c" and
+          eval: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s"
+  shows "v = c \<and> normal s"
+proof -
+  have  True and 
+        "\<And> c v s0 s. \<lbrakk> constVal e = Some c; G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s\<rbrakk> 
+                      \<Longrightarrow> v = c \<and> normal s"
+        and True and True
+  proof (induct  rule: var_expr_stmt.induct)
+    case NewC hence False by simp thus ?case ..
+  next
+    case NewA hence False by simp thus ?case ..
+  next
+    case Cast hence False by simp thus ?case ..
+  next
+    case Inst hence False by simp thus ?case ..
+  next
+    case (Lit val c v s0 s)
+    have "constVal (Lit val) = Some c" .
+    moreover
+    have "G\<turnstile>Norm s0 \<midarrow>Lit val-\<succ>v\<rightarrow> s" .
+    then obtain "v=val" and "normal s"
+      by cases simp
+    ultimately show "v=c \<and> normal s" by simp
+  next
+    case (UnOp unop e c v s0 s)
+    have const: "constVal (UnOp unop e) = Some c" .
+    then obtain ce where ce: "constVal e = Some ce" by simp
+    have "G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>v\<rightarrow> s" .
+    then obtain ve where ve: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>ve\<rightarrow> s" and
+                          v: "v = eval_unop unop ve" 
+      by cases simp
+    from ce ve
+    obtain eq_ve_ce: "ve=ce" and nrm_s: "normal s"
+      by (rule UnOp.hyps [elim_format]) rules
+    from eq_ve_ce const ce v 
+    have "v=c" 
+      by simp
+    from this nrm_s
+    show ?case ..
+  next
+    case (BinOp binop e1 e2 c v s0 s)
+    have const: "constVal (BinOp binop e1 e2) = Some c" .
+    then obtain c1 c2 where c1: "constVal e1 = Some c1" and
+                            c2: "constVal e2 = Some c2" and
+                             c: "c = eval_binop binop c1 c2"
+      by simp
+    have "G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>v\<rightarrow> s" .
+    then obtain v1 s1 v2
+      where v1: "G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1" and
+            v2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
+                               else In1r Skip)\<succ>\<rightarrow> (In1 v2, s)" and
+             v: "v = eval_binop binop v1 v2"
+      by cases simp
+    from c1 v1
+    obtain eq_v1_c1: "v1 = c1" and 
+             nrm_s1: "normal s1"
+      by (rule BinOp.hyps [elim_format]) rules
+    show ?case
+    proof (cases "need_second_arg binop v1")
+      case True
+      with v2 nrm_s1  obtain s1' 
+	where "G\<turnstile>Norm s1' \<midarrow>e2-\<succ>v2\<rightarrow> s" 
+	by (cases s1) simp
+      with c2 obtain "v2 = c2" "normal s"
+	by (rule BinOp.hyps [elim_format]) rules
+      with c c1 c2 eq_v1_c1 v 
+      show ?thesis by simp
+    next
+      case False
+      with nrm_s1 v2
+      have "s=s1"
+	by (cases s1) (auto elim!: eval_elim_cases)
+      moreover
+      from False c v eq_v1_c1  
+      have "v = c"
+	by (simp add: eval_binop_arg2_indep)
+      ultimately 
+      show ?thesis
+	using nrm_s1 by simp
+    qed  (* hallo ehco *)
+  next
+    case Super hence False by simp thus ?case ..
+  next
+    case Acc hence False by simp thus ?case ..
+  next
+    case Ass hence False by simp thus ?case ..
+  next
+    case (Cond b e1 e2 c v s0 s)
+    have c: "constVal (b ? e1 : e2) = Some c" .
+    then obtain cb c1 c2 where
+      cb: "constVal b  = Some cb" and
+      c1: "constVal e1 = Some c1" and
+      c2: "constVal e2 = Some c2"
+      by (auto split: bool.splits)
+    have "G\<turnstile>Norm s0 \<midarrow>b ? e1 : e2-\<succ>v\<rightarrow> s" .
+    then obtain vb s1
+      where     vb: "G\<turnstile>Norm s0 \<midarrow>b-\<succ>vb\<rightarrow> s1" and
+            eval_v: "G\<turnstile>s1 \<midarrow>(if the_Bool vb then e1 else e2)-\<succ>v\<rightarrow> s"
+      by cases simp 
+    from cb vb
+    obtain eq_vb_cb: "vb = cb" and nrm_s1: "normal s1"
+      by (rule Cond.hyps [elim_format]) rules 
+    show ?case
+    proof (cases "the_Bool vb")
+      case True
+      with c cb c1 eq_vb_cb 
+      have "c = c1"
+	by simp
+      moreover
+      from True eval_v nrm_s1 obtain s1' 
+	where "G\<turnstile>Norm s1' \<midarrow>e1-\<succ>v\<rightarrow> s"
+	by (cases s1) simp
+      with c1 obtain "c1 = v" "normal s"
+	by (rule Cond.hyps [elim_format]) rules 
+      ultimately show ?thesis by simp
+    next
+      case False
+      with c cb c2 eq_vb_cb 
+      have "c = c2"
+	by simp
+      moreover
+      from False eval_v nrm_s1 obtain s1' 
+	where "G\<turnstile>Norm s1' \<midarrow>e2-\<succ>v\<rightarrow> s"
+	by (cases s1) simp
+      with c2 obtain "c2 = v" "normal s"
+	by (rule Cond.hyps [elim_format]) rules 
+      ultimately show ?thesis by simp
+    qed
+  next
+    case Call hence False by simp thus ?case ..
+  qed simp_all
+  with const eval
+  show ?thesis
+    by rules
+qed
+  
+lemmas constVal_eval_elim = constVal_eval [THEN conjE]
+
+lemma eval_unop_type: 
+  "typeof dt (eval_unop unop v) = Some (PrimT (unop_type unop))"
+  by (cases unop) simp_all
+
+lemma eval_binop_type:
+  "typeof dt (eval_binop binop v1 v2) = Some (PrimT (binop_type binop))"
+  by (cases binop) simp_all
+
+lemma constVal_Boolean: 
+ assumes const: "constVal e = Some c" and
+            wt: "Env\<turnstile>e\<Colon>-PrimT Boolean"
+  shows "typeof empty_dt c = Some (PrimT Boolean)"
+proof - 
+  have True and 
+       "\<And> c. \<lbrakk>constVal e = Some c;Env\<turnstile>e\<Colon>-PrimT Boolean\<rbrakk> 
+                \<Longrightarrow> typeof empty_dt c = Some (PrimT Boolean)"
+       and True and True 
+  proof (induct rule: var_expr_stmt.induct)
+    case NewC hence False by simp thus ?case ..
+  next
+    case NewA hence False by simp thus ?case ..
+  next
+    case Cast hence False by simp thus ?case ..
+  next
+    case Inst hence False by simp thus ?case ..
+  next
+    case (Lit v c)
+    have "constVal (Lit v) = Some c" .
+    hence "c=v" by simp
+    moreover have "Env\<turnstile>Lit v\<Colon>-PrimT Boolean" .
+    hence "typeof empty_dt v = Some (PrimT Boolean)"
+      by cases simp
+    ultimately show ?case by simp
+  next
+    case (UnOp unop e c)
+    have "Env\<turnstile>UnOp unop e\<Colon>-PrimT Boolean" .
+    hence "Boolean = unop_type unop" by cases simp
+    moreover have "constVal (UnOp unop e) = Some c" .
+    then obtain ce where "c = eval_unop unop ce" by auto
+    ultimately show ?case by (simp add: eval_unop_type)
+  next
+    case (BinOp binop e1 e2 c)
+    have "Env\<turnstile>BinOp binop e1 e2\<Colon>-PrimT Boolean" .
+    hence "Boolean = binop_type binop" by cases simp
+    moreover have "constVal (BinOp binop e1 e2) = Some c" .
+    then obtain c1 c2 where "c = eval_binop binop c1 c2" by auto
+    ultimately show ?case by (simp add: eval_binop_type)
+  next
+    case Super hence False by simp thus ?case ..
+  next
+    case Acc hence False by simp thus ?case ..
+  next
+    case Ass hence False by simp thus ?case ..
+  next
+    case (Cond b e1 e2 c)
+    have c: "constVal (b ? e1 : e2) = Some c" .
+    then obtain cb c1 c2 where
+      cb: "constVal b  = Some cb" and
+      c1: "constVal e1 = Some c1" and
+      c2: "constVal e2 = Some c2"
+      by (auto split: bool.splits)
+    have wt: "Env\<turnstile>b ? e1 : e2\<Colon>-PrimT Boolean" .
+    then
+    obtain T1 T2
+      where "Env\<turnstile>b\<Colon>-PrimT Boolean" and
+            wt_e1: "Env\<turnstile>e1\<Colon>-PrimT Boolean" and
+            wt_e2: "Env\<turnstile>e2\<Colon>-PrimT Boolean"
+      by cases (auto dest: widen_Boolean2)
+    show ?case
+    proof (cases "the_Bool cb")
+      case True
+      from c1 wt_e1 
+      have  "typeof empty_dt c1 = Some (PrimT Boolean)" 
+	by (rule Cond.hyps)
+      with True c cb c1 show ?thesis by simp
+    next
+      case False
+      from c2 wt_e2 
+      have "typeof empty_dt c2 = Some (PrimT Boolean)" 
+	by (rule Cond.hyps)
+      with False c cb c2 show ?thesis by simp
+    qed
+  next
+    case Call hence False by simp thus ?case ..
+  qed simp_all
+  with const wt
+  show ?thesis
+    by rules
+qed	
+      
+lemma assigns_if_good_approx:
+  assumes
+     eval: "prg Env\<turnstile> s0 \<midarrow>e-\<succ>b\<rightarrow> s1"   and
+   normal: "normal s1" and
+     bool: "Env\<turnstile> e\<Colon>-PrimT Boolean"
+  shows "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+proof -
+  -- {* To properly perform induction on the evaluation relation we have to
+        generalize the lemma to terms not only expressions. *}
+  {fix t val
+   assume eval': "prg Env\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (val,s1)"  
+   assume bool': "Env\<turnstile> t\<Colon>Inl (PrimT Boolean)"
+   assume expr:  "\<exists> expr. t=In1l expr"
+   have "assigns_if (the_Bool (the_In1 val)) (the_In1l t) 
+                \<subseteq> dom (locals (store s1))"
+   using eval' normal bool' expr
+   proof (induct)
+     case Abrupt thus ?case by simp
+   next
+     case (NewC C a s0 s1 s2)
+     have "Env\<turnstile>NewC C\<Colon>-PrimT Boolean" .
+     hence False 
+       by cases simp
+     thus ?case ..
+   next
+     case (NewA T a e i s0 s1 s2 s3)
+     have "Env\<turnstile>New T[e]\<Colon>-PrimT Boolean" .
+     hence False 
+       by cases simp
+     thus ?case ..
+   next
+     case (Cast T e s0 s1 s2 b)
+     have s2: "s2 = abupd (raise_if (\<not> prg Env,snd s1\<turnstile>b fits T) ClassCast) s1".
+     have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))" 
+     proof -
+       have "normal s2" .
+       with s2 have "normal s1"
+	 by (cases s1) simp
+       moreover
+       have "Env\<turnstile>Cast T e\<Colon>-PrimT Boolean" .
+       hence "Env\<turnstile>e\<Colon>- PrimT Boolean" 
+	 by (cases) (auto dest: cast_Boolean2)
+       ultimately show ?thesis 
+	 by (rule Cast.hyps [elim_format]) auto
+     qed
+     also from s2 
+     have "\<dots> \<subseteq> dom (locals (store s2))"
+       by simp
+     finally show ?case by simp
+   next
+     case (Inst T b e s0 s1 v)
+     have "prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1" and "normal s1" .
+     hence "assignsE e \<subseteq> dom (locals (store s1))"
+       by (rule assignsE_good_approx)
+     thus ?case
+       by simp
+   next
+     case (Lit s v)
+     have "Env\<turnstile>Lit v\<Colon>-PrimT Boolean" .
+     hence "typeof empty_dt v = Some (PrimT Boolean)"
+       by cases simp
+     then obtain b where "v=Bool b"
+       by (cases v) (simp_all add: empty_dt_def)
+     thus ?case
+       by simp
+   next
+     case (UnOp e s0 s1 unop v)
+     have bool: "Env\<turnstile>UnOp unop e\<Colon>-PrimT Boolean" .
+     hence bool_e: "Env\<turnstile>e\<Colon>-PrimT Boolean" 
+       by cases (cases unop,simp_all)
+     show ?case
+     proof (cases "constVal (UnOp unop e)")
+       case None
+       have "normal s1" .
+       moreover note bool_e
+       ultimately have "assigns_if (the_Bool v) e \<subseteq> dom (locals (store s1))"
+	 by (rule UnOp.hyps [elim_format]) auto
+       moreover
+       from bool have "unop = UNot"
+	 by cases (cases unop, simp_all)
+       moreover note None
+       ultimately 
+       have "assigns_if (the_Bool (eval_unop unop v)) (UnOp unop e) 
+              \<subseteq> dom (locals (store s1))"
+	 by simp
+       thus ?thesis by simp
+     next
+       case (Some c)
+       moreover
+       have "prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1" .
+       hence "prg Env\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>eval_unop unop v\<rightarrow> s1" 
+	 by (rule eval.UnOp)
+       with Some
+       have "eval_unop unop v=c"
+	 by (rule constVal_eval_elim) simp
+       moreover
+       from Some bool
+       obtain b where "c=Bool b"
+	 by (rule constVal_Boolean [elim_format])
+	    (cases c, simp_all add: empty_dt_def)
+       ultimately
+       have "assigns_if (the_Bool (eval_unop unop v)) (UnOp unop e) = {}"
+	 by simp
+       thus ?thesis by simp
+     qed
+   next
+     case (BinOp binop e1 e2 s0 s1 s2 v1 v2)
+     have bool: "Env\<turnstile>BinOp binop e1 e2\<Colon>-PrimT Boolean" .
+     show ?case
+     proof (cases "constVal (BinOp binop e1 e2)") 
+       case (Some c)
+       moreover
+       from BinOp.hyps
+       have
+	 "prg Env\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>eval_binop binop v1 v2\<rightarrow> s2"
+	 by - (rule eval.BinOp) 
+       with Some
+       have "eval_binop binop v1 v2=c"
+	 by (rule constVal_eval_elim) simp
+       moreover
+       from Some bool
+       obtain b where "c = Bool b"
+	 by (rule constVal_Boolean [elim_format])
+	    (cases c, simp_all add: empty_dt_def)
+       ultimately
+       have "assigns_if (the_Bool (eval_binop binop v1 v2)) (BinOp binop e1 e2) 
+             = {}"
+	 by simp 
+       thus ?thesis by simp
+     next
+       case None
+       show ?thesis
+       proof (cases "binop=CondAnd \<or> binop=CondOr")
+	 case True
+	 from bool obtain bool_e1: "Env\<turnstile>e1\<Colon>-PrimT Boolean" and
+                          bool_e2: "Env\<turnstile>e2\<Colon>-PrimT Boolean"
+	   using True by cases auto
+	 have "assigns_if (the_Bool v1) e1 \<subseteq> dom (locals (store s1))"
+	 proof -
+	   from BinOp have "normal s1"
+	     by - (erule eval_no_abrupt_lemma [rule_format])
+	   from this bool_e1
+	   show ?thesis
+	     by (rule BinOp.hyps [elim_format]) auto
+	 qed
+	 also
+	 from BinOp.hyps
+	 have "\<dots> \<subseteq> dom (locals (store s2))"
+	   by - (erule dom_locals_eval_mono_elim,simp)
+	 finally
+	 have e1_s2: "assigns_if (the_Bool v1) e1 \<subseteq> dom (locals (store s2))".
+	 from True show ?thesis
+	 proof
+	   assume condAnd: "binop = CondAnd"
+	   show ?thesis
+	   proof (cases "the_Bool (eval_binop binop v1 v2)")
+	     case True
+	     with condAnd 
+	     have need_second: "need_second_arg binop v1"
+	       by (simp add: need_second_arg_def)
+	     have "normal s2" . 
+	     hence "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
+	       by (rule BinOp.hyps [elim_format]) 
+                  (simp add: need_second bool_e2)+
+	     with e1_s2
+	     have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
+		    \<subseteq> dom (locals (store s2))"
+	       by (rule Un_least)
+	     with True condAnd None show ?thesis
+	       by simp
+	   next
+	     case False
+	     note binop_False = this
+	     show ?thesis
+	     proof (cases "need_second_arg binop v1")
+	       case True
+	       with binop_False condAnd
+	       obtain "the_Bool v1=True" and "the_Bool v2 = False"
+		 by (simp add: need_second_arg_def)
+	       moreover
+	       have "normal s2" . 
+	       hence "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
+		 by (rule BinOp.hyps [elim_format]) (simp add: True bool_e2)+
+	       with e1_s2
+	       have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
+		      \<subseteq> dom (locals (store s2))"
+		 by (rule Un_least)
+	       moreover note binop_False condAnd None
+	       ultimately show ?thesis
+		 by auto
+	     next
+	       case False
+	       with binop_False condAnd
+	       have "the_Bool v1=False"
+		 by (simp add: need_second_arg_def)
+	       with e1_s2 
+	       show ?thesis
+		 using binop_False condAnd None by auto
+	     qed
+	   qed
+	 next
+	   assume condOr: "binop = CondOr"
+	   show ?thesis
+	   proof (cases "the_Bool (eval_binop binop v1 v2)")
+	     case False
+	     with condOr 
+	     have need_second: "need_second_arg binop v1"
+	       by (simp add: need_second_arg_def)
+	     have "normal s2" . 
+	     hence "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
+	       by (rule BinOp.hyps [elim_format]) 
+                  (simp add: need_second bool_e2)+
+	     with e1_s2
+	     have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
+		    \<subseteq> dom (locals (store s2))"
+	       by (rule Un_least)
+	     with False condOr None show ?thesis
+	       by simp
+	   next
+	     case True
+	     note binop_True = this
+	     show ?thesis
+	     proof (cases "need_second_arg binop v1")
+	       case True
+	       with binop_True condOr
+	       obtain "the_Bool v1=False" and "the_Bool v2 = True"
+		 by (simp add: need_second_arg_def)
+	       moreover
+	       have "normal s2" . 
+	       hence "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
+		 by (rule BinOp.hyps [elim_format]) (simp add: True bool_e2)+
+	       with e1_s2
+	       have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
+		      \<subseteq> dom (locals (store s2))"
+		 by (rule Un_least)
+	       moreover note binop_True condOr None
+	       ultimately show ?thesis
+		 by auto
+	     next
+	       case False
+	       with binop_True condOr
+	       have "the_Bool v1=True"
+		 by (simp add: need_second_arg_def)
+	       with e1_s2 
+	       show ?thesis
+		 using binop_True condOr None by auto
+	     qed
+	   qed
+	 qed  
+       next
+	 case False
+	 have "\<not> (binop = CondAnd \<or> binop = CondOr)" .
+	 from BinOp.hyps
+	 have
+	  "prg Env\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>eval_binop binop v1 v2\<rightarrow> s2"
+	   by - (rule eval.BinOp)  
+	 moreover have "normal s2" .
+	 ultimately
+	 have "assignsE (BinOp binop e1 e2) \<subseteq> dom (locals (store s2))"
+	   by (rule assignsE_good_approx)
+	 with False None
+	 show ?thesis 
+	   by simp
+       qed
+     qed
+   next     
+     case Super 
+     have "Env\<turnstile>Super\<Colon>-PrimT Boolean" .
+     hence False 
+       by cases simp
+     thus ?case ..
+   next
+     case (Acc f s0 s1 v va)
+     have "prg Env\<turnstile>Norm s0 \<midarrow>va=\<succ>(v, f)\<rightarrow> s1"  and "normal s1".
+     hence "assignsV va \<subseteq> dom (locals (store s1))"
+       by (rule assignsV_good_approx)
+     thus ?case by simp
+   next
+     case (Ass e f s0 s1 s2 v va w)
+     hence "prg Env\<turnstile>Norm s0 \<midarrow>va := e-\<succ>v\<rightarrow> assign f v s2"
+       by - (rule eval.Ass)
+     moreover have "normal (assign f v s2)" .
+     ultimately 
+     have "assignsE (va := e) \<subseteq> dom (locals (store (assign f v s2)))"
+       by (rule assignsE_good_approx)
+     thus ?case by simp
+   next
+     case (Cond b e0 e1 e2 s0 s1 s2 v)
+     have "Env\<turnstile>e0 ? e1 : e2\<Colon>-PrimT Boolean" .
+     then obtain wt_e1: "Env\<turnstile>e1\<Colon>-PrimT Boolean" and
+                 wt_e2: "Env\<turnstile>e2\<Colon>-PrimT Boolean"
+       by cases (auto dest: widen_Boolean2)
+     have eval_e0: "prg Env\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1" .
+     have e0_s2: "assignsE e0 \<subseteq> dom (locals (store s2))"
+     proof -
+       note eval_e0 
+       moreover
+       have "normal s2" .
+       with Cond.hyps have "normal s1"
+	 by - (erule eval_no_abrupt_lemma [rule_format],simp)
+       ultimately
+       have "assignsE e0 \<subseteq> dom (locals (store s1))"
+	 by (rule assignsE_good_approx)
+       also
+       from Cond
+       have "\<dots> \<subseteq> dom (locals (store s2))"
+	 by - (erule dom_locals_eval_mono [elim_format],simp)
+       finally show ?thesis .
+     qed
+     show ?case
+     proof (cases "constVal e0")
+       case None
+       have "assigns_if (the_Bool v) e1 \<inter>  assigns_if (the_Bool v) e2 
+	      \<subseteq> dom (locals (store s2))"
+       proof (cases "the_Bool b")
+	 case True
+	 have "normal s2" .
+	 hence "assigns_if (the_Bool v) e1 \<subseteq> dom (locals (store s2))"    
+	   by (rule Cond.hyps [elim_format]) (simp_all add: wt_e1 True)
+	 thus ?thesis
+	   by blast
+       next
+	 case False
+	 have "normal s2" .
+	 hence "assigns_if (the_Bool v) e2 \<subseteq> dom (locals (store s2))"    
+	   by (rule Cond.hyps [elim_format]) (simp_all add: wt_e2 False)
+	 thus ?thesis
+	   by blast
+       qed
+       with e0_s2
+       have "assignsE e0 \<union> 
+             (assigns_if (the_Bool v) e1 \<inter>  assigns_if (the_Bool v) e2)
+	       \<subseteq> dom (locals (store s2))"
+	 by (rule Un_least)
+       with None show ?thesis
+	 by simp
+     next
+       case (Some c)
+       from this eval_e0 have eq_b_c: "b=c" 
+	 by (rule constVal_eval_elim) 
+       show ?thesis
+       proof (cases "the_Bool c")
+	 case True
+	 have "normal s2" .
+	 hence "assigns_if (the_Bool v) e1 \<subseteq> dom (locals (store s2))"    
+	   by (rule Cond.hyps [elim_format]) (simp_all add: eq_b_c True)
+	 with e0_s2
+	 have "assignsE e0 \<union> assigns_if (the_Bool v) e1  \<subseteq> \<dots>"
+	   by (rule Un_least)
+	 with Some True show ?thesis
+	   by simp
+       next
+	 case False
+	 have "normal s2" .
+	 hence "assigns_if (the_Bool v) e2 \<subseteq> dom (locals (store s2))"    
+	   by (rule Cond.hyps [elim_format]) (simp_all add: eq_b_c False)
+	 with e0_s2
+	 have "assignsE e0 \<union> assigns_if (the_Bool v) e2  \<subseteq> \<dots>"
+	   by (rule Un_least)
+	 with Some False show ?thesis
+	   by simp
+       qed
+     qed
+   next
+     case (Call D a accC args e mn mode pTs s0 s1 s2 s3 s3' s4 statT v vs)
+     hence
+     "prg Env\<turnstile>Norm s0 \<midarrow>({accC,statT,mode}e\<cdot>mn( {pTs}args))-\<succ>v\<rightarrow> 
+                       (set_lvars (locals (store s2)) s4)"
+       by - (rule eval.Call)
+     hence "assignsE ({accC,statT,mode}e\<cdot>mn( {pTs}args)) 
+              \<subseteq>  dom (locals (store ((set_lvars (locals (store s2))) s4)))"
+       by (rule assignsE_good_approx)
+     thus ?case by simp
+   next
+     case Methd show ?case by simp
+   next
+     case Body show ?case by simp
+   qed simp+ -- {* all the statements and variables *}
+ }
+ note generalized = this
+ from eval bool show ?thesis
+   by (rule generalized [elim_format]) simp+
+qed
+
+lemma assigns_if_good_approx': 
+  assumes   eval: "G\<turnstile>s0 \<midarrow>e-\<succ>b\<rightarrow> s1"  
+     and  normal: "normal s1" 
+     and    bool: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>- (PrimT Boolean)"
+  shows "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+proof -
+  from eval have "prg \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>s0 \<midarrow>e-\<succ>b\<rightarrow> s1" by simp
+  from this normal bool show ?thesis
+    by (rule assigns_if_good_approx)
+qed
+
+ 
+lemma subset_Intl: "A \<subseteq> C \<Longrightarrow> A \<inter> B \<subseteq> C" 
+  by blast
+
+lemma subset_Intr: "B \<subseteq> C \<Longrightarrow> A \<inter> B \<subseteq> C" 
+  by blast
+
+lemma da_good_approx: 
+  assumes  eval: "prg Env\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" and
+           wt: "Env\<turnstile>t\<Colon>T"     (is "?Wt Env t T") and
+           da: "Env\<turnstile> dom (locals (store s0)) \<guillemotright>t\<guillemotright> A"  (is "?Da Env s0 t A") and 
+           wf: "wf_prog (prg Env)" 
+     shows "(normal s1 \<longrightarrow> (nrm A \<subseteq>  dom (locals (store s1)))) \<and>
+            (\<forall> l. abrupt s1 = Some (Jump (Break l)) \<and> normal s0 
+                       \<longrightarrow> (brk A l \<subseteq> dom (locals (store s1)))) \<and>
+            (abrupt s1 = Some (Jump Ret) \<and> normal s0 
+                       \<longrightarrow> Result \<in> dom (locals (store s1)))"
+     (is "?NormalAssigned s1 A \<and> ?BreakAssigned s0 s1 A \<and> ?ResAssigned s0 s1")
+proof -
+  note inj_term_simps [simp]
+  obtain G where G: "prg Env = G" by (cases Env) simp
+  with eval have eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" by simp
+  from G wf have wf: "wf_prog G" by simp
+  let ?HypObj = "\<lambda> t s0 s1.
+      \<forall> Env T A. ?Wt Env t T \<longrightarrow>  ?Da Env s0 t A \<longrightarrow> prg Env = G 
+       \<longrightarrow> ?NormalAssigned s1 A \<and> ?BreakAssigned s0 s1 A \<and> ?ResAssigned  s0 s1"
+  -- {* Goal in object logic variant *} 
+  from eval
+  show "\<And> Env T A. \<lbrakk>?Wt Env t T; ?Da Env s0 t A; prg Env = G\<rbrakk> 
+        \<Longrightarrow> ?NormalAssigned s1 A \<and> ?BreakAssigned s0 s1 A \<and> ?ResAssigned s0 s1"
+        (is "PROP ?Hyp t s0 s1")
+  proof (induct)
+    case (Abrupt s t xc Env T A)
+    have da: "Env\<turnstile> dom (locals s) \<guillemotright>t\<guillemotright> A" using Abrupt.prems by simp 
+    have "?NormalAssigned (Some xc,s) A" 
+      by simp
+    moreover
+    have "?BreakAssigned (Some xc,s) (Some xc,s) A"
+      by simp
+    moreover have "?ResAssigned (Some xc,s) (Some xc,s)"
+      by simp
+    ultimately show ?case by (intro conjI)
+  next
+    case (Skip s Env T A)
+    have da: "Env\<turnstile> dom (locals (store (Norm s))) \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> A" 
+      using Skip.prems by simp
+    hence "nrm A = dom (locals (store (Norm s)))"
+      by (rule da_elim_cases) simp
+    hence "?NormalAssigned (Norm s) A"
+      by auto
+    moreover
+    have "?BreakAssigned (Norm s) (Norm s) A" 
+      by simp
+    moreover have "?ResAssigned (Norm s) (Norm s)"
+      by simp
+    ultimately show ?case by (intro conjI)
+  next
+    case (Expr e s0 s1 v Env T A)
+    from Expr.prems
+    show "?NormalAssigned s1 A \<and> ?BreakAssigned (Norm s0) s1 A 
+           \<and> ?ResAssigned (Norm s0) s1"
+      by (elim wt_elim_cases da_elim_cases) 
+         (rule Expr.hyps, auto)
+  next 
+    case (Lab c j s0 s1 Env T A)
+    have G: "prg Env = G" .
+    from Lab.prems
+    obtain C l where
+      da_c: "Env\<turnstile> dom (locals (snd (Norm s0))) \<guillemotright>\<langle>c\<rangle>\<guillemotright> C" and
+         A: "nrm A = nrm C \<inter> (brk C) l" "brk A = rmlab l (brk C)" and
+         j: "j = Break l"
+      by - (erule da_elim_cases, simp)
+    from Lab.prems
+    have wt_c: "Env\<turnstile>c\<Colon>\<surd>"
+      by - (erule wt_elim_cases, simp)
+    from wt_c da_c G and Lab.hyps
+    have norm_c: "?NormalAssigned s1 C" and 
+          brk_c: "?BreakAssigned (Norm s0) s1 C" and
+          res_c: "?ResAssigned (Norm s0) s1"
+      by simp_all
+    have "?NormalAssigned (abupd (absorb j) s1) A"
+    proof
+      assume normal: "normal (abupd (absorb j) s1)"
+      show "nrm A \<subseteq> dom (locals (store (abupd (absorb j) s1)))"
+      proof (cases "abrupt s1")
+	case None
+	with norm_c A 
+	show ?thesis
+	  by auto
+      next
+	case Some
+	with normal j 
+	have "abrupt s1 = Some (Jump (Break l))"
+	  by (auto dest: absorb_Some_NoneD)
+	with brk_c A
+  	show ?thesis
+	  by auto
+      qed
+    qed
+    moreover 
+    have "?BreakAssigned (Norm s0) (abupd (absorb j) s1) A"
+    proof -
+      {
+	fix l'
+	assume break: "abrupt (abupd (absorb j) s1) = Some (Jump (Break l'))"
+	with j 
+	have "l\<noteq>l'"
+	  by (cases s1) (auto dest!: absorb_Some_JumpD)
+	hence "(rmlab l (brk C)) l'= (brk C) l'"
+	  by (simp)
+	with break brk_c A
+	have 
+          "(brk A l' \<subseteq> dom (locals (store (abupd (absorb j) s1))))"
+	  by (cases s1) auto
+      }
+      then show ?thesis
+	by simp
+    qed
+    moreover
+    from res_c have "?ResAssigned (Norm s0) (abupd (absorb j) s1)"
+      by (cases s1) (simp add: absorb_def)
+    ultimately show ?case by (intro conjI)
+  next
+    case (Comp c1 c2 s0 s1 s2 Env T A)
+    have G: "prg Env = G" .
+    from Comp.prems
+    obtain C1 C2
+      where da_c1: "Env\<turnstile> dom (locals (snd (Norm s0))) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1" and 
+            da_c2: "Env\<turnstile> nrm C1 \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2" and
+                A: "nrm A = nrm C2" "brk A = (brk C1) \<Rightarrow>\<inter> (brk C2)"
+      by (elim da_elim_cases) simp
+    from Comp.prems
+    obtain wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and
+           wt_c2: "Env\<turnstile>c2\<Colon>\<surd>"
+      by (elim wt_elim_cases) simp
+    have "PROP ?Hyp (In1r c1) (Norm s0) s1" .
+    with wt_c1 da_c1 G 
+    obtain nrm_c1: "?NormalAssigned s1 C1" and 
+           brk_c1: "?BreakAssigned (Norm s0) s1 C1" and
+           res_c1: "?ResAssigned (Norm s0) s1"
+      by simp
+    show ?case
+    proof (cases "normal s1")
+      case True
+      with nrm_c1 have "nrm C1 \<subseteq> dom (locals (snd s1))" by rules
+      with da_c2 obtain C2' 
+	where  da_c2': "Env\<turnstile> dom (locals (snd s1)) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
+               nrm_c2: "nrm C2 \<subseteq> nrm C2'"                  and
+               brk_c2: "\<forall> l. brk C2 l \<subseteq> brk C2' l"
+        by (rule da_weakenE) rules
+      have "PROP ?Hyp (In1r c2) s1 s2" .
+      with wt_c2 da_c2' G
+      obtain nrm_c2': "?NormalAssigned s2 C2'" and 
+             brk_c2': "?BreakAssigned s1 s2 C2'" and
+             res_c2 : "?ResAssigned s1 s2" 
+	by simp
+      from nrm_c2' nrm_c2 A
+      have "?NormalAssigned s2 A" 
+	by blast
+      moreover from brk_c2' brk_c2 A
+      have "?BreakAssigned s1 s2 A" 
+	by fastsimp
+      with True 
+      have "?BreakAssigned (Norm s0) s2 A" by simp
+      moreover from res_c2 True
+      have "?ResAssigned (Norm s0) s2"
+	by simp
+      ultimately show ?thesis by (intro conjI)
+    next
+      case False
+      have "G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2" .
+      with False have eq_s1_s2: "s2=s1" by auto
+      with False have "?NormalAssigned s2 A" by blast
+      moreover
+      have "?BreakAssigned (Norm s0) s2 A"
+      proof (cases "\<exists> l. abrupt s1 = Some (Jump (Break l))")
+	case True
+	then obtain l where l: "abrupt s1 = Some (Jump (Break l))" .. 
+	with brk_c1
+	have "brk C1 l \<subseteq> dom (locals (store s1))"
+	  by simp
+	with A eq_s1_s2 
+	have "brk A l \<subseteq> dom (locals (store s2))" 
+	  by auto
+	with l eq_s1_s2
+	show ?thesis by simp
+      next
+	case False
+	with eq_s1_s2 show ?thesis by simp
+      qed
+      moreover from False res_c1 eq_s1_s2
+      have "?ResAssigned (Norm s0) s2"
+	by simp
+      ultimately show ?thesis by (intro conjI)
+    qed
+  next
+    case (If b c1 c2 e s0 s1 s2 Env T A)
+    have G: "prg Env = G" .
+    with If.hyps have eval_e: "prg Env \<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1" by simp
+    from If.prems
+    obtain E C1 C2 where
+      da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E" and
+     da_c1: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                   \<union> assigns_if True e) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1" and
+     da_c2: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                   \<union> assigns_if False e) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2" and
+     A: "nrm A = nrm C1 \<inter> nrm C2"  "brk A = brk C1 \<Rightarrow>\<inter> brk C2"
+      by (elim da_elim_cases) 
+    from If.prems
+    obtain 
+       wt_e:  "Env\<turnstile>e\<Colon>- PrimT Boolean" and 
+       wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and
+       wt_c2: "Env\<turnstile>c2\<Colon>\<surd>" 
+      by (elim wt_elim_cases)
+    from wt_e da_e G
+    obtain nrm_s1: "?NormalAssigned s1 E" and 
+           brk_s1: "?BreakAssigned (Norm s0) s1 E" and
+           res_s1: "?ResAssigned (Norm s0) s1"
+      by (elim If.hyps [elim_format]) simp+
+    from If.hyps have
+     s0_s1:"dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by (elim dom_locals_eval_mono_elim)
+    show ?case
+    proof (cases "normal s1")
+      case True
+      note normal_s1 = this 
+      show ?thesis
+      proof (cases "the_Bool b")
+	case True
+	from eval_e normal_s1 wt_e
+	have "assigns_if True e \<subseteq> dom (locals (store s1))"
+	  by (rule assigns_if_good_approx [elim_format]) (simp add: True)
+	with s0_s1
+	have "dom (locals (store ((Norm s0)::state))) \<union> assigns_if True e \<subseteq> \<dots>"
+	  by (rule Un_least)
+	with da_c1 obtain C1'
+	  where da_c1': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
+	        nrm_c1: "nrm C1 \<subseteq> nrm C1'"                  and
+                brk_c1: "\<forall> l. brk C1 l \<subseteq> brk C1' l"
+          by (rule da_weakenE) rules
+	from If.hyps True have "PROP ?Hyp (In1r c1) s1 s2" by simp
+	with wt_c1 da_c1'
+	obtain nrm_c1': "?NormalAssigned s2 C1'" and 
+               brk_c1': "?BreakAssigned s1 s2 C1'" and
+               res_c1: "?ResAssigned s1 s2"
+	  using G by simp
+	from nrm_c1' nrm_c1 A
+	have "?NormalAssigned s2 A" 
+	  by blast
+	moreover from brk_c1' brk_c1 A
+	have "?BreakAssigned s1 s2 A" 
+	  by fastsimp
+	with normal_s1
+	have "?BreakAssigned (Norm s0) s2 A" by simp
+	moreover from res_c1 normal_s1 have "?ResAssigned (Norm s0) s2"
+	  by simp
+	ultimately show ?thesis by (intro conjI)
+      next
+	case False
+	from eval_e normal_s1 wt_e
+	have "assigns_if False e \<subseteq> dom (locals (store s1))"
+	  by (rule assigns_if_good_approx [elim_format]) (simp add: False)
+	with s0_s1
+	have "dom (locals (store ((Norm s0)::state)))\<union> assigns_if False e \<subseteq> \<dots>"
+	  by (rule Un_least)
+	with da_c2 obtain C2'
+	  where da_c2': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
+	        nrm_c2: "nrm C2 \<subseteq> nrm C2'"                  and
+                brk_c2: "\<forall> l. brk C2 l \<subseteq> brk C2' l"
+          by (rule da_weakenE) rules
+	from If.hyps False have "PROP ?Hyp (In1r c2) s1 s2" by simp
+	with wt_c2 da_c2'
+	obtain nrm_c2': "?NormalAssigned s2 C2'" and 
+               brk_c2': "?BreakAssigned s1 s2 C2'" and
+	       res_c2: "?ResAssigned s1 s2"
+	  using G by simp
+	from nrm_c2' nrm_c2 A
+	have "?NormalAssigned s2 A" 
+	  by blast
+	moreover from brk_c2' brk_c2 A
+	have "?BreakAssigned s1 s2 A" 
+	  by fastsimp
+	with normal_s1
+	have "?BreakAssigned (Norm s0) s2 A" by simp
+	moreover from res_c2 normal_s1 have "?ResAssigned (Norm s0) s2"
+	  by simp
+	ultimately show ?thesis by (intro conjI)
+      qed
+    next
+      case False 
+      then obtain abr where abr: "abrupt s1 = Some abr"
+	by (cases s1) auto
+      moreover
+      from eval_e _ wt_e have "\<And> l. abrupt s1 \<noteq> Some (Jump (Break l))"
+	by (rule eval_expression_no_jump) (simp_all add: G wf)
+      moreover
+      have "s2 = s1"
+      proof -
+	have "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2" .
+        with abr show ?thesis  
+	  by (cases s1) simp
+      qed
+      ultimately show ?thesis using res_s1 by simp
+    qed
+  next
+-- {* 
+\par
+*} (* dummy text command to break paragraph for latex;
+              large paragraphs exhaust memory of debian pdflatex *)
+    case (Loop b c e l s0 s1 s2 s3 Env T A)
+    have G: "prg Env = G" .
+    with Loop.hyps have eval_e: "prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1" 
+      by (simp (no_asm_simp))
+    from Loop.prems
+    obtain E C where
+      da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E" and
+      da_c: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                   \<union> assigns_if True e) \<guillemotright>\<langle>c\<rangle>\<guillemotright> C" and
+     A: "nrm A = nrm C \<inter> 
+              (dom (locals (store ((Norm s0)::state))) \<union> assigns_if False e)"
+        "brk A = brk C"
+      by (elim da_elim_cases) 
+    from Loop.prems
+    obtain 
+       wt_e: "Env\<turnstile>e\<Colon>-PrimT Boolean" and 
+       wt_c: "Env\<turnstile>c\<Colon>\<surd>"
+      by (elim wt_elim_cases)
+    from wt_e da_e G
+    obtain res_s1: "?ResAssigned (Norm s0) s1"  
+      by (elim Loop.hyps [elim_format]) simp+
+    from Loop.hyps have
+      s0_s1:"dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by (elim dom_locals_eval_mono_elim)
+    show ?case
+    proof (cases "normal s1")
+      case True
+      note normal_s1 = this
+      show ?thesis
+      proof (cases "the_Bool b")
+	case True  
+	with Loop.hyps obtain
+          eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" and 
+          eval_while: "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
+	  by simp
+	from Loop.hyps True 
+	have "?HypObj (In1r c) s1 s2" by simp
+	note hyp_c = this [rule_format]
+	from Loop.hyps True
+	have "?HypObj (In1r (l\<bullet> While(e) c)) (abupd (absorb (Cont l)) s2) s3"
+	  by simp
+	note hyp_while = this [rule_format]
+	from eval_e normal_s1 wt_e
+	have "assigns_if True e \<subseteq> dom (locals (store s1))"
+	  by (rule assigns_if_good_approx [elim_format]) (simp add: True)
+	with s0_s1
+	have "dom (locals (store ((Norm s0)::state))) \<union> assigns_if True e \<subseteq> \<dots>"
+	  by (rule Un_least)
+	with da_c obtain C'
+	  where da_c': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'" and
+	     nrm_C_C': "nrm C \<subseteq> nrm C'"                  and
+             brk_C_C': "\<forall> l. brk C l \<subseteq> brk C' l"
+          by (rule da_weakenE) rules
+	from hyp_c wt_c da_c'
+	obtain nrm_C': "?NormalAssigned s2 C'" and 
+          brk_C': "?BreakAssigned s1 s2 C'" and
+          res_s2: "?ResAssigned s1 s2"
+	  using G by simp
+	show ?thesis
+	proof (cases "normal s2 \<or> abrupt s2 = Some (Jump (Cont l))")
+	  case True
+	  from Loop.prems obtain
+	    wt_while: "Env\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T" and
+            da_while: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) 
+                           \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A"
+	    by simp
+	  have "dom (locals (store ((Norm s0)::state)))
+                  \<subseteq> dom (locals (store (abupd (absorb (Cont l)) s2)))"
+	  proof -
+	    note s0_s1
+	    also from eval_c 
+	    have "dom (locals (store s1)) \<subseteq> dom (locals (store s2))"
+	      by (rule dom_locals_eval_mono_elim)
+	    also have "\<dots> \<subseteq> dom (locals (store (abupd (absorb (Cont l)) s2)))"
+	      by simp
+	    finally show ?thesis .
+	  qed
+	  with da_while obtain A'
+	    where 
+	    da_while': "Env\<turnstile> dom (locals (store (abupd (absorb (Cont l)) s2))) 
+                       \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A'" 
+	    and  nrm_A_A': "nrm A \<subseteq> nrm A'"                  
+            and  brk_A_A': "\<forall> l. brk A l \<subseteq> brk A' l"
+	    by (rule da_weakenE) simp
+	  with wt_while hyp_while
+	  obtain nrm_A': "?NormalAssigned s3 A'" and
+                 brk_A': "?BreakAssigned (abupd (absorb (Cont l)) s2) s3 A'" and
+                 res_s3: "?ResAssigned (abupd (absorb (Cont l)) s2) s3"
+	    using G by simp 
+	  from nrm_A_A' nrm_A'
+	  have "?NormalAssigned s3 A" 
+	    by blast
+	  moreover 
+	  have "?BreakAssigned (Norm s0) s3 A" 
+	  proof -
+	    from brk_A_A' brk_A' 
+	    have "?BreakAssigned (abupd (absorb (Cont l)) s2) s3 A" 
+	      by fastsimp
+	    moreover
+	    from True have "normal (abupd (absorb (Cont l)) s2)"
+	      by (cases s2) auto
+	    ultimately show ?thesis
+	      by simp
+	  qed
+	  moreover from res_s3 True have "?ResAssigned (Norm s0) s3"
+	    by auto
+	  ultimately show ?thesis by (intro conjI)
+	next
+	  case False
+	  then obtain abr where 
+	    "abrupt s2 = Some abr" and
+	    "abrupt (abupd (absorb (Cont l)) s2) = Some abr"
+	    by auto
+	  with eval_while 
+	  have eq_s3_s2: "s3=s2"
+	    by auto
+	  with nrm_C_C' nrm_C' A
+	  have "?NormalAssigned s3 A"
+	    by fastsimp
+	  moreover
+	  from eq_s3_s2 brk_C_C' brk_C' normal_s1 A
+	  have "?BreakAssigned (Norm s0) s3 A"
+	    by fastsimp
+	  moreover 
+	  from eq_s3_s2 res_s2 normal_s1 have "?ResAssigned (Norm s0) s3"
+	    by simp
+	  ultimately show ?thesis by (intro conjI)
+	qed
+      next
+	case False
+	with Loop.hyps have eq_s3_s1: "s3=s1"
+	  by simp
+	from eq_s3_s1 res_s1 
+	have res_s3: "?ResAssigned (Norm s0) s3"
+	  by simp
+	from eval_e True wt_e
+	have "assigns_if False e \<subseteq> dom (locals (store s1))"
+	  by (rule assigns_if_good_approx [elim_format]) (simp add: False)
+	with s0_s1
+	have "dom (locals (store ((Norm s0)::state)))\<union>assigns_if False e \<subseteq> \<dots>"
+	  by (rule Un_least)
+	hence "nrm C \<inter> 
+               (dom (locals (store ((Norm s0)::state))) \<union> assigns_if False e)
+               \<subseteq> dom (locals (store s1))"
+	  by (rule subset_Intr)
+	with normal_s1 A eq_s3_s1
+	have "?NormalAssigned s3 A"
+	  by simp
+	moreover
+	from normal_s1 eq_s3_s1
+	have "?BreakAssigned (Norm s0) s3 A"
+	  by simp
+	moreover note res_s3
+	ultimately show ?thesis by (intro conjI)
+      qed
+    next
+      case False
+      then obtain abr where abr: "abrupt s1 = Some abr"
+	by (cases s1) auto
+      moreover
+      from eval_e _ wt_e have no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+	by (rule eval_expression_no_jump) (simp_all add: wf G)
+      moreover
+      have eq_s3_s1: "s3=s1"
+      proof (cases "the_Bool b")
+	case True  
+	with Loop.hyps obtain
+          eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" and 
+          eval_while: "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
+	  by simp
+	from eval_c abr have "s2=s1" by auto
+	moreover from calculation no_jmp have "abupd (absorb (Cont l)) s2=s2"
+	  by (cases s1) (simp add: absorb_def)
+	ultimately show ?thesis
+	  using eval_while abr
+	  by auto
+      next
+	case False
+	with Loop.hyps show ?thesis by simp
+      qed
+      moreover
+      from eq_s3_s1 res_s1 
+      have res_s3: "?ResAssigned (Norm s0) s3"
+	by simp
+      ultimately show ?thesis
+	by simp 
+    qed
+  next 
+    case (Jmp j s Env T A)
+    have "?NormalAssigned (Some (Jump j),s) A" by simp
+    moreover
+    from Jmp.prems
+    obtain ret: "j = Ret \<longrightarrow> Result \<in> dom (locals (store (Norm s)))" and
+           brk: "brk A = (case j of
+                           Break l \<Rightarrow> \<lambda> k. if k=l 
+                                     then dom (locals (store ((Norm s)::state)))
+                                     else UNIV     
+                         | Cont l  \<Rightarrow> \<lambda> k. UNIV
+                         | Ret     \<Rightarrow> \<lambda> k. UNIV)"
+      by (elim da_elim_cases) simp
+    from brk have "?BreakAssigned (Norm s) (Some (Jump j),s) A"
+      by simp
+    moreover from ret have "?ResAssigned (Norm s) (Some (Jump j),s)"
+      by simp
+    ultimately show ?case by (intro conjI)
+  next
+    case (Throw a e s0 s1 Env T A)
+    have G: "prg Env = G" .
+    from Throw.prems obtain E where 
+      da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E"
+      by (elim da_elim_cases)
+    from Throw.prems
+      obtain eT where wt_e: "Env\<turnstile>e\<Colon>-eT"
+	by (elim wt_elim_cases)
+    have "?NormalAssigned (abupd (throw a) s1) A"
+      by (cases s1) (simp add: throw_def)
+    moreover
+    have "?BreakAssigned (Norm s0) (abupd (throw a) s1) A"
+    proof -
+      from G Throw.hyps have eval_e: "prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1" 
+	by (simp (no_asm_simp))
+      from eval_e _ wt_e 
+      have "\<And> l. abrupt s1 \<noteq> Some (Jump (Break l))"
+	by (rule eval_expression_no_jump) (simp_all add: wf G) 
+      hence "\<And> l. abrupt (abupd (throw a) s1) \<noteq> Some (Jump (Break l))"
+	by (cases s1) (simp add: throw_def abrupt_if_def)
+      thus ?thesis
+	by simp
+    qed
+    moreover
+    from wt_e da_e G have "?ResAssigned (Norm s0) s1"
+      by (elim Throw.hyps [elim_format]) simp+
+    hence "?ResAssigned (Norm s0) (abupd (throw a) s1)"
+      by (cases s1) (simp add: throw_def abrupt_if_def)
+    ultimately show ?case by (intro conjI)
+  next
+    case (Try C c1 c2 s0 s1 s2 s3 vn Env T A)
+    have G: "prg Env = G" .
+    from Try.prems obtain C1 C2 where
+      da_c1: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1"  and
+      da_c2:
+       "Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>
+        \<turnstile> (dom (locals (store ((Norm s0)::state))) \<union> {VName vn}) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2" and
+      A: "nrm A = nrm C1 \<inter> nrm C2" "brk A = brk C1 \<Rightarrow>\<inter> brk C2"
+      by (elim da_elim_cases) simp
+    from Try.prems obtain 
+      wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and
+      wt_c2: "Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>\<turnstile>c2\<Colon>\<surd>"
+      by (elim wt_elim_cases)
+    have sxalloc: "prg Env\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" using Try.hyps G 
+      by (simp (no_asm_simp))
+    have "PROP ?Hyp (In1r c1) (Norm s0) s1" .
+    with wt_c1 da_c1 G
+    obtain nrm_C1: "?NormalAssigned s1 C1" and
+           brk_C1: "?BreakAssigned (Norm s0) s1 C1" and
+           res_s1: "?ResAssigned (Norm s0) s1"
+      by simp
+    show ?case
+    proof (cases "normal s1")
+      case True
+      with nrm_C1 have "nrm C1 \<inter> nrm C2 \<subseteq> dom (locals (store s1))"
+	by auto
+      moreover
+      have "s3=s1"
+      proof -
+	from sxalloc True have eq_s2_s1: "s2=s1"
+	  by (cases s1) (auto elim: sxalloc_elim_cases)
+	with True have "\<not>  G,s2\<turnstile>catch C"
+	  by (simp add: catch_def)
+	with Try.hyps have "s3=s2"
+	  by simp
+	with eq_s2_s1 show ?thesis by simp
+      qed
+      ultimately show ?thesis
+	using True A res_s1 by simp
+    next
+      case False
+      note not_normal_s1 = this
+      show ?thesis
+      proof (cases "\<exists> l. abrupt s1 = Some (Jump (Break l))")
+	case True
+	then obtain l where l: "abrupt s1 = Some (Jump (Break l))"
+	  by auto
+	with brk_C1 have "(brk C1 \<Rightarrow>\<inter> brk C2) l \<subseteq> dom (locals (store s1))"
+	  by auto
+	moreover have "s3=s1"
+	proof -
+	  from sxalloc l have eq_s2_s1: "s2=s1"
+	    by (cases s1) (auto elim: sxalloc_elim_cases)
+	  with l have "\<not>  G,s2\<turnstile>catch C"
+	    by (simp add: catch_def)
+	  with Try.hyps have "s3=s2"
+	    by simp
+	  with eq_s2_s1 show ?thesis by simp
+	qed
+	ultimately show ?thesis
+	  using l A res_s1 by simp
+      next
+	case False
+	note abrupt_no_break = this
+	show ?thesis
+	proof (cases "G,s2\<turnstile>catch C")
+	  case True
+	  with Try.hyps have "?HypObj (In1r c2) (new_xcpt_var vn s2) s3"
+            by simp
+          note hyp_c2 = this [rule_format]
+          have "(dom (locals (store ((Norm s0)::state))) \<union> {VName vn}) 
+                  \<subseteq> dom (locals (store (new_xcpt_var vn s2)))"
+          proof -
+            have "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1" .
+            hence "dom (locals (store ((Norm s0)::state))) 
+                    \<subseteq> dom (locals (store s1))"
+              by (rule dom_locals_eval_mono_elim)
+            also
+            from sxalloc
+            have "\<dots> \<subseteq> dom (locals (store s2))"
+              by (rule dom_locals_sxalloc_mono)
+            also 
+            have "\<dots> \<subseteq> dom (locals (store (new_xcpt_var vn s2)))" 
+              by (cases s2) (simp add: new_xcpt_var_def, blast) 
+            also
+            have "{VName vn} \<subseteq> \<dots>"
+              by (cases s2) simp
+            ultimately show ?thesis
+              by (rule Un_least)
+          qed
+          with da_c2 
+          obtain C2' where
+            da_C2': "Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>
+                     \<turnstile> dom (locals (store (new_xcpt_var vn s2))) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'"
+           and nrm_C2': "nrm C2 \<subseteq> nrm C2'" 
+           and brk_C2': "\<forall> l. brk C2 l \<subseteq> brk C2' l"
+            by (rule da_weakenE) simp
+          from wt_c2 da_C2' G and hyp_c2
+          obtain nrmAss_C2: "?NormalAssigned s3 C2'" and
+                 brkAss_C2: "?BreakAssigned (new_xcpt_var vn s2) s3 C2'" and
+                 resAss_s3: "?ResAssigned (new_xcpt_var vn s2) s3"
+            by simp
+          from nrmAss_C2 nrm_C2' A 
+          have "?NormalAssigned s3 A"
+            by auto
+          moreover
+          have "?BreakAssigned (Norm s0) s3 A"
+          proof -
+            from brkAss_C2 have "?BreakAssigned (Norm s0) s3 C2'"
+              by (cases s2) (auto simp add: new_xcpt_var_def)
+            with brk_C2' A show ?thesis 
+              by fastsimp
+          qed
+	  moreover
+	  from resAss_s3 have "?ResAssigned (Norm s0) s3"
+	    by (cases s2) ( simp add: new_xcpt_var_def)
+          ultimately show ?thesis by (intro conjI)
+        next
+          case False
+          with Try.hyps 
+          have eq_s3_s2: "s3=s2" by simp
+          moreover from sxalloc not_normal_s1 abrupt_no_break
+          obtain "\<not> normal s2" 
+                 "\<forall> l. abrupt s2 \<noteq> Some (Jump (Break l))"
+            by - (rule sxalloc_cases,auto)
+	  ultimately obtain 
+	    "?NormalAssigned s3 A" and "?BreakAssigned (Norm s0) s3 A"
+	    by (cases s2) auto
+	  moreover have "?ResAssigned (Norm s0) s3"
+	  proof (cases "abrupt s1 = Some (Jump Ret)")
+	    case True
+	    with sxalloc have "s2=s1"
+	      by (elim sxalloc_cases) auto
+	    with res_s1 eq_s3_s2 show ?thesis by simp
+	  next
+	    case False
+	    with sxalloc
+	    have "abrupt s2 \<noteq> Some (Jump Ret)"
+	      by (rule sxalloc_no_jump) 
+	    with eq_s3_s2 show ?thesis
+	      by simp
+	  qed
+          ultimately show ?thesis by (intro conjI)
+        qed
+      qed
+    qed
+  next
+    case (Fin c1 c2 s0 s1 s2 s3 x1 Env T A)
+    have G: "prg Env = G" .
+    from Fin.prems obtain C1 C2 where 
+      da_C1: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1" and
+      da_C2: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2" and
+      nrm_A: "nrm A = nrm C1 \<union> nrm C2" and
+      brk_A: "brk A = ((brk C1) \<Rightarrow>\<union>\<^sub>\<forall> (nrm C2)) \<Rightarrow>\<inter> (brk C2)"
+      by (elim da_elim_cases) simp
+    from Fin.prems obtain 
+      wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and
+      wt_c2: "Env\<turnstile>c2\<Colon>\<surd>"
+      by (elim wt_elim_cases)
+    have "PROP ?Hyp (In1r c1) (Norm s0) (x1,s1)" .
+    with wt_c1 da_C1 G
+    obtain nrmAss_C1: "?NormalAssigned (x1,s1) C1" and
+           brkAss_C1: "?BreakAssigned (Norm s0) (x1,s1) C1" and
+           resAss_s1: "?ResAssigned (Norm s0) (x1,s1)"
+      by simp
+    obtain nrmAss_C2: "?NormalAssigned s2 C2" and
+           brkAss_C2: "?BreakAssigned (Norm s1) s2 C2" and
+           resAss_s2: "?ResAssigned (Norm s1) s2"
+    proof -
+      from Fin.hyps
+      have "dom (locals (store ((Norm s0)::state))) 
+             \<subseteq> dom (locals (store (x1,s1)))"
+        by - (rule dom_locals_eval_mono_elim) 
+      with da_C2 obtain C2'
+        where
+        da_C2': "Env\<turnstile> dom (locals (store (x1,s1))) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
+        nrm_C2': "nrm C2 \<subseteq> nrm C2'" and
+        brk_C2': "\<forall> l. brk C2 l \<subseteq> brk C2' l"
+        by (rule da_weakenE) simp
+      have "PROP ?Hyp (In1r c2) (Norm s1) s2" .
+      with wt_c2 da_C2' G
+      obtain nrmAss_C2': "?NormalAssigned s2 C2'" and
+             brkAss_C2': "?BreakAssigned (Norm s1) s2 C2'" and
+	     resAss_s2': "?ResAssigned (Norm s1) s2"
+        by simp
+      from nrmAss_C2' nrm_C2' have "?NormalAssigned s2 C2"
+        by blast
+      moreover
+      from brkAss_C2' brk_C2' have "?BreakAssigned (Norm s1) s2 C2"
+        by fastsimp
+      ultimately
+      show ?thesis
+        using that resAss_s2' by simp
+    qed
+    have s3: "s3 = (if \<exists>err. x1 = Some (Error err) then (x1, s1)
+                       else abupd (abrupt_if (x1 \<noteq> None) x1) s2)" .
+    have s1_s2: "dom (locals s1) \<subseteq> dom (locals (store s2))"
+    proof -  
+      have "G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2" .
+      thus ?thesis
+        by (rule dom_locals_eval_mono_elim) simp
+    qed
+
+    have "?NormalAssigned s3 A"
+    proof 
+      assume normal_s3: "normal s3"
+      show "nrm A \<subseteq> dom (locals (snd s3))"
+      proof -
+        have "nrm C1 \<subseteq> dom (locals (snd s3))"
+        proof -
+          from normal_s3 s3
+          have "normal (x1,s1)"
+            by (cases s2) (simp add: abrupt_if_def)
+          with normal_s3 nrmAss_C1 s3 s1_s2
+          show ?thesis
+            by fastsimp
+        qed
+        moreover 
+        have "nrm C2 \<subseteq> dom (locals (snd s3))"
+        proof -
+          from normal_s3 s3
+          have "normal s2"
+            by (cases s2) (simp add: abrupt_if_def)
+          with normal_s3 nrmAss_C2 s3 s1_s2
+          show ?thesis
+            by fastsimp
+        qed
+        ultimately have "nrm C1 \<union> nrm C2 \<subseteq> \<dots>"
+          by (rule Un_least)
+        with nrm_A show ?thesis
+          by simp
+      qed
+    qed
+    moreover
+    {
+      fix l assume brk_s3: "abrupt s3 = Some (Jump (Break l))"
+      have "brk A l \<subseteq> dom (locals (store s3))"
+      proof (cases "normal s2")
+        case True
+        with brk_s3 s3 
+        have s2_s3: "dom (locals (store s2)) \<subseteq> dom (locals (store s3))"
+          by simp
+        have "brk C1 l \<subseteq> dom (locals (store s3))"
+        proof -
+          from True brk_s3 s3 have "x1=Some (Jump (Break l))"
+            by (cases s2) (simp add: abrupt_if_def)
+          with brkAss_C1 s1_s2 s2_s3
+          show ?thesis
+            by simp (blast intro: subset_trans)
+        qed
+        moreover from True nrmAss_C2 s2_s3
+        have "nrm C2 \<subseteq> dom (locals (store s3))"
+          by - (rule subset_trans, simp_all)
+        ultimately 
+        have "((brk C1) \<Rightarrow>\<union>\<^sub>\<forall> (nrm C2)) l \<subseteq> \<dots>"
+          by blast
+        with brk_A show ?thesis
+          by simp blast
+      next
+        case False
+        note not_normal_s2 = this
+        have "s3=s2"
+        proof (cases "normal (x1,s1)")
+          case True with not_normal_s2 s3 show ?thesis
+            by (cases s2) (simp add: abrupt_if_def)
+        next
+          case False with not_normal_s2 s3 brk_s3 show ?thesis
+            by (cases s2) (simp add: abrupt_if_def)
+        qed
+        with brkAss_C2 brk_s3
+        have "brk C2 l \<subseteq> dom (locals (store s3))"
+          by simp
+        with brk_A show ?thesis
+          by simp blast
+      qed
+    }
+    hence "?BreakAssigned (Norm s0) s3 A"
+      by simp
+    moreover
+    {
+      assume abr_s3: "abrupt s3 = Some (Jump Ret)"
+      have "Result \<in> dom (locals (store s3))"
+      proof (cases "x1 = Some (Jump Ret)")
+	case True
+	note ret_x1 = this
+	with resAss_s1 have res_s1: "Result \<in> dom (locals s1)"
+	  by simp
+	moreover have "dom (locals (store ((Norm s1)::state))) 
+                         \<subseteq> dom (locals (store s2))"
+	  by (rule dom_locals_eval_mono_elim)
+	ultimately have "Result \<in> dom (locals (store s2))"
+	  by - (rule subsetD,auto)
+	with res_s1 s3 show ?thesis
+	  by simp
+      next
+	case False
+	with s3 abr_s3 obtain  "abrupt s2 = Some (Jump Ret)" and "s3=s2"
+	  by (cases s2) (simp add: abrupt_if_def)
+	with resAss_s2 show ?thesis
+	  by simp
+      qed
+    }
+    hence "?ResAssigned (Norm s0) s3"
+      by simp
+    ultimately show ?case by (intro conjI)
+  next 
+    case (Init C c s0 s1 s2 s3 Env T A)
+    have G: "prg Env = G" .
+    from Init.hyps
+    have eval: "prg Env\<turnstile> Norm s0 \<midarrow>Init C\<rightarrow> s3"
+      apply (simp only: G) 
+      apply (rule eval.Init, assumption)
+      apply (cases "inited C (globs s0) ")
+      apply   simp
+      apply (simp only: if_False )
+      apply (elim conjE,intro conjI,assumption+,simp)
+      done (* auto or simp alone always do too much *)
+    have "the (class G C) = c" .
+    with Init.prems 
+    have c: "class G C = Some c"
+      by (elim wt_elim_cases) auto
+    from Init.prems obtain
+       nrm_A: "nrm A = dom (locals (store ((Norm s0)::state)))"
+      by (elim da_elim_cases) simp
+    show ?case
+    proof (cases "inited C (globs s0)")
+      case True
+      with Init.hyps have "s3=Norm s0" by simp
+      thus ?thesis
+	using nrm_A by simp
+    next
+      case False
+      from Init.hyps False G
+      obtain eval_initC: 
+              "prg Env\<turnstile>Norm ((init_class_obj G C) s0) 
+                       \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1" and
+             eval_init: "prg Env\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<rightarrow> s2" and
+             s3: "s3=(set_lvars (locals (store s1))) s2"
+	by simp
+      have "?NormalAssigned s3 A"
+      proof
+	show "nrm A \<subseteq> dom (locals (store s3))"
+	proof -
+	  from nrm_A have "nrm A \<subseteq> dom (locals (init_class_obj G C s0))"
+	    by simp
+	  also from eval_initC have "\<dots> \<subseteq> dom (locals (store s1))"
+	    by (rule dom_locals_eval_mono_elim) simp
+	  also from s3 have "\<dots> \<subseteq> dom (locals (store s3))"
+	    by (cases s1) (cases s2, simp add: init_lvars_def2)
+	  finally show ?thesis .
+	qed
+      qed
+      moreover
+      from eval 
+      have "\<And> j. abrupt s3 \<noteq> Some (Jump j)"
+	by (rule eval_statement_no_jump) (auto simp add: wf c G)
+      then obtain "?BreakAssigned (Norm s0) s3 A" 
+              and "?ResAssigned (Norm s0) s3"
+	by simp
+      ultimately show ?thesis by (intro conjI)
+    qed
+  next 
+    case (NewC C a s0 s1 s2 Env T A)
+    have G: "prg Env = G" .
+    from NewC.prems
+    obtain A: "nrm A = dom (locals (store ((Norm s0)::state)))"
+              "brk A = (\<lambda> l. UNIV)"
+      by (elim da_elim_cases) simp
+    from wf NewC.prems 
+    have wt_init: "Env\<turnstile>(Init C)\<Colon>\<surd>"
+      by  (elim wt_elim_cases) (drule is_acc_classD,simp)
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s2))"
+    proof -
+      have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+        by (rule dom_locals_eval_mono_elim)
+      also
+      have "\<dots> \<subseteq> dom (locals (store s2))"
+        by (rule dom_locals_halloc_mono)
+      finally show ?thesis .
+    qed
+    with A have "?NormalAssigned s2 A"
+      by simp
+    moreover
+    {
+      fix j have "abrupt s2 \<noteq> Some (Jump j)"
+      proof -
+        have eval: "prg Env\<turnstile> Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
+	  by (simp only: G) (rule eval.NewC)
+        from NewC.prems 
+        obtain T' where  "T=Inl T'"
+          by (elim wt_elim_cases) simp
+        with NewC.prems have "Env\<turnstile>NewC C\<Colon>-T'" 
+          by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_expression_no_jump) (simp_all add: G wf)
+      qed
+    }
+    hence "?BreakAssigned (Norm s0) s2 A" and "?ResAssigned (Norm s0) s2"
+      by simp_all      
+    ultimately show ?case by (intro conjI)
+  next
+    case (NewA elT a e i s0 s1 s2 s3 Env T A) 
+    have G: "prg Env = G" .
+    from NewA.prems obtain
+      da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
+      by (elim da_elim_cases)
+    from NewA.prems obtain 
+      wt_init: "Env\<turnstile>init_comp_ty elT\<Colon>\<surd>" and 
+      wt_size: "Env\<turnstile>e\<Colon>-PrimT Integer"
+      by (elim wt_elim_cases) (auto dest:  wt_init_comp_ty')
+    have halloc:"G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr elT (the_Intg i)\<succ>a\<rightarrow>s3".
+    have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by (rule dom_locals_eval_mono_elim)
+    with da_e obtain A' where
+                da_e': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A'" 
+        and  nrm_A_A': "nrm A \<subseteq> nrm A'"                  
+        and  brk_A_A': "\<forall> l. brk A l \<subseteq> brk A' l"
+      by (rule da_weakenE) simp
+    have "PROP ?Hyp (In1l e) s1 s2" .
+    with wt_size da_e' G obtain 
+      nrmAss_A': "?NormalAssigned s2 A'" and
+      brkAss_A': "?BreakAssigned s1 s2 A'"
+      by simp
+    have s2_s3: "dom (locals (store s2)) \<subseteq> dom (locals (store s3))"
+    proof -
+      have "dom (locals (store s2)) 
+             \<subseteq> dom (locals (store (abupd (check_neg i) s2)))"
+        by (simp)
+      also have "\<dots> \<subseteq> dom (locals (store s3))"
+        by (rule dom_locals_halloc_mono)
+      finally show ?thesis .
+    qed 
+    have "?NormalAssigned s3 A"
+    proof 
+      assume normal_s3: "normal s3"
+      show "nrm A \<subseteq> dom (locals (store s3))"
+      proof -
+        from halloc normal_s3 
+        have "normal (abupd (check_neg i) s2)"
+          by cases simp_all
+        hence "normal s2"
+          by (cases s2) simp
+        with nrmAss_A' nrm_A_A' s2_s3 show ?thesis
+          by auto
+      qed
+    qed
+    moreover
+    {
+      fix j have "abrupt s3 \<noteq> Some (Jump j)"
+      proof -
+        have eval: "prg Env\<turnstile> Norm s0 \<midarrow>New elT[e]-\<succ>Addr a\<rightarrow> s3"
+	  by (simp only: G) (rule eval.NewA)
+        from NewA.prems 
+        obtain T' where  "T=Inl T'"
+          by (elim wt_elim_cases) simp
+        with NewA.prems have "Env\<turnstile>New elT[e]\<Colon>-T'" 
+          by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_expression_no_jump) (simp_all add: G wf)
+      qed
+    }
+    hence "?BreakAssigned (Norm s0) s3 A" and "?ResAssigned (Norm s0) s3"
+      by simp_all
+    ultimately show ?case by (intro conjI)
+  next
+    case (Cast cT e s0 s1 s2 v Env T A)
+    have G: "prg Env = G" .
+    from Cast.prems obtain
+      da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
+      by (elim da_elim_cases)
+    from Cast.prems obtain eT where
+      wt_e: "Env\<turnstile>e\<Colon>-eT"
+      by (elim wt_elim_cases) 
+    have "PROP ?Hyp (In1l e) (Norm s0) s1" .
+    with wt_e da_e G obtain 
+      nrmAss_A: "?NormalAssigned s1 A" and
+      brkAss_A: "?BreakAssigned (Norm s0) s1 A"
+      by simp
+    have s2: "s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits cT) ClassCast) s1" .
+    hence s1_s2: "dom (locals (store s1)) \<subseteq> dom (locals (store s2))"
+      by simp
+    have "?NormalAssigned s2 A"
+    proof 
+      assume "normal s2"
+      with s2 have "normal s1"
+        by (cases s1) simp
+      with nrmAss_A s1_s2 
+      show "nrm A \<subseteq> dom (locals (store s2))"
+        by blast
+    qed
+    moreover
+    {
+      fix j have "abrupt s2 \<noteq> Some (Jump j)"
+      proof -
+        have eval: "prg Env\<turnstile> Norm s0 \<midarrow>Cast cT e-\<succ>v\<rightarrow> s2"
+	  by (simp only: G) (rule eval.Cast)
+        from Cast.prems 
+        obtain T' where  "T=Inl T'"
+          by (elim wt_elim_cases) simp
+        with Cast.prems have "Env\<turnstile>Cast cT e\<Colon>-T'" 
+          by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_expression_no_jump) (simp_all add: G wf)
+      qed
+    }
+    hence "?BreakAssigned (Norm s0) s2 A" and "?ResAssigned (Norm s0) s2"
+      by simp_all
+    ultimately show ?case by (intro conjI)
+  next
+    case (Inst iT b e s0 s1 v Env T A)
+    have G: "prg Env = G" .
+    from Inst.prems obtain
+      da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
+      by (elim da_elim_cases)
+    from Inst.prems obtain eT where
+      wt_e: "Env\<turnstile>e\<Colon>-eT"
+      by (elim wt_elim_cases) 
+    have "PROP ?Hyp (In1l e) (Norm s0) s1" .
+    with wt_e da_e G obtain 
+      "?NormalAssigned s1 A" and
+      "?BreakAssigned (Norm s0) s1 A" and
+      "?ResAssigned (Norm s0) s1"
+      by simp
+    thus ?case by (intro conjI)
+  next
+    case (Lit s v Env T A)
+    from Lit.prems
+    have "nrm A = dom (locals (store ((Norm s)::state)))"
+      by (elim da_elim_cases) simp
+    thus ?case by simp
+  next
+    case (UnOp e s0 s1 unop v Env T A)
+     have G: "prg Env = G" .
+    from UnOp.prems obtain
+      da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
+      by (elim da_elim_cases)
+    from UnOp.prems obtain eT where
+      wt_e: "Env\<turnstile>e\<Colon>-eT"
+      by (elim wt_elim_cases) 
+    have "PROP ?Hyp (In1l e) (Norm s0) s1" .
+    with wt_e da_e G obtain 
+      "?NormalAssigned s1 A" and
+      "?BreakAssigned (Norm s0) s1 A" and
+      "?ResAssigned (Norm s0) s1"
+      by simp
+    thus ?case by (intro conjI)
+  next
+    case (BinOp binop e1 e2 s0 s1 s2 v1 v2 Env T A)
+    have G: "prg Env = G". 
+    from BinOp.hyps 
+    have 
+     eval: "prg Env\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<rightarrow> s2"
+      by (simp only: G) (rule eval.BinOp)
+    have s0_s1: "dom (locals (store ((Norm s0)::state))) 
+                  \<subseteq> dom (locals (store s1))"
+      by (rule dom_locals_eval_mono_elim)
+    also have s1_s2: "dom (locals (store s1)) \<subseteq>  dom (locals (store s2))"
+      by (rule dom_locals_eval_mono_elim)
+    finally 
+    have s0_s2: "dom (locals (store ((Norm s0)::state))) 
+                  \<subseteq> dom (locals (store s2))" . 
+    from BinOp.prems obtain e1T e2T
+      where wt_e1: "Env\<turnstile>e1\<Colon>-e1T" 
+       and  wt_e2: "Env\<turnstile>e2\<Colon>-e2T" 
+       and  wt_binop: "wt_binop (prg Env) binop e1T e2T"
+       and  T: "T=Inl (PrimT (binop_type binop))"
+      by (elim wt_elim_cases) simp
+    have "?NormalAssigned s2 A"
+    proof 
+      assume normal_s2: "normal s2"
+      have   normal_s1: "normal s1"
+        by (rule eval_no_abrupt_lemma [rule_format])
+      show "nrm A \<subseteq> dom (locals (store s2))"
+      proof (cases "binop=CondAnd")
+        case True
+        note CondAnd = this
+        from BinOp.prems obtain
+           nrm_A: "nrm A = dom (locals (store ((Norm s0)::state))) 
+                            \<union> (assigns_if True (BinOp CondAnd e1 e2) \<inter> 
+                                 assigns_if False (BinOp CondAnd e1 e2))"
+          by (elim da_elim_cases) (simp_all add: CondAnd)
+        from T BinOp.prems CondAnd
+        have "Env\<turnstile>BinOp binop e1 e2\<Colon>-PrimT Boolean"
+          by (simp)
+        with eval normal_s2
+        have ass_if: "assigns_if (the_Bool (eval_binop binop v1 v2)) 
+                                 (BinOp binop e1 e2)
+                        \<subseteq> dom (locals (store s2))"
+          by (rule assigns_if_good_approx) 
+        have "(assigns_if True (BinOp CondAnd e1 e2) \<inter> 
+                         assigns_if False (BinOp CondAnd e1 e2)) \<subseteq> \<dots>"
+        proof (cases "the_Bool (eval_binop binop v1 v2)")
+          case True
+          with ass_if CondAnd  
+          have "assigns_if True (BinOp CondAnd e1 e2) 
+                 \<subseteq> dom (locals (store s2))"
+	    by simp
+          thus ?thesis by blast
+        next
+          case False
+          with ass_if CondAnd
+          have "assigns_if False (BinOp CondAnd e1 e2) 
+                 \<subseteq> dom (locals (store s2))"
+            by (simp only: False)
+          thus ?thesis by blast
+        qed
+        with s0_s2
+        have "dom (locals (store ((Norm s0)::state))) 
+                \<union> (assigns_if True (BinOp CondAnd e1 e2) \<inter> 
+                   assigns_if False (BinOp CondAnd e1 e2)) \<subseteq> \<dots>"
+          by (rule Un_least)
+        thus ?thesis by (simp only: nrm_A)
+      next
+        case False
+        note notCondAnd = this
+        show ?thesis
+        proof (cases "binop=CondOr")
+          case True
+          note CondOr = this
+          from BinOp.prems obtain
+            nrm_A: "nrm A = dom (locals (store ((Norm s0)::state))) 
+                             \<union> (assigns_if True (BinOp CondOr e1 e2) \<inter> 
+                                 assigns_if False (BinOp CondOr e1 e2))"
+            by (elim da_elim_cases) (simp_all add: CondOr)
+          from T BinOp.prems CondOr
+          have "Env\<turnstile>BinOp binop e1 e2\<Colon>-PrimT Boolean"
+            by (simp)
+          with eval normal_s2
+          have ass_if: "assigns_if (the_Bool (eval_binop binop v1 v2)) 
+                                 (BinOp binop e1 e2)
+                          \<subseteq> dom (locals (store s2))"
+            by (rule assigns_if_good_approx) 
+          have "(assigns_if True (BinOp CondOr e1 e2) \<inter> 
+                         assigns_if False (BinOp CondOr e1 e2)) \<subseteq> \<dots>"
+          proof (cases "the_Bool (eval_binop binop v1 v2)")
+            case True
+            with ass_if CondOr 
+            have "assigns_if True (BinOp CondOr e1 e2) 
+                    \<subseteq> dom (locals (store s2))"
+              by (simp)
+            thus ?thesis by blast
+          next
+            case False
+            with ass_if CondOr
+            have "assigns_if False (BinOp CondOr e1 e2) 
+                    \<subseteq> dom (locals (store s2))"
+              by (simp)
+            thus ?thesis by blast
+          qed
+          with s0_s2
+          have "dom (locals (store ((Norm s0)::state))) 
+                 \<union> (assigns_if True (BinOp CondOr e1 e2) \<inter> 
+                    assigns_if False (BinOp CondOr e1 e2)) \<subseteq> \<dots>"
+            by (rule Un_least)
+          thus ?thesis by (simp only: nrm_A)
+        next
+          case False
+          with notCondAnd obtain notAndOr: "binop\<noteq>CondAnd" "binop\<noteq>CondOr"
+            by simp
+          from BinOp.prems obtain E1
+            where da_e1: "Env\<turnstile> dom (locals (snd (Norm s0))) \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1"  
+             and  da_e2: "Env\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A"
+            by (elim da_elim_cases) (simp_all add: notAndOr)
+          have "PROP ?Hyp (In1l e1) (Norm s0) s1" .
+          with wt_e1 da_e1 G normal_s1
+          obtain "?NormalAssigned s1 E1"  
+            by simp
+          with normal_s1 have "nrm E1 \<subseteq> dom (locals (store s1))" by rules
+          with da_e2 obtain A' 
+            where  da_e2': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A'" and
+                 nrm_A_A': "nrm A \<subseteq> nrm A'"                  
+            by (rule da_weakenE) rules
+          from notAndOr have "need_second_arg binop v1" by simp
+          with BinOp.hyps 
+          have "PROP ?Hyp (In1l e2) s1 s2" by simp
+          with wt_e2 da_e2' G
+          obtain "?NormalAssigned s2 A'"  
+            by simp
+          with nrm_A_A' normal_s2
+          show "nrm A \<subseteq> dom (locals (store s2))" 
+            by blast
+        qed
+      qed
+    qed
+    moreover
+    {
+      fix j have "abrupt s2 \<noteq> Some (Jump j)"
+      proof -
+        from BinOp.prems T 
+        have "Env\<turnstile>In1l (BinOp binop e1 e2)\<Colon>Inl (PrimT (binop_type binop))"
+          by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_expression_no_jump) (simp_all add: G wf) 
+      qed
+    }
+    hence "?BreakAssigned (Norm s0) s2 A" and "?ResAssigned (Norm s0) s2"
+      by simp_all
+    ultimately show ?case by (intro conjI) 
+  next
+    case (Super s Env T A)
+    from Super.prems
+    have "nrm A = dom (locals (store ((Norm s)::state)))"
+      by (elim da_elim_cases) simp
+    thus ?case by simp
+  next
+    case (Acc upd s0 s1 w v Env T A)
+    show ?case
+    proof (cases "\<exists> vn. v = LVar vn")
+      case True
+      then obtain vn where vn: "v=LVar vn"..
+      from Acc.prems
+      have "nrm A = dom (locals (store ((Norm s0)::state)))"
+        by (simp only: vn) (elim da_elim_cases,simp_all)
+      moreover have "G\<turnstile>Norm s0 \<midarrow>v=\<succ>(w, upd)\<rightarrow> s1" .
+      hence "s1=Norm s0"
+        by (simp only: vn) (elim eval_elim_cases,simp)
+      ultimately show ?thesis by simp
+    next
+      case False
+      have G: "prg Env = G" .
+      from False Acc.prems
+      have da_v: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>v\<rangle>\<guillemotright> A"
+        by (elim da_elim_cases) simp_all 
+      from Acc.prems obtain vT where
+        wt_v: "Env\<turnstile>v\<Colon>=vT"
+        by (elim wt_elim_cases) 
+      have "PROP ?Hyp (In2 v) (Norm s0) s1" .
+      with wt_v da_v G obtain 
+        "?NormalAssigned s1 A" and
+        "?BreakAssigned (Norm s0) s1 A" and
+	"?ResAssigned (Norm s0) s1"
+        by simp
+      thus ?thesis by (intro conjI)
+    qed
+  next
+    case (Ass e upd s0 s1 s2 v var w Env T A)
+    have G: "prg Env = G" .
+    from Ass.prems obtain varT eT where
+      wt_var: "Env\<turnstile>var\<Colon>=varT" and
+      wt_e:   "Env\<turnstile>e\<Colon>-eT"
+      by (elim wt_elim_cases) simp
+    have eval_var: "prg Env\<turnstile>Norm s0 \<midarrow>var=\<succ>(w, upd)\<rightarrow> s1" 
+      using Ass.hyps by (simp only: G)
+    have "?NormalAssigned (assign upd v s2) A"
+    proof 
+      assume normal_ass_s2: "normal (assign upd v s2)"
+      from normal_ass_s2 
+      have normal_s2: "normal s2"
+        by (cases s2) (simp add: assign_def Let_def)
+      hence normal_s1: "normal s1"
+        by - (rule eval_no_abrupt_lemma [rule_format])
+      have hyp_var: "PROP ?Hyp (In2 var) (Norm s0) s1" .
+      have hyp_e: "PROP ?Hyp (In1l e) s1 s2" .
+      show "nrm A \<subseteq> dom (locals (store (assign upd v s2)))"
+      proof (cases "\<exists> vn. var = LVar vn")
+	case True
+	then obtain vn where vn: "var=LVar vn"..
+	from Ass.prems obtain E where
+	  da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E" and
+	  nrm_A: "nrm A = nrm E \<union> {vn}"
+	  by (elim da_elim_cases) (insert vn,auto)
+	obtain E' where
+	  da_e': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'" and
+	  E_E': "nrm E \<subseteq> nrm E'"
+	proof -
+	  have "dom (locals (store ((Norm s0)::state))) 
+                  \<subseteq> dom (locals (store s1))"
+	    by (rule dom_locals_eval_mono_elim)
+	  with da_e show ?thesis
+	    by (rule da_weakenE)
+	qed
+	from G eval_var vn 
+	have eval_lvar: "G\<turnstile>Norm s0 \<midarrow>LVar vn=\<succ>(w, upd)\<rightarrow> s1" 
+	  by simp
+	then have upd: "upd = snd (lvar vn (store s1))"
+	  by cases (cases "lvar vn (store s1)",simp)
+	have "nrm E \<subseteq> dom (locals (store (assign upd v s2)))"
+	proof -
+	  from hyp_e wt_e da_e' G normal_s2
+	  have "nrm E' \<subseteq> dom (locals (store s2))"
+	    by simp
+	  also
+	  from upd
+	  have "dom (locals (store s2))  \<subseteq> dom (locals (store (upd v s2)))"
+	    by (simp add: lvar_def) blast
+	  hence "dom (locals (store s2)) 
+	             \<subseteq> dom (locals (store (assign upd v s2)))"
+	    by (rule dom_locals_assign_mono)
+	  finally
+	  show ?thesis using E_E' 
+	    by blast
+	qed
+	moreover
+	from upd normal_s2
+	have "{vn} \<subseteq> dom (locals (store (assign upd v s2)))"
+	  by (auto simp add: assign_def Let_def lvar_def upd split: split_split)
+	ultimately
+	show "nrm A \<subseteq> \<dots>"
+	  by (rule Un_least [elim_format]) (simp add: nrm_A)
+      next
+	case False
+	from Ass.prems obtain V where
+	  da_var: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>var\<rangle>\<guillemotright> V" and
+	  da_e:   "Env\<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
+	  by (elim da_elim_cases) (insert False,simp+)
+        from hyp_var wt_var da_var G normal_s1
+        have "nrm V \<subseteq> dom (locals (store s1))"
+          by simp
+        with da_e obtain A' 
+          where  da_e': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A'" and
+                 nrm_A_A': "nrm A \<subseteq> nrm A'"                  
+          by (rule da_weakenE) rules
+        from hyp_e wt_e da_e' G normal_s2
+        obtain "nrm A' \<subseteq> dom (locals (store s2))"   
+          by simp
+        with nrm_A_A' have "nrm A \<subseteq> \<dots>" 
+          by blast
+        also have "\<dots> \<subseteq> dom (locals (store (assign upd v s2)))"
+        proof -
+          from eval_var normal_s1
+          have "dom (locals (store s2)) \<subseteq> dom (locals (store (upd v s2)))"
+            by (cases rule: dom_locals_eval_mono_elim)
+               (cases s2, simp)
+          thus ?thesis
+            by (rule dom_locals_assign_mono)
+        qed
+        finally show ?thesis .
+      qed
+    qed
+    moreover 
+    {
+      fix j have "abrupt (assign upd v s2) \<noteq> Some (Jump j)"
+      proof -
+        have eval: "prg Env\<turnstile>Norm s0 \<midarrow>var:=e-\<succ>v\<rightarrow> (assign upd v s2)"
+          by (simp only: G) (rule eval.Ass)
+        from Ass.prems 
+        obtain T' where  "T=Inl T'"
+          by (elim wt_elim_cases) simp
+        with Ass.prems have "Env\<turnstile>var:=e\<Colon>-T'" by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_expression_no_jump) (simp_all add: G wf)
+      qed
+    }
+    hence "?BreakAssigned (Norm s0) (assign upd v s2) A"
+      and "?ResAssigned (Norm s0) (assign upd v s2)"       
+      by simp_all
+    ultimately show ?case by (intro conjI)
+  next
+-- {* 
+\par
+*} (* dummy text command to break paragraph for latex;
+              large paragraphs exhaust memory of debian pdflatex *)
+    case (Cond b e0 e1 e2 s0 s1 s2 v Env T A)
+    have G: "prg Env = G" .
+    have "?NormalAssigned s2 A"
+    proof 
+      assume normal_s2: "normal s2"
+      show "nrm A \<subseteq> dom (locals (store s2))"
+      proof (cases "Env\<turnstile>(e0 ? e1 : e2)\<Colon>-(PrimT Boolean)")
+        case True
+        with Cond.prems 
+        have nrm_A: "nrm A = dom (locals (store ((Norm s0)::state))) 
+                             \<union> (assigns_if True  (e0 ? e1 : e2) \<inter> 
+                                 assigns_if False (e0 ? e1 : e2))"
+          by (elim da_elim_cases) simp_all
+        have eval: "prg Env\<turnstile>Norm s0 \<midarrow>(e0 ? e1 : e2)-\<succ>v\<rightarrow> s2"
+          by (simp only: G) (rule eval.Cond)
+        from eval 
+        have "dom (locals (store ((Norm s0)::state)))\<subseteq>dom (locals (store s2))"
+          by (rule dom_locals_eval_mono_elim)
+        moreover
+        from eval normal_s2 True
+        have ass_if: "assigns_if (the_Bool v) (e0 ? e1 : e2)
+                        \<subseteq> dom (locals (store s2))"
+          by (rule assigns_if_good_approx)
+        have "assigns_if True  (e0 ? e1:e2) \<inter> assigns_if False (e0 ? e1:e2)
+               \<subseteq> dom (locals (store s2))"
+        proof (cases "the_Bool v")
+          case True 
+          from ass_if 
+          have "assigns_if True  (e0 ? e1:e2) \<subseteq> dom (locals (store s2))"
+            by (simp only: True)
+          thus ?thesis by blast
+        next
+          case False 
+          from ass_if 
+          have "assigns_if False  (e0 ? e1:e2) \<subseteq> dom (locals (store s2))"
+            by (simp only: False)
+          thus ?thesis by blast
+        qed
+        ultimately show "nrm A \<subseteq> dom (locals (store s2))"
+          by (simp only: nrm_A) (rule Un_least)
+      next
+        case False
+        with Cond.prems obtain E1 E2 where
+         da_e1: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                         \<union> assigns_if True e0) \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1" and
+         da_e2: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                        \<union> assigns_if False e0) \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2" and
+         nrm_A: "nrm A = nrm E1 \<inter> nrm E2"
+          by (elim da_elim_cases) simp_all
+        from Cond.prems obtain e1T e2T where
+          wt_e0: "Env\<turnstile>e0\<Colon>- PrimT Boolean" and
+          wt_e1: "Env\<turnstile>e1\<Colon>-e1T" and
+          wt_e2: "Env\<turnstile>e2\<Colon>-e2T" 
+          by (elim wt_elim_cases)
+        have s0_s1: "dom (locals (store ((Norm s0)::state))) 
+                       \<subseteq> dom (locals (store s1))"
+          by (rule dom_locals_eval_mono_elim)
+        have eval_e0: "prg Env\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1" by (simp only: G)
+        have normal_s1: "normal s1"
+          by (rule eval_no_abrupt_lemma [rule_format])
+        show ?thesis
+        proof (cases "the_Bool b")
+          case True
+          from True Cond.hyps have "PROP ?Hyp (In1l e1) s1 s2" by simp
+          moreover
+          from eval_e0 normal_s1 wt_e0
+          have "assigns_if True e0 \<subseteq> dom (locals (store s1))"
+            by (rule assigns_if_good_approx [elim_format]) (simp only: True)
+          with s0_s1 
+          have "dom (locals (store ((Norm s0)::state))) 
+                 \<union> assigns_if True e0 \<subseteq> \<dots>"
+            by (rule Un_least)
+          with da_e1 obtain E1' where
+                  da_e1': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'" and
+              nrm_E1_E1': "nrm E1 \<subseteq> nrm E1'"
+            by (rule da_weakenE) rules
+          ultimately have "nrm E1' \<subseteq> dom (locals (store s2))"
+            using wt_e1 G normal_s2 by simp 
+          with nrm_E1_E1' show ?thesis
+            by (simp only: nrm_A) blast
+        next
+          case False
+          from False Cond.hyps have "PROP ?Hyp (In1l e2) s1 s2" by simp
+          moreover
+          from eval_e0 normal_s1 wt_e0
+          have "assigns_if False e0 \<subseteq> dom (locals (store s1))"
+            by (rule assigns_if_good_approx [elim_format]) (simp only: False)
+          with s0_s1 
+          have "dom (locals (store ((Norm s0)::state))) 
+                 \<union> assigns_if False e0 \<subseteq> \<dots>"
+            by (rule Un_least)
+          with da_e2 obtain E2' where
+                  da_e2': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2'" and
+              nrm_E2_E2': "nrm E2 \<subseteq> nrm E2'"
+            by (rule da_weakenE) rules
+          ultimately have "nrm E2' \<subseteq> dom (locals (store s2))"
+            using wt_e2 G normal_s2 by simp 
+          with nrm_E2_E2' show ?thesis
+            by (simp only: nrm_A) blast
+        qed
+      qed
+    qed
+    moreover
+    {
+      fix j have "abrupt s2 \<noteq> Some (Jump j)"
+      proof -
+        have eval: "prg Env\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<rightarrow> s2"
+          by (simp only: G) (rule eval.Cond)
+        from Cond.prems 
+        obtain T' where  "T=Inl T'"
+          by (elim wt_elim_cases) simp
+        with Cond.prems have "Env\<turnstile>e0 ? e1 : e2\<Colon>-T'" by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_expression_no_jump) (simp_all add: G wf)
+      qed
+    } 
+    hence "?BreakAssigned (Norm s0) s2 A" and "?ResAssigned (Norm s0) s2"
+      by simp_all
+    ultimately show ?case by (intro conjI)
+  next
+    case (Call D a accC args e mn mode pTs s0 s1 s2 s3 s3' s4 statT v vs 
+           Env T A)
+    have G: "prg Env = G" .
+    have "?NormalAssigned (restore_lvars s2 s4) A"
+    proof 
+      assume normal_restore_lvars: "normal (restore_lvars s2 s4)"
+      show "nrm A \<subseteq> dom (locals (store (restore_lvars s2 s4)))"
+      proof -
+        from Call.prems obtain E where
+             da_e: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>\<langle>e\<rangle>\<guillemotright> E" and
+          da_args: "Env\<turnstile> nrm E \<guillemotright>\<langle>args\<rangle>\<guillemotright> A" 
+          by (elim da_elim_cases)
+        from Call.prems obtain eT argsT where
+             wt_e: "Env\<turnstile>e\<Colon>-eT" and
+          wt_args: "Env\<turnstile>args\<Colon>\<doteq>argsT"
+          by (elim wt_elim_cases)
+        have s3: "s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a vs s2" .
+        have s3': "s3' = check_method_access G accC statT mode 
+                                           \<lparr>name=mn,parTs=pTs\<rparr> a s3" .
+        have normal_s2: "normal s2"
+        proof -
+          from normal_restore_lvars have "normal s4"
+            by simp
+          then have "normal s3'"
+            by - (rule eval_no_abrupt_lemma [rule_format])
+          with s3' have "normal s3"
+            by (cases s3) (simp add: check_method_access_def Let_def)
+          with s3 show "normal s2"
+            by (cases s2) (simp add: init_lvars_def Let_def)
+        qed
+        then have normal_s1: "normal s1"
+          by  - (rule eval_no_abrupt_lemma [rule_format])
+        have "PROP ?Hyp (In1l e) (Norm s0) s1" .
+        with da_e wt_e G normal_s1
+        have "nrm E \<subseteq> dom (locals (store s1))"
+          by simp
+        with da_args obtain A' where
+          da_args': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<guillemotright> A'" and
+          nrm_A_A': "nrm A \<subseteq> nrm A'"
+          by (rule da_weakenE) rules
+        have "PROP ?Hyp (In3 args) s1 s2" .
+        with da_args' wt_args G normal_s2
+        have "nrm A' \<subseteq> dom (locals (store s2))"
+          by simp
+        with nrm_A_A' have "nrm A \<subseteq> dom (locals (store s2))"
+          by blast
+        also have "\<dots> \<subseteq> dom (locals (store (restore_lvars s2 s4)))"
+          by (cases s4) simp
+        finally show ?thesis .
+      qed
+    qed
+    moreover
+    {
+      fix j have "abrupt (restore_lvars s2 s4) \<noteq> Some (Jump j)"
+      proof -
+        have eval: "prg Env\<turnstile>Norm s0 \<midarrow>({accC,statT,mode}e\<cdot>mn( {pTs}args))-\<succ>v
+                            \<rightarrow> (restore_lvars s2 s4)"
+          by (simp only: G) (rule eval.Call)
+        from Call.prems 
+        obtain T' where  "T=Inl T'"
+          by (elim wt_elim_cases) simp
+        with Call.prems have "Env\<turnstile>({accC,statT,mode}e\<cdot>mn( {pTs}args))\<Colon>-T'" 
+          by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_expression_no_jump) (simp_all add: G wf)
+      qed
+    } 
+    hence "?BreakAssigned (Norm s0) (restore_lvars s2 s4) A"
+      and "?ResAssigned (Norm s0) (restore_lvars s2 s4)"
+      by simp_all
+    ultimately show ?case by (intro conjI)
+  next
+    case (Methd D s0 s1 sig v Env T A)
+    have G: "prg Env = G". 
+    from Methd.prems obtain m where
+       m:      "methd (prg Env) D sig = Some m" and
+       da_body: "Env\<turnstile>(dom (locals (store ((Norm s0)::state)))) 
+                  \<guillemotright>\<langle>Body (declclass m) (stmt (mbody (mthd m)))\<rangle>\<guillemotright> A"
+      by - (erule da_elim_cases)
+    from Methd.prems m obtain
+      isCls: "is_class (prg Env) D" and
+      wt_body: "Env \<turnstile>In1l (Body (declclass m) (stmt (mbody (mthd m))))\<Colon>T"
+      by - (erule wt_elim_cases,simp)
+    have "PROP ?Hyp (In1l (body G D sig)) (Norm s0) s1" .
+    moreover
+    from wt_body have "Env\<turnstile>In1l (body G D sig)\<Colon>T"
+      using isCls m G by (simp add: body_def2)
+    moreover 
+    from da_body have "Env\<turnstile>(dom (locals (store ((Norm s0)::state)))) 
+                          \<guillemotright>\<langle>body G D sig\<rangle>\<guillemotright> A"
+      using isCls m G by (simp add: body_def2)
+    ultimately show ?case
+      using G by simp
+  next
+    case (Body D c s0 s1 s2 s3 Env T A)
+    have G: "prg Env = G" .
+    from Body.prems 
+    have nrm_A: "nrm A = dom (locals (store ((Norm s0)::state)))"
+      by (elim da_elim_cases) simp
+    have eval: "prg Env\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
+                        \<rightarrow>abupd (absorb Ret) s3"
+      by (simp only: G) (rule eval.Body)
+    hence "nrm A \<subseteq> dom (locals (store (abupd (absorb Ret) s3)))"
+      by  (simp only: nrm_A) (rule dom_locals_eval_mono_elim)
+    hence "?NormalAssigned (abupd (absorb Ret) s3) A"
+      by simp
+    moreover
+    from eval have "\<And> j. abrupt (abupd (absorb Ret) s3) \<noteq> Some (Jump j)"
+      by (rule Body_no_jump) simp
+    hence "?BreakAssigned (Norm s0) (abupd (absorb Ret) s3) A" and
+          "?ResAssigned (Norm s0) (abupd (absorb Ret) s3)"
+      by simp_all
+    ultimately show ?case by (intro conjI)
+  next
+    case (LVar s vn Env T A)
+    from LVar.prems
+    have "nrm A = dom (locals (store ((Norm s)::state)))"
+      by (elim da_elim_cases) simp
+    thus ?case by simp
+  next
+    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v Env T A)
+    have G: "prg Env = G" .
+    have "?NormalAssigned s3 A"
+    proof 
+      assume normal_s3: "normal s3"
+      show "nrm A \<subseteq> dom (locals (store s3))"
+      proof -
+        have fvar: "(v, s2') = fvar statDeclC stat fn a s2" and
+               s3: "s3 = check_field_access G accC statDeclC fn stat a s2'" .
+        from FVar.prems
+        have da_e: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
+          by (elim da_elim_cases)
+        from FVar.prems obtain eT where
+          wt_e: "Env\<turnstile>e\<Colon>-eT"
+          by (elim wt_elim_cases)
+        have "(dom (locals (store ((Norm s0)::state)))) 
+               \<subseteq> dom (locals (store s1))"
+          by (rule dom_locals_eval_mono_elim)
+        with da_e obtain A' where
+          da_e': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A'" and
+          nrm_A_A': "nrm A \<subseteq> nrm A'"
+          by (rule da_weakenE) rules
+        have normal_s2: "normal s2"
+        proof -
+          from normal_s3 s3 
+          have "normal s2'"
+            by (cases s2') (simp add: check_field_access_def Let_def)
+          with fvar 
+          show "normal s2"
+            by (cases s2) (simp add: fvar_def2)
+        qed
+        have "PROP ?Hyp (In1l e) s1 s2" . 
+        with da_e' wt_e G normal_s2
+        have "nrm A' \<subseteq> dom (locals (store s2))"
+          by simp
+        with nrm_A_A' have "nrm A \<subseteq> dom (locals (store s2))"
+          by blast
+        also have "\<dots> \<subseteq> dom (locals (store s3))"
+        proof -
+          from fvar have "s2' = snd (fvar statDeclC stat fn a s2)" 
+            by (cases "fvar statDeclC stat fn a s2") simp
+          hence "dom (locals (store s2)) \<subseteq>  dom (locals (store s2'))"
+            by (simp) (rule dom_locals_fvar_mono)
+          also from s3 have "\<dots> \<subseteq> dom (locals (store s3))"
+            by (cases s2') (simp add: check_field_access_def Let_def)
+          finally show ?thesis .
+        qed
+        finally show ?thesis .
+      qed
+    qed
+    moreover
+    {
+      fix j have "abrupt s3 \<noteq> Some (Jump j)"
+      proof -
+        obtain w upd where v: "(w,upd)=v"
+          by (cases v) auto
+        have eval: "prg Env\<turnstile>Norm s0\<midarrow>({accC,statDeclC,stat}e..fn)=\<succ>(w,upd)\<rightarrow>s3"
+          by (simp only: G v) (rule eval.FVar)
+        from FVar.prems 
+        obtain T' where  "T=Inl T'"
+          by (elim wt_elim_cases) simp
+        with FVar.prems have "Env\<turnstile>({accC,statDeclC,stat}e..fn)\<Colon>=T'" 
+          by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_var_no_jump [THEN conjunct1]) (simp_all add: G wf)
+      qed
+    } 
+    hence "?BreakAssigned (Norm s0) s3 A" and "?ResAssigned (Norm s0) s3"
+      by simp_all
+    ultimately show ?case by (intro conjI)
+  next
+    case (AVar a e1 e2 i s0 s1 s2 s2' v Env T A)
+    have G: "prg Env = G" .
+    have "?NormalAssigned s2' A"
+    proof 
+      assume normal_s2': "normal s2'"
+      show "nrm A \<subseteq> dom (locals (store s2'))"
+      proof -
+        have avar: "(v, s2') = avar G i a s2" .
+        from AVar.prems obtain E1 where
+          da_e1: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1" and
+          da_e2: "Env\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A" 
+          by (elim da_elim_cases)
+        from AVar.prems obtain e1T e2T where
+             wt_e1: "Env\<turnstile>e1\<Colon>-e1T" and
+             wt_e2: "Env\<turnstile>e2\<Colon>-e2T"
+          by (elim wt_elim_cases)
+        from avar normal_s2' 
+        have normal_s2: "normal s2"
+          by (cases s2) (simp add: avar_def2)
+        hence "normal s1"
+          by - (rule eval_no_abrupt_lemma [rule_format])
+        moreover have "PROP ?Hyp (In1l e1) (Norm s0) s1" .
+        ultimately have "nrm E1 \<subseteq> dom (locals (store s1))" 
+          using da_e1 wt_e1 G by simp
+        with da_e2 obtain A' where
+          da_e2': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A'" and
+          nrm_A_A': "nrm A \<subseteq> nrm A'"
+          by (rule da_weakenE) rules
+        have "PROP ?Hyp (In1l e2) s1 s2" .
+        with da_e2' wt_e2 G normal_s2
+        have "nrm A' \<subseteq> dom (locals (store s2))"
+          by simp
+        with nrm_A_A' have "nrm A \<subseteq> dom (locals (store s2))"
+          by blast
+        also have "\<dots> \<subseteq> dom (locals (store s2'))"
+        proof -
+           from avar have "s2' = snd (avar G i a s2)" 
+            by (cases "(avar G i a s2)") simp
+          thus "dom (locals (store s2)) \<subseteq>  dom (locals (store s2'))"
+            by (simp) (rule dom_locals_avar_mono)
+        qed
+        finally show ?thesis .
+      qed
+    qed
+    moreover
+    {
+      fix j have "abrupt s2' \<noteq> Some (Jump j)"
+      proof -
+        obtain w upd where v: "(w,upd)=v"
+          by (cases v) auto
+        have eval: "prg Env\<turnstile>Norm s0\<midarrow>(e1.[e2])=\<succ>(w,upd)\<rightarrow>s2'"
+          by  (simp only: G v) (rule eval.AVar)
+        from AVar.prems 
+        obtain T' where  "T=Inl T'"
+          by (elim wt_elim_cases) simp
+        with AVar.prems have "Env\<turnstile>(e1.[e2])\<Colon>=T'" 
+          by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_var_no_jump [THEN conjunct1]) (simp_all add: G wf)
+      qed
+    } 
+    hence "?BreakAssigned (Norm s0) s2' A" and "?ResAssigned (Norm s0) s2'"
+      by simp_all
+    ultimately show ?case by (intro conjI)
+  next
+    case (Nil s0 Env T A)
+    from Nil.prems
+    have "nrm A = dom (locals (store ((Norm s0)::state)))"
+      by (elim da_elim_cases) simp
+    thus ?case by simp
+  next 
+    case (Cons e es s0 s1 s2 v vs Env T A)
+    have G: "prg Env = G" .
+    have "?NormalAssigned s2 A"
+    proof 
+      assume normal_s2: "normal s2"
+      show "nrm A \<subseteq> dom (locals (store s2))"
+      proof -
+        from Cons.prems obtain E where
+          da_e: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>\<langle>e\<rangle>\<guillemotright> E" and
+          da_es: "Env\<turnstile> nrm E \<guillemotright>\<langle>es\<rangle>\<guillemotright> A" 
+          by (elim da_elim_cases)
+        from Cons.prems obtain eT esT where
+             wt_e: "Env\<turnstile>e\<Colon>-eT" and
+             wt_es: "Env\<turnstile>es\<Colon>\<doteq>esT"
+          by (elim wt_elim_cases)
+        have "normal s1"
+          by - (rule eval_no_abrupt_lemma [rule_format])
+        moreover have "PROP ?Hyp (In1l e) (Norm s0) s1" .
+        ultimately have "nrm E \<subseteq> dom (locals (store s1))" 
+          using da_e wt_e G by simp
+        with da_es obtain A' where
+          da_es': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>es\<rangle>\<guillemotright> A'" and
+          nrm_A_A': "nrm A \<subseteq> nrm A'"
+          by (rule da_weakenE) rules
+        have "PROP ?Hyp (In3 es) s1 s2" .
+        with da_es' wt_es G normal_s2
+        have "nrm A' \<subseteq> dom (locals (store s2))"
+          by simp
+        with nrm_A_A' show "nrm A \<subseteq> dom (locals (store s2))"
+          by blast
+      qed
+    qed
+    moreover
+    {
+      fix j have "abrupt s2 \<noteq> Some (Jump j)"
+      proof -
+        have eval: "prg Env\<turnstile>Norm s0\<midarrow>(e # es)\<doteq>\<succ>v#vs\<rightarrow>s2"
+          by (simp only: G) (rule eval.Cons)
+        from Cons.prems 
+        obtain T' where  "T=Inr T'"
+          by (elim wt_elim_cases) simp
+        with Cons.prems have "Env\<turnstile>(e # es)\<Colon>\<doteq>T'" 
+          by simp
+        from eval _ this
+        show ?thesis
+          by (rule eval_expression_list_no_jump) (simp_all add: G wf)
+      qed
+    } 
+    hence "?BreakAssigned (Norm s0) s2 A" and "?ResAssigned (Norm s0) s2"
+      by simp_all
+    ultimately show ?case by (intro conjI)
+  qed
+qed
+
+lemma da_good_approxE [consumes 4]:
+  "\<lbrakk>prg Env\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v, s1); Env\<turnstile>t\<Colon>T; Env\<turnstile> dom (locals (store s0)) \<guillemotright>t\<guillemotright> A;
+   wf_prog (prg Env);
+   \<lbrakk>normal s1 \<Longrightarrow> nrm A \<subseteq> dom (locals (store s1));
+     \<And> l. \<lbrakk>abrupt s1 = Some (Jump (Break l)); normal s0\<rbrakk>
+           \<Longrightarrow> brk A l \<subseteq> dom (locals (store s1));
+     \<lbrakk>abrupt s1 = Some (Jump Ret);normal s0\<rbrakk>\<Longrightarrow>Result \<in> dom (locals (store s1))
+   \<rbrakk> \<Longrightarrow> P
+  \<rbrakk> \<Longrightarrow> P"
+by (drule (3) da_good_approx) simp
+
+
+(* Same as above but with explicit environment; 
+   It would be nicer to find a "normalized" way to right down those things
+   in Bali.
+*)
+lemma da_good_approxE' [consumes 4]:
+  assumes eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v, s1)"
+     and    wt: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T"
+     and    da: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>t\<guillemotright> A"
+     and    wf: "wf_prog G"
+     and  elim: "\<lbrakk>normal s1 \<Longrightarrow> nrm A \<subseteq> dom (locals (store s1));
+                 \<And> l. \<lbrakk>abrupt s1 = Some (Jump (Break l)); normal s0\<rbrakk>
+                       \<Longrightarrow> brk A l \<subseteq> dom (locals (store s1));
+                  \<lbrakk>abrupt s1 = Some (Jump Ret);normal s0\<rbrakk>
+                  \<Longrightarrow>Result \<in> dom (locals (store s1))
+                 \<rbrakk> \<Longrightarrow> P"
+  shows "P"
+proof -
+  from eval have "prg \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v, s1)" by simp
+  moreover note wt da
+  moreover from wf have "wf_prog (prg \<lparr>prg=G,cls=C,lcl=L\<rparr>)" by simp
+  ultimately show ?thesis
+    using elim by (rule da_good_approxE) rules+
+qed
+
+ML {*
+Addsimprocs [wt_expr_proc,wt_var_proc,wt_exprs_proc,wt_stmt_proc]
+*}
+end
\ No newline at end of file
--- a/src/HOL/Bali/Eval.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Eval.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -97,17 +97,50 @@
 \end{itemize}
 *}
 
+
 types vvar  =         "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
       vals  =        "(val, vvar, val list) sum3"
 translations
      "vvar" <= (type) "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
-     "vals" <= (type)"(val, vvar, val list) sum3"
+     "vals" <= (type)"(val, vvar, val list) sum3" 
+
+text {* To avoid redundancy and to reduce the number of rules, there is only 
+ one evaluation rule for each syntactic term. This is also true for variables
+ (e.g. see the rules below for @{text LVar}, @{text FVar} and @{text AVar}). 
+ So evaluation of a variable must capture both possible further uses: 
+ read (rule @{text Acc}) or write (rule @{text Ass}) to the variable. 
+ Therefor a variable evaluates to a special value @{term vvar}, which is 
+ a pair, consisting of the current value (for later read access) and an update 
+ function (for later write access). Because
+ during assignment to an array variable an exception may occur if the types
+ don't match, the update function is very generic: it transforms the 
+ full state. This generic update function causes some technical trouble during
+ some proofs (e.g. type safety, correctness of definite assignment). There we
+ need to prove some additional invariant on this update function to prove the
+ assignment correct, since the update function could potentially alter the whole
+ state in an arbitrary manner. This invariant must be carried around through
+ the whole induction.
+ So for future approaches it may be better not to take
+ such a generic update function, but only to store the address and the kind
+ of variable (array (+ element type), local variable or field) for later 
+ assignment. 
+*}
 
 syntax (xsymbols)
   dummy_res :: "vals" ("\<diamondsuit>")
 translations
   "\<diamondsuit>" == "In1 Unit"
 
+syntax 
+  val_inj_vals:: "expr \<Rightarrow> term" ("\<lfloor>_\<rfloor>\<^sub>e" 1000)
+  var_inj_vals::  "var \<Rightarrow> term"  ("\<lfloor>_\<rfloor>\<^sub>v" 1000)
+  lst_inj_vals:: "expr list \<Rightarrow> term" ("\<lfloor>_\<rfloor>\<^sub>l" 1000)
+
+translations 
+  "\<lfloor>e\<rfloor>\<^sub>e" \<rightharpoonup> "In1 e"
+  "\<lfloor>v\<rfloor>\<^sub>v" \<rightharpoonup> "In2 v"
+  "\<lfloor>es\<rfloor>\<^sub>l" \<rightharpoonup> "In3 es"
+
 constdefs
   arbitrary3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals"
  "arbitrary3 \<equiv> sum3_case (In1 \<circ> sum_case (\<lambda>x. arbitrary) (\<lambda>x. Unit))
@@ -179,6 +212,16 @@
 apply (simp (no_asm))
 done
 
+lemma catch_Jump [simp]: "\<not>G,(Some (Jump j),s)\<turnstile>catch tn"
+apply (unfold catch_def)
+apply (simp (no_asm))
+done
+
+lemma catch_Error [simp]: "\<not>G,(Some (Error e),s)\<turnstile>catch tn"
+apply (unfold catch_def)
+apply (simp (no_asm))
+done
+
 constdefs
   new_xcpt_var :: "vname \<Rightarrow> state \<Rightarrow> state"
  "new_xcpt_var vn \<equiv> 
@@ -229,6 +272,9 @@
 "assign f v (Some x,s) = (Some x,s)" 
 by (simp add: assign_def Let_def split_beta)
 
+lemma assign_Some1 [simp]:  "\<not> normal s \<Longrightarrow> assign f v s = s" 
+by (auto simp add: assign_def Let_def split_beta)
+
 lemma assign_supd [simp]: 
 "assign (\<lambda>v. supd (f v)) v (x,s)  
   = (x, if x = None then f v s else s)"
@@ -242,6 +288,7 @@
 apply auto
 done
 
+
 (*
 lemma assign_raise_if [simp]: 
   "assign (\<lambda>v s. \<lparr>abrupt=(raise_if (b (store s) v) xcpt) (abrupt s),
@@ -266,14 +313,6 @@
 
 constdefs
 
-(*
-  target  :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname"
- "target m s a' t 
-    \<equiv> if m = IntVir
-	 then obj_class (lookup_obj s a') 
-         else the_Class (RefT t)"
-*)
-
  invocation_class  :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname"
  "invocation_class m s a' statT 
     \<equiv> (case m of
@@ -296,11 +335,6 @@
 lemma dynclass_SuperM [simp]: 
  "invocation_class SuperM s a' statT = the_Class (RefT statT)"
 by (simp add: invocation_class_def)
-(*
-lemma invocation_class_notIntVir [simp]: 
- "m \<noteq> IntVir \<Longrightarrow> invocation_class m s a' statT = the_Class (RefT statT)"
-by (simp add: invocation_class_def)
-*)
 
 lemma invocation_class_Static [simp]: 
   "invocation_class Static s a' statT = (if (\<exists> statC. statT = ClassT statC) 
@@ -318,26 +352,24 @@
               (case k of
                  EName e 
                    \<Rightarrow> (case e of 
-                         VNam v \<Rightarrow> (init_vals (table_of (lcls (mbody m)))
-                                                     ((pars m)[\<mapsto>]pvs)) v
-                       | Res    \<Rightarrow> Some (default_val (resTy m)))
+                         VNam v \<Rightarrow> (empty ((pars m)[\<mapsto>]pvs)) v
+                       | Res    \<Rightarrow> None)
                | This 
                    \<Rightarrow> (if mode=Static then None else Some a'))
       in set_lvars l (if mode = Static then x else np a' x,s)"
 
 
 
-lemma init_lvars_def2: "init_lvars G C sig mode a' pvs (x,s) =  
+lemma init_lvars_def2: --{* better suited for simplification *} 
+"init_lvars G C sig mode a' pvs (x,s) =  
   set_lvars 
     (\<lambda> k. 
        (case k of
           EName e 
             \<Rightarrow> (case e of 
                   VNam v 
-                  \<Rightarrow> (init_vals 
-                       (table_of (lcls (mbody (mthd (the (methd G C sig))))))
-                                 ((pars (mthd (the (methd G C sig))))[\<mapsto>]pvs)) v
-               | Res \<Rightarrow> Some (default_val (resTy (mthd (the (methd G C sig))))))
+                  \<Rightarrow> (empty ((pars (mthd (the (methd G C sig))))[\<mapsto>]pvs)) v
+               | Res \<Rightarrow> None)
         | This 
             \<Rightarrow> (if mode=Static then None else Some a')))
     (if mode = Static then x else np a' x,s)"
@@ -350,7 +382,7 @@
  "body G C sig \<equiv> let m = the (methd G C sig) 
                  in Body (declclass m) (stmt (mbody (mthd m)))"
 
-lemma body_def2: 
+lemma body_def2: --{* better suited for simplification *} 
 "body G C sig = Body  (declclass (the (methd G C sig))) 
                       (stmt (mbody (mthd (the (methd G C sig)))))"
 apply (unfold body_def Let_def)
@@ -371,14 +403,7 @@
 	          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'); 
@@ -391,7 +416,8 @@
       in ((the (cs n),f)
          ,abupd (raise_if (\<not>i in_bounds k) IndOutBound \<circ> np a') s)"
 
-lemma fvar_def2: "fvar C stat fn a' s =  
+lemma fvar_def2: --{* better suited for simplification *} 
+"fvar C stat fn a' s =  
   ((the 
      (values 
       (the (globs (store s) (if stat then Stat C else Heap (the_Addr a')))) 
@@ -405,7 +431,8 @@
 apply (simp (no_asm) add: Let_def split_beta)
 done
 
-lemma avar_def2: "avar G i' a' s =  
+lemma avar_def2: --{* better suited for simplification *} 
+"avar G i' a' s =  
   ((the ((snd(snd(the_Arr (globs (store s) (Heap (the_Addr a')))))) 
            (Inr (the_Intg i')))
    ,(\<lambda>v (x,s').  (raise_if (\<not>G,s'\<turnstile>v fits (fst(the_Arr (globs (store s)
@@ -448,59 +475,6 @@
        
 section "evaluation judgments"
 
-consts eval_unop :: "unop \<Rightarrow> val \<Rightarrow> val"
-primrec
-"eval_unop UPlus   v = Intg (the_Intg v)"
-"eval_unop UMinus  v = Intg (- (the_Intg v))"
-"eval_unop UBitNot v = Intg 42"                -- "FIXME: Not yet implemented"
-"eval_unop UNot    v = Bool (\<not> the_Bool v)"
-
-consts eval_binop :: "binop \<Rightarrow> val \<Rightarrow> val \<Rightarrow> val"
-primrec
-"eval_binop Mul     v1 v2 = Intg ((the_Intg v1) * (the_Intg v2))" 
-"eval_binop Div     v1 v2 = Intg ((the_Intg v1) div (the_Intg v2))"
-"eval_binop Mod     v1 v2 = Intg ((the_Intg v1) mod (the_Intg v2))"
-"eval_binop Plus    v1 v2 = Intg ((the_Intg v1) + (the_Intg v2))"
-"eval_binop Minus   v1 v2 = Intg ((the_Intg v1) - (the_Intg v2))"
-
--- "Be aware of the explicit coercion of the shift distance to nat"
-"eval_binop LShift  v1 v2 = Intg ((the_Intg v1) *   (2^(nat (the_Intg v2))))"
-"eval_binop RShift  v1 v2 = Intg ((the_Intg v1) div (2^(nat (the_Intg v2))))"
-"eval_binop RShiftU v1 v2 = Intg 42" --"FIXME: Not yet implemented"
-
-"eval_binop Less    v1 v2 = Bool ((the_Intg v1) < (the_Intg v2))" 
-"eval_binop Le      v1 v2 = Bool ((the_Intg v1) \<le> (the_Intg v2))"
-"eval_binop Greater v1 v2 = Bool ((the_Intg v2) < (the_Intg v1))"
-"eval_binop Ge      v1 v2 = Bool ((the_Intg v2) \<le> (the_Intg v1))"
-
-"eval_binop Eq      v1 v2 = Bool (v1=v2)"
-"eval_binop Neq     v1 v2 = Bool (v1\<noteq>v2)"
-"eval_binop BitAnd  v1 v2 = Intg 42" -- "FIXME: Not yet implemented"
-"eval_binop And     v1 v2 = Bool ((the_Bool v1) \<and> (the_Bool v2))"
-"eval_binop BitXor  v1 v2 = Intg 42" -- "FIXME: Not yet implemented"
-"eval_binop Xor     v1 v2 = Bool ((the_Bool v1) \<noteq> (the_Bool v2))"
-"eval_binop BitOr   v1 v2 = Intg 42" -- "FIXME: Not yet implemented"
-"eval_binop Or      v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
-"eval_binop CondAnd v1 v2 = Bool ((the_Bool v1) \<and> (the_Bool v2))"
-"eval_binop CondOr  v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
-
-constdefs need_second_arg :: "binop \<Rightarrow> val \<Rightarrow> bool"
-"need_second_arg binop v1 \<equiv> \<not> ((binop=CondAnd \<and>  \<not> the_Bool v1) \<or>
-                               (binop=CondOr  \<and> the_Bool v1))"
-text {* @{term CondAnd} and @{term CondOr} only evalulate the second argument
- if the value isn't already determined by the first argument*}
-
-lemma need_second_arg_CondAnd [simp]: "need_second_arg CondAnd (Bool b) = b" 
-by (simp add: need_second_arg_def)
-
-lemma need_second_arg_CondOr [simp]: "need_second_arg CondOr (Bool b) = (\<not> b)" 
-by (simp add: need_second_arg_def)
-
-lemma need_second_arg_strict[simp]: 
- "\<lbrakk>binop\<noteq>CondAnd; binop\<noteq>CondOr\<rbrakk> \<Longrightarrow> need_second_arg binop b"
-by (cases binop) 
-   (simp_all add: need_second_arg_def)
-
 consts
   eval   :: "prog \<Rightarrow> (state \<times> term    \<times> vals \<times> state) set"
   halloc::  "prog \<Rightarrow> (state \<times> obj_tag \<times> loc  \<times> state) set"
@@ -546,7 +520,7 @@
   "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>     (x,s')" <= "(s     ,x,s') \<in> sxalloc G"
   "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>        s' " == "(s     ,  s') \<in> sxalloc G"
 
-inductive "halloc G" intros (* allocating objects on the heap, cf. 12.5 *)
+inductive "halloc G" intros --{* allocating objects on the heap, cf. 12.5 *}
 
   Abrupt: 
   "G\<turnstile>(Some x,s) \<midarrow>halloc oi\<succ>arbitrary\<rightarrow> (Some x,s)"
@@ -557,79 +531,84 @@
             \<Longrightarrow>
 	    G\<turnstile>Norm s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,init_obj G oi' (Heap a) s)"
 
-inductive "sxalloc G" intros (* allocating exception objects for
-	 	 	      standard exceptions (other than OutOfMemory) *)
+inductive "sxalloc G" intros --{* allocating exception objects for
+	 	 	      standard exceptions (other than OutOfMemory) *}
 
   Norm:	 "G\<turnstile> Norm              s   \<midarrow>sxalloc\<rightarrow>  Norm             s"
 
+  Jmp:   "G\<turnstile>(Some (Jump j), s)  \<midarrow>sxalloc\<rightarrow> (Some (Jump j), s)"
+
+  Error: "G\<turnstile>(Some (Error e), s)  \<midarrow>sxalloc\<rightarrow> (Some (Error e), s)"
+
   XcptL: "G\<turnstile>(Some (Xcpt (Loc a) ),s)  \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s)"
 
   SXcpt: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>halloc (CInst (SXcpt xn))\<succ>a\<rightarrow> (x,s1)\<rbrakk> \<Longrightarrow>
 	  G\<turnstile>(Some (Xcpt (Std xn)),s0) \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s1)"
 
+
 text {* 
 \par
 *} (* dummy text command to break paragraph for latex;
               large paragraphs exhaust memory of debian pdflatex *)
 
+
 inductive "eval G" intros
 
-(* propagation of abrupt completion *)
+--{* propagation of abrupt completion *}
 
-  (* cf. 14.1, 15.5 *)
+  --{* cf. 14.1, 15.5 *}
   Abrupt: 
    "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (arbitrary3 t,(Some xc,s))"
 
 
-(* execution of statements *)
+--{* execution of statements *}
 
-  (* cf. 14.5 *)
+  --{* cf. 14.5 *}
   Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<rightarrow> Norm s"
 
-  (* cf. 14.7 *)
+  --{* cf. 14.7 *}
   Expr:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 				  G\<turnstile>Norm s0 \<midarrow>Expr e\<rightarrow> s1"
 
   Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<rightarrow> s1\<rbrakk> \<Longrightarrow>
                                 G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<rightarrow> abupd (absorb l) s1"
-  (* cf. 14.2 *)
+  --{* cf. 14.2 *}
   Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<rightarrow> s1;
 	  G\<turnstile>     s1 \<midarrow>c2 \<rightarrow> s2\<rbrakk> \<Longrightarrow>
 				 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<rightarrow> s2"
 
-
-  (* cf. 14.8.2 *)
+  --{* cf. 14.8.2 *}
   If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
 	  G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<rightarrow> s2"
 
-  (* cf. 14.10, 14.10.1 *)
-  (*      G\<turnstile>Norm s0 \<midarrow>If(e) (c;; While(e) c) Else Skip\<rightarrow> s3 *)
-  (* A "continue jump" from the while body c is handled by 
-     this rule. If a continue jump with the proper label was invoked inside c
-     this label (Cont l) is deleted out of the abrupt component of the state 
-     before the iterative evaluation of the while statement.
-     A "break jump" is handled by the Lab Statement (Lab l (while\<dots>).
-  *)
+  --{* cf. 14.10, 14.10.1 *}
+  
+  --{* A continue jump from the while body @{term c} is handled by 
+     this rule. If a continue jump with the proper label was invoked inside 
+     @{term c} this label (Cont l) is deleted out of the abrupt component of 
+     the state before the iterative evaluation of the while statement.
+     A break jump is handled by the Lab Statement @{text "Lab l (while\<dots>)"}.
+  *}
   Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
-	  if normal s1 \<and> the_Bool b 
+	  if the_Bool b 
              then (G\<turnstile>s1 \<midarrow>c\<rightarrow> s2 \<and> 
                    G\<turnstile>(abupd (absorb (Cont l)) s2) \<midarrow>l\<bullet> While(e) c\<rightarrow> s3)
 	     else s3 = s1\<rbrakk> \<Longrightarrow>
 			      G\<turnstile>Norm s0 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
 
-  Do: "G\<turnstile>Norm s \<midarrow>Do j\<rightarrow> (Some (Jump j), s)"
+  Jmp: "G\<turnstile>Norm s \<midarrow>Jmp j\<rightarrow> (Some (Jump j), s)"
    
-  (* cf. 14.16 *)
+  --{* cf. 14.16 *}
   Throw: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 				 G\<turnstile>Norm s0 \<midarrow>Throw e\<rightarrow> abupd (throw a') s1"
 
-  (* cf. 14.18.1 *)
+  --{* cf. 14.18.1 *}
   Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; 
 	  if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2\<rbrakk> \<Longrightarrow>
 		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(C vn) c2\<rightarrow> s3"
 
-  (* cf. 14.18.2 *)
+  --{* cf. 14.18.2 *}
   Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1,s1);
 	  G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2;
           s3=(if (\<exists> err. x1=Some (Error err)) 
@@ -637,7 +616,7 @@
               else abupd (abrupt_if (x1\<noteq>None) x1) s2) \<rbrakk> 
           \<Longrightarrow>
           G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<rightarrow> s3"
-  (* cf. 12.4.2, 8.5 *)
+  --{* cf. 12.4.2, 8.5 *}
   Init:	"\<lbrakk>the (class G C) = c;
 	  if inited C (globs s0) then s3 = Norm s0
 	  else (G\<turnstile>Norm (init_class_obj G C s0) 
@@ -645,50 +624,51 @@
 	       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
+   --{* 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
+      (1) The current class object (the static fields) are initialised
+           (@{text init_class_obj}),
+      (2) 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
+      But we can't just naively toggle 1 and 2. By calling 
+      @{text 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). 
+      @{text "AxCompl.thy"}. Static initialisation requires an induction on 
+      the number of classes not yet initialised (or to be more precise, 
+      classes were 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 *)
+      when allocating class objects, we can leave things as they are for 
+      convenience. 
+   *}
+--{* evaluation of expressions *}
 
-  (* cf. 15.8.1, 12.4.1 *)
+  --{* cf. 15.8.1, 12.4.1 *}
   NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1;
 	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
 
-  (* cf. 15.9.1, 12.4.1 *)
+  --{* cf. 15.9.1, 12.4.1 *}
   NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<rightarrow> s2; 
 	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
 	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<rightarrow> s3"
 
-  (* cf. 15.15 *)
+  --{* cf. 15.15 *}
   Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
 	  s2 = abupd (raise_if (\<not>G,store s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
 			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<rightarrow> s2"
 
-  (* cf. 15.19.2 *)
+  --{* cf. 15.19.2 *}
   Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
 	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
 			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<rightarrow> s1"
 
-  (* cf. 15.7.1 *)
+  --{* cf. 15.7.1 *}
   Lit:	"G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<rightarrow> Norm s"
 
   UnOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> 
@@ -700,25 +680,43 @@
           \<rbrakk> 
          \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<rightarrow> s2"
    
-  (* cf. 15.10.2 *)
+  --{* cf. 15.10.2 *}
   Super: "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<rightarrow> Norm s"
 
-  (* cf. 15.2 *)
+  --{* cf. 15.2 *}
   Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<rightarrow> s1"
 
-  (* cf. 15.25.1 *)
+  --{* cf. 15.25.1 *}
   Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<rightarrow> s1;
           G\<turnstile>     s1 \<midarrow>e-\<succ>v  \<rightarrow> s2\<rbrakk> \<Longrightarrow>
 				   G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v\<rightarrow> assign f v s2"
 
-  (* cf. 15.24 *)
+  --{* cf. 15.24 *}
   Cond:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1;
           G\<turnstile>     s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 			    G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<rightarrow> s2"
 
 
-  (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5 *)
+-- {* The interplay of  @{term Call}, @{term Methd} and @{term Body}:
+      Method invocation is split up into these three rules:
+      \begin{itemize}
+      \item [@{term Call}] Calculates the target address and evaluates the
+                           arguments of the method, and then performs dynamic
+                           or static lookup of the method, corresponding to the
+                           call mode. Then the @{term Methd} rule is evaluated
+                           on the calculated declaration class of the method
+                           invocation.
+      \item [@{term Methd}] A syntactic bridge for the folded method body.
+                            It is used by the axiomatic semantics to add the
+                            proper hypothesis for recursive calls of the method.
+      \item [@{term Body}] An extra syntactic entity for the unfolded method
+                           body was introduced to properly trigger class 
+                           initialisation. Without class initialisation we 
+                           could just evaluate the body statement. 
+      \end{itemize}
+   *}
+  --{* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5 *}
   Call:	
   "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1; G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2;
     D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>;
@@ -727,59 +725,66 @@
     G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<rightarrow> s4\<rbrakk>
    \<Longrightarrow>
        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
-*)
+--{* The accessibility check is after @{term init_lvars}, to keep it simple. 
+   @{term 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"
-  (* The local variables l are just a dummy here. The are only used by
-     the smallstep semantics *)
-  (* cf. 14.15, 12.4.1 *)
-  Body:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<rightarrow> s2\<rbrakk> \<Longrightarrow>
-           G\<turnstile>Norm s0 \<midarrow>Body D c
-            -\<succ>the (locals (store s2) Result)\<rightarrow>abupd (absorb Ret) s2"
-  (* The local variables l are just a dummy here. The are only used by
-     the smallstep semantics *)
-(* evaluation of variables *)
+  
+  Body:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<rightarrow> s2; 
+          s3 = (if (\<exists> l. abrupt s2 = Some (Jump (Break l)) \<or>  
+                         abrupt s2 = Some (Jump (Cont l)))
+                  then abupd (\<lambda> x. Some (Error CrossMethodJump)) s2 
+                  else s2)\<rbrakk> \<Longrightarrow>
+           G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
+              \<rightarrow>abupd (absorb Ret) s3"
+  --{* cf. 14.15, 12.4.1 *}
+  --{* We filter out a break/continue in @{term s2}, so that we can proof 
+     definite assignment
+     correct, without the need of conformance of the state. By this the
+     different parts of the typesafety proof can be disentangled a little. *}
 
-  (* cf. 15.13.1, 15.7.2 *)
+--{* evaluation of variables *}
+
+  --{* cf. 15.13.1, 15.7.2 *}
   LVar:	"G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<rightarrow> Norm s"
 
-  (* cf. 15.10.1, 12.4.1 *)
+  --{* cf. 15.10.1, 12.4.1 *}
   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
-  *)
+ --{* The accessibility check is after @{term fvar}, to keep it simple. 
+    @{term fvar} already tests for the absence of a null-pointer reference 
+    in case of an instance field
+  *}
 
-  (* cf. 15.12.1, 15.25.1 *)
+  --{* cf. 15.12.1, 15.25.1 *}
   AVar:	"\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2;
 	  (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
 	              G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<rightarrow> s2'"
 
 
-(* evaluation of expression lists *)
+--{* evaluation of expression lists *}
 
-  (* cf. 15.11.4.2 *)
+  --{* cf. 15.11.4.2 *}
   Nil:
 				    "G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<rightarrow> Norm s0"
 
-  (* cf. 15.6.4 *)
+  --{* cf. 15.6.4 *}
   Cons:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<rightarrow> s1;
           G\<turnstile>     s1 \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 				   G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<rightarrow> s2"
 
 (* Rearrangement of premisses:
-[0,1(Abrupt),2(Skip),8(Do),4(Lab),30(Nil),31(Cons),27(LVar),17(Cast),18(Inst),
+[0,1(Abrupt),2(Skip),8(Jmp),4(Lab),30(Nil),31(Cons),27(LVar),17(Cast),18(Inst),
  17(Lit),18(UnOp),19(BinOp),20(Super),21(Acc),3(Expr),5(Comp),25(Methd),26(Body),23(Cond),6(If),
  7(Loop),11(Fin),9(Throw),13(NewC),14(NewA),12(Init),22(Ass),10(Try),28(FVar),
  29(AVar),24(Call)]
 *)
+
 ML {*
 bind_thm ("eval_induct_", rearrange_prems 
 [0,1,2,8,4,30,31,27,15,16,
@@ -807,12 +812,16 @@
 
 inductive_cases sxalloc_elim_cases:
  	"G\<turnstile> Norm                 s  \<midarrow>sxalloc\<rightarrow> s'"
- 	"G\<turnstile>(Some (Xcpt (Loc a )),s) \<midarrow>sxalloc\<rightarrow> s'"
+        "G\<turnstile>(Some (Jump j),s) \<midarrow>sxalloc\<rightarrow> s'"
+ 	"G\<turnstile>(Some (Error e),s) \<midarrow>sxalloc\<rightarrow> s'"
+        "G\<turnstile>(Some (Xcpt (Loc a )),s) \<midarrow>sxalloc\<rightarrow> s'"
  	"G\<turnstile>(Some (Xcpt (Std xn)),s) \<midarrow>sxalloc\<rightarrow> s'"
 inductive_cases sxalloc_cases: "G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'"
 
 lemma sxalloc_elim_cases2: "\<lbrakk>G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s';  
        \<And>s.   \<lbrakk>s' = Norm s\<rbrakk> \<Longrightarrow> P;  
+       \<And>j s. \<lbrakk>s' = (Some (Jump j),s)\<rbrakk> \<Longrightarrow> P;
+       \<And>e s. \<lbrakk>s' = (Some (Error e),s)\<rbrakk> \<Longrightarrow> P;
        \<And>a s. \<lbrakk>s' = (Some (Xcpt (Loc a)),s)\<rbrakk> \<Longrightarrow> P  
       \<rbrakk> \<Longrightarrow> P"
 apply cut_tac 
@@ -827,10 +836,10 @@
 *}
 inductive_cases eval_cases: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> vs'"
 
-inductive_cases eval_elim_cases:
+inductive_cases eval_elim_cases [cases set]:
         "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 (Jmp 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'"
@@ -882,7 +891,12 @@
 apply auto
 done
 
-
+text {* The following simplification procedures set up the proper injections of
+ terms and their corresponding values in the evaluation relation:
+ E.g. an expression 
+ (injection @{term In1l} into terms) always evaluates to ordinary values 
+ (injection @{term In1} into generalised values @{term vals}). 
+*}
 ML_setup {*
 fun eval_fun nam inj rhs =
 let
@@ -1133,6 +1147,16 @@
 apply force
 done
 
+lemma eval_Body: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<rightarrow> s2;
+                   res=the (locals (store s2) Result);
+                   s3 = (if (\<exists> l. abrupt s2 = Some (Jump (Break l)) \<or>
+                                  abrupt s2 = Some (Jump (Cont l))) 
+                   then abupd (\<lambda> x. Some (Error CrossMethodJump)) s2 
+                   else s2);
+                   s4=abupd (absorb Ret) s3\<rbrakk> \<Longrightarrow>
+ G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>res\<rightarrow>s4"
+by (auto elim: eval.Body)
+
 lemma eval_binop_arg2_indep:
 "\<not> need_second_arg binop v1 \<Longrightarrow> eval_binop binop v1 x = eval_binop binop v1 y"
 by (cases binop)
@@ -1195,11 +1219,7 @@
 lemma split_pairD: "(x,y) = p \<Longrightarrow> x = fst p & y = snd p"
 by auto
 
-lemma eval_Body: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<rightarrow> s2;
-                   res=the (locals (store s2) Result);
-                   s3=abupd (absorb Ret) s2\<rbrakk> \<Longrightarrow>
- G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>res\<rightarrow>s3"
-by (auto elim: eval.Body)
+
 
 lemma unique_eval [rule_format (no_asm)]: 
   "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws \<Longrightarrow> (\<forall>ws'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws' \<longrightarrow> ws' = ws)"
--- a/src/HOL/Bali/Evaln.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Evaln.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -7,13 +7,26 @@
           statements
 *}
 
-theory Evaln = Eval + TypeSafe:
+theory Evaln = TypeSafe:
+
 
 text {*
-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.
+Variant of @{term eval} relation with counter for bounded recursive depth. 
+In principal @{term evaln} could replace @{term eval}.
+
+Validity of the axiomatic semantics builds on @{term evaln}. 
+For recursive method calls the axiomatic semantics rule assumes the method ok 
+to derive a proof for the body. To prove the method rule sound we need to 
+perform induction on the recursion depth. 
+For the completeness proof of the axiomatic semantics the notion of the most
+general formula is used. The most general formula right now builds on the 
+ordinary evaluation relation @{term eval}. 
+So sometimes we have to switch between @{term evaln} and @{term eval} and vice 
+versa. To make
+this switch easy @{term evaln} also does all the technical accessibility tests 
+@{term check_field_access} and @{term check_method_access} like @{term eval}. 
+If it would omit them @{term evaln} and @{term eval} would only be equivalent 
+for welltyped, and definitely assigned terms.
 *}
 
 consts
@@ -63,18 +76,19 @@
 
 inductive "evaln G" intros
 
-(* propagation of abrupt completion *)
+--{* propagation of abrupt completion *}
 
   Abrupt:   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (arbitrary3 t,(Some xc,s))"
 
 
-(* evaluation of variables *)
+--{* evaluation of variables *}
 
   LVar:	"G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<midarrow>n\<rightarrow> Norm s"
 
-  FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init 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'"
+  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;
+          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\<midarrow>n\<rightarrow> s3"
 
   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>
@@ -83,7 +97,7 @@
 
 
 
-(* evaluation of expressions *)
+--{* evaluation of expressions *}
 
   NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1;
 	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
@@ -127,19 +141,25 @@
   Call:	
   "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2;
     D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>; 
-    G\<turnstile>init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs s2
-            \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<midarrow>n\<rightarrow> s3\<rbrakk>
+    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\<midarrow>n\<rightarrow> s4
+   \<rbrakk>
    \<Longrightarrow> 
-    G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<midarrow>n\<rightarrow> (restore_lvars s2 s3)"
+    G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<midarrow>n\<rightarrow> (restore_lvars s2 s4)"
 
   Methd:"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 				G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<midarrow>Suc n\<rightarrow> s1"
 
-  Body:	"\<lbrakk>G\<turnstile>Norm s0\<midarrow>Init D\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2\<rbrakk>\<Longrightarrow>
+  Body:	"\<lbrakk>G\<turnstile>Norm s0\<midarrow>Init D\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2;
+          s3 = (if (\<exists> l. abrupt s2 = Some (Jump (Break l)) \<or>  
+                         abrupt s2 = Some (Jump (Cont l)))
+                  then abupd (\<lambda> x. Some (Error CrossMethodJump)) s2 
+                  else s2)\<rbrakk>\<Longrightarrow>
          G\<turnstile>Norm s0 \<midarrow>Body D c
-          -\<succ>the (locals (store s2) Result)\<midarrow>n\<rightarrow>abupd (absorb Ret) s2"
+          -\<succ>the (locals (store s2) Result)\<midarrow>n\<rightarrow>abupd (absorb Ret) s3"
 
-(* evaluation of expression lists *)
+--{* evaluation of expression lists *}
 
   Nil:
 				"G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<midarrow>n\<rightarrow> Norm s0"
@@ -149,7 +169,7 @@
 			     G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<midarrow>n\<rightarrow> s2"
 
 
-(* execution of statements *)
+--{* execution of statements *}
 
   Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<midarrow>n\<rightarrow> Norm s"
 
@@ -168,13 +188,13 @@
 		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<midarrow>n\<rightarrow> s2"
 
   Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
-	  if normal s1 \<and> the_Bool b 
+	  if the_Bool b 
              then (G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2 \<and> 
                    G\<turnstile>(abupd (absorb (Cont l)) s2) \<midarrow>l\<bullet> While(e) c\<midarrow>n\<rightarrow> s3)
 	     else s3 = s1\<rbrakk> \<Longrightarrow>
 			      G\<turnstile>Norm s0 \<midarrow>l\<bullet> While(e) c\<midarrow>n\<rightarrow> s3"
   
-  Do: "G\<turnstile>Norm s \<midarrow>Do j\<midarrow>n\<rightarrow> (Some (Jump j), s)"
+  Jmp: "G\<turnstile>Norm s \<midarrow>Jmp j\<midarrow>n\<rightarrow> (Some (Jump j), s)"
   
   Throw:"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 				 G\<turnstile>Norm s0 \<midarrow>Throw e\<midarrow>n\<rightarrow> abupd (throw a') s1"
@@ -185,8 +205,11 @@
 		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(tn vn) c2\<midarrow>n\<rightarrow> s3"
 
   Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> (x1,s1);
-	  G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
-              G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<midarrow>n\<rightarrow> abupd (abrupt_if (x1\<noteq>None) x1) s2"
+	  G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n\<rightarrow> s2;
+          s3=(if (\<exists> err. x1=Some (Error err)) 
+              then (x1,s1) 
+              else abupd (abrupt_if (x1\<noteq>None) x1) s2)\<rbrakk> \<Longrightarrow>
+              G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<midarrow>n\<rightarrow> s3"
   
   Init:	"\<lbrakk>the (class G C) = c;
 	  if inited C (globs s0) then s3 = Norm s0
@@ -202,12 +225,17 @@
 
 declare split_if     [split del] split_if_asm     [split del]
         option.split [split del] option.split_asm [split del]
+        not_None_eq [simp del] 
+        split_paired_All [simp del] split_paired_Ex [simp del]
+ML_setup {*
+simpset_ref() := simpset() delloop "split_all_tac"
+*}
 inductive_cases evaln_cases: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> vs'"
 
 inductive_cases evaln_elim_cases:
 	"G\<turnstile>(Some xc, s) \<midarrow>t                        \<succ>\<midarrow>n\<rightarrow> vs'"
 	"G\<turnstile>Norm s \<midarrow>In1r Skip                      \<succ>\<midarrow>n\<rightarrow> xs'"
-        "G\<turnstile>Norm s \<midarrow>In1r (Do j)                    \<succ>\<midarrow>n\<rightarrow> xs'"
+        "G\<turnstile>Norm s \<midarrow>In1r (Jmp j)                   \<succ>\<midarrow>n\<rightarrow> xs'"
         "G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> c)                    \<succ>\<midarrow>n\<rightarrow> xs'"
 	"G\<turnstile>Norm s \<midarrow>In3  ([])                      \<succ>\<midarrow>n\<rightarrow> vs'"
 	"G\<turnstile>Norm s \<midarrow>In3  (e#es)                    \<succ>\<midarrow>n\<rightarrow> vs'"
@@ -236,9 +264,15 @@
 	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                 \<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'"
+        "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]
-
+        not_None_eq [simp] 
+        split_paired_All [simp] split_paired_Ex [simp]
+ML_setup {*
+simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
+*}
 lemma evaln_Inj_elim: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (w,s') \<Longrightarrow> case t of In1 ec \<Rightarrow>  
   (case ec of Inl e \<Rightarrow> (\<exists>v. w = In1 v) | Inr c \<Rightarrow> w = \<diamondsuit>)  
   | In2 e \<Rightarrow> (\<exists>v. w = In2 v) | In3 e \<Rightarrow> (\<exists>v. w = In3 v)"
@@ -248,6 +282,13 @@
 apply auto
 done
 
+text {* The following simplification procedures set up the proper injections of
+ terms and their corresponding values in the evaluation relation:
+ E.g. an expression 
+ (injection @{term In1l} into terms) always evaluates to ordinary values 
+ (injection @{term In1} into generalised values @{term vals}). 
+*}
+
 ML_setup {*
 fun enf nam inj rhs =
 let
@@ -390,663 +431,47 @@
 apply auto
 done
 
-(* ##### FIXME: To WellType *)
-lemma wt_elim_BinOp:
-  "\<lbrakk>E,dt\<Turnstile>In1l (BinOp binop e1 e2)\<Colon>T;
-    \<And>T1 T2 T3.
-       \<lbrakk>E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2; wt_binop (prg E) binop T1 T2;
-          E,dt\<Turnstile>(if b then In1l e2 else In1r Skip)\<Colon>T3;
-          T = Inl (PrimT (binop_type binop))\<rbrakk>
-       \<Longrightarrow> P\<rbrakk>
-\<Longrightarrow> P"
-apply (erule wt_elim_cases)
-apply (cases b)
-apply auto
-done
+
+
 
 section {* evaln implies eval *}
+
 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" 
-       
+  assumes evaln: "G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)" 
   shows "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 UnOp
-    with wf show ?case
-      by - (erule wt_elim_cases, rule eval.UnOp,auto dest: eval_type_sound)
-  next
-    case BinOp
-    with wf show ?case 
-      by - (erule wt_elim_BinOp, blast intro!: eval.BinOp dest: eval_type_sound)
-  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)
-    txt {* Repeats large parts of the type soundness proof. One should factor
-      out some lemmata about the relations and conformance of @{text
-      s2}, @{text s3} and @{text 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) fastsimp+
-    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,
-                auto 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 c1 c2 n s0 s1 s2 x1 L accC T)
-    have hyp_c1: "PROP ?EqEval (Norm s0) (x1,s1) (In1r c1) \<diamondsuit>" .
-    have hyp_c2: "PROP ?EqEval (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
-    have eval_c1: "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1, s1)"
-      by (rule hyp_c1)
-    with wf wt_c1 conf_s0
-    obtain       conf_s1: "Norm s1\<Colon>\<preceq>(G, L)" and 
-           error_free_s1: "error_free (x1,s1)"
-      by (auto dest!: eval_type_sound intro: conforms_NormI)
-    from conf_s1 wt_c2
-    have eval_c2: "G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2"
-      by (rule hyp_c2)
-    with eval_c1 error_free_s1
-    show "G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<rightarrow> abupd (abrupt_if (x1 \<noteq> None) x1) s2"
-      by (auto intro: eval.Fin simp add: error_free_def)
-  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
-    with cls show ?case
-      by (rule eval.Init)
-  qed 
-qed
+using evaln 
+proof (induct)
+  case (Loop b c e l n s0 s1 s2 s3)
+  have "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1".
+  moreover
+  have "if 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"
+    using Loop.hyps by simp
+  ultimately show ?case by (rule eval.Loop)
+next
+  case (Try c1 c2 n s0 s1 s2 s3 C vn)
+  have "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1".
+  moreover
+  have "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2".
+  moreover
+  have "if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2"
+    using Try.hyps by simp
+  ultimately show ?case by (rule eval.Try)
+next
+  case (Init C c n s0 s1 s2 s3)
+  have "the (class G C) = c".
+  moreover
+  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 = (set_lvars (locals (store s1))) s2"
+    using Init.hyps by simp
+  ultimately show ?case by (rule eval.Init)
+qed (rule eval.intros,(assumption+ | assumption?))+
 
 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)
@@ -1068,8 +493,13 @@
 
 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
+by (fast intro: le_maxI1 le_maxI2)
+
+corollary evaln_max2E [consumes 2]:
+  "\<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>n1\<rightarrow> ws1; G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>n2\<rightarrow> ws2; 
+    \<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max n1 n2\<rightarrow> ws1;G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max n1 n2\<rightarrow> ws2 \<rbrakk> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
+by (drule (1) evaln_max2) simp
+
 
 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>
@@ -1080,6 +510,16 @@
 apply (fast intro!: le_maxI1 le_maxI2)
 done
 
+corollary evaln_max3E: 
+"\<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;
+   \<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws1;
+    G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws2; 
+    G\<turnstile>s3 \<midarrow>t3\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws3
+   \<rbrakk> \<Longrightarrow> P
+  \<rbrakk> \<Longrightarrow> P"
+by (drule (2) evaln_max3) simp
+
+
 lemma le_max3I1: "(n2::nat) \<le> max n1 (max n2 n3)"
 proof -
   have "n2 \<le> max n2 n3"
@@ -1102,755 +542,328 @@
   show ?thesis .
 qed
 
+ML {*
+Delsimprocs [wt_expr_proc,wt_var_proc,wt_exprs_proc,wt_stmt_proc]
+*}
+
 section {* eval implies evaln *}
 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"  
+  assumes eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
   shows  "\<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 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)
+using eval 
+proof (induct)
+  case (Abrupt s t xc)
+  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)
+  then obtain n where
+    "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+    by (rules)
+  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)
+  then obtain n where
+    "G\<turnstile>Norm s0 \<midarrow>c\<midarrow>n\<rightarrow> s1"
+    by (rules)
+  then have "G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<midarrow>n\<rightarrow> abupd (absorb l) s1"
+    by (rule evaln.Lab)
+  then show ?case ..
+next
+  case (Comp c1 c2 s0 s1 s2)
+  then obtain n1 n2 where
+    "G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n1\<rightarrow> s1"
+    "G\<turnstile>s1 \<midarrow>c2\<midarrow>n2\<rightarrow> s2"
+    by (rules)
+  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)
+  then 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 (rules)
+  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)
+  from Loop.hyps obtain n1 where
+    "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n1\<rightarrow> s1"
+    by (rules)
+  moreover from Loop.hyps obtain n2 where
+    "if 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"
+    by simp (rules intro: evaln_nonstrict le_maxI1 le_maxI2)
+  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 useful 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 s3 x1 L accC T)
-    have s3: "s3 = (if \<exists>err. x1 = Some (Error err) 
-                       then (x1, s1)
-                       else abupd (abrupt_if (x1 \<noteq> None) x1) s2)" .
-    from Fin 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" and
-      error_free_s1: "error_free (x1,s1)"
-      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)
-    with error_free_s1 s3
-    show "\<exists>n. G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<midarrow>n\<rightarrow> s3"
-      by (auto simp add: error_free_def)
-  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
+    apply   (auto intro: evaln_nonstrict intro: le_maxI2)
+    done
+  then show ?case ..
+next
+  case (Jmp j s)
+  have "G\<turnstile>Norm s \<midarrow>Jmp j\<midarrow>n\<rightarrow> (Some (Jump j), s)"
+    by (rule evaln.Jmp)
+  then show ?case ..
+next
+  case (Throw a e s0 s1)
+  then obtain n where
+    "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1"
+    by (rules)
+  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)
+  from Try.hyps obtain n1 where
+    "G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n1\<rightarrow> s1"
+    by (rules)
+  moreover 
+  have sxalloc: "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" .
+  moreover
+  from Try.hyps 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"
+    by fastsimp 
+  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 s3 x1)
+  from Fin 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 rules
+  moreover
+  have s3: "s3 = (if \<exists>err. x1 = Some (Error err) 
+                     then (x1, s1)
+                     else abupd (abrupt_if (x1 \<noteq> None) x1) s2)" .
+  ultimately 
+  have 
+    "G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<midarrow>max n1 n2\<rightarrow> s3"
+    by (blast intro: evaln.Fin dest: evaln_max2)
+  then show ?case ..
+next
+  case (Init C c s0 s1 s2 s3)
+  have     cls: "the (class G C) = c" .
+  moreover from Init.hyps 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 (UnOp e s0 s1 unop 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)
-    hence "G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>eval_unop unop v\<midarrow>n\<rightarrow> s1"
-      by (rule evaln.UnOp)
-    then show ?case ..
-  next
-    case (BinOp binop e1 e2 s0 s1 s2 v1 v2 L accC T)
-    with wf obtain n1 n2 where 
-      "G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<midarrow>n1\<rightarrow> s1"
-      "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
+    by (auto intro: evaln_nonstrict le_maxI1 le_maxI2)
+  ultimately 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)
+  then obtain n where 
+    "G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1"
+    by (rules)
+  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)
+  then 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 (rules)
+  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)
+  then obtain n where
+    "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+    by (rules)
+  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)
+  then obtain n where
+    "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+    by (rules)
+  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)
+  have "G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<midarrow>n\<rightarrow> Norm s"
+    by (rule evaln.Lit)
+  then show ?case ..
+next
+  case (UnOp e s0 s1 unop v )
+  then obtain n where
+    "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+    by (rules)
+  hence "G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>eval_unop unop v\<midarrow>n\<rightarrow> s1"
+    by (rule evaln.UnOp)
+  then show ?case ..
+next
+  case (BinOp binop e1 e2 s0 s1 s2 v1 v2)
+  then obtain n1 n2 where 
+    "G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<midarrow>n1\<rightarrow> s1"
+    "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
                else In1r Skip)\<succ>\<midarrow>n2\<rightarrow> (In1 v2, s2)"    
-      by (blast elim!: wt_elim_BinOp dest: eval_type_sound)
-    hence "G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<midarrow>max n1 n2
-           \<rightarrow> s2"
-      by (blast intro!: evaln.BinOp dest: evaln_max2)
-    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) fastsimp+
-    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> 
+    by (rules)
+  hence "G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<midarrow>max n1 n2
+          \<rightarrow> s2"
+    by (blast intro!: evaln.BinOp dest: evaln_max2)
+  then show ?case ..
+next
+  case (Super s )
+  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)
+  then obtain n where
+    "G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v, f)\<midarrow>n\<rightarrow> s1"
+    by (rules)
+  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)
+  then 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 (rules)
+  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)
+  then obtain n1 n2 where 
+    "G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<midarrow>n1\<rightarrow> s1"
+    "G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<midarrow>n2\<rightarrow> s2"
+    by (rules)
+  then
+  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)
+  then obtain n1 n2 where
+    "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<midarrow>n1\<rightarrow> s1"
+    "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n2\<rightarrow> s2"
+    by rules
+  moreover
+  have "invDeclC = invocation_declclass G mode (store s2) a' statT 
+                       \<lparr>name=mn,parTs=pTs'\<rparr>" .
+  moreover
+  have "s3 = init_lvars G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> mode a' vs s2" .
+  moreover
+  have "s3'=check_method_access G accC' statT mode \<lparr>name=mn,parTs=pTs'\<rparr> a' s3".
+  moreover 
+  from Call.hyps
+  obtain m where 
+    "G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name=mn, parTs=pTs'\<rparr>-\<succ>v\<midarrow>m\<rightarrow> s4"
+    by rules
+  ultimately
+  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,
-                auto 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 
+    by (auto intro!: evaln.Call le_maxI1 le_max3I1 le_max3I2)
+  thus ?case ..
+next
+  case (Methd D s0 s1 sig v )
+  then obtain n where
+    "G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<midarrow>n\<rightarrow> s1"
+    by rules
+  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 s3 )
+  from Body.hyps obtain n1 n2 where 
+    evaln_init: "G\<turnstile>Norm s0 \<midarrow>Init D\<midarrow>n1\<rightarrow> s1" and
+    evaln_c: "G\<turnstile>s1 \<midarrow>c\<midarrow>n2\<rightarrow> s2"
+    by (rules)
+  moreover
+  have "s3 = (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or> 
+                     fst s2 = Some (Jump (Cont l))
+                 then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 
+                 else s2)".
+  ultimately
+  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) (fastsimp 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
+       \<rightarrow> abupd (absorb Ret) s3"
+    by (rules intro: evaln.Body dest: evaln_max2)
+  then show ?case ..
+next
+  case (LVar s vn )
+  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)
+  then obtain n1 n2 where
+    "G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<midarrow>n1\<rightarrow> s1"
+    "G\<turnstile>s1 \<midarrow>e-\<succ>a\<midarrow>n2\<rightarrow> s2"
+    by rules
+  moreover
+  have "s3 = check_field_access G accC statDeclC fn stat a s2'" 
+       "(v, s2') = fvar statDeclC stat fn a s2" .
+  ultimately
+  have "G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v\<midarrow>max n1 n2\<rightarrow> s3"
+    by (rules intro: evaln.FVar dest: evaln_max2)
+  then show ?case ..
+next
+  case (AVar a e1 e2 i s0 s1 s2 s2' v )
+  then 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 rules
+  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)
+  show ?case by (rules intro: evaln.Nil)
+next
+  case (Cons e es s0 s1 s2 v vs)
+  then 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 rules
+  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
-
+       
 end
--- a/src/HOL/Bali/Example.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Example.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -160,7 +160,8 @@
               \<lparr>access=Public,static=False,pars=[z],resT=Class Ext,
 	       mbody=\<lparr>lcls=[]
                      ,stmt=Expr({Ext,Ext,False}Cast (Class Ext) (!!z)..vee := 
-       	                                                     Lit (Intg 1))\<rparr>
+       	                                                     Lit (Intg 1)) ;;
+                                Return (Lit Null)\<rparr>
 	      \<rparr>)"
 
 constdefs
@@ -883,17 +884,44 @@
 
 declare member_is_static_simp [simp]
 declare wt.Skip [rule del] wt.Init [rule del]
+ML {* bind_thms ("wt_intros",map (rewrite_rule [id_def]) (thms "wt.intros")) *}
+lemmas wtIs = wt_Call wt_Super wt_FVar wt_StatRef wt_intros
+lemmas daIs = assigned.select_convs da_Skip da_NewC da_Lit da_Super da.intros
+
 lemmas Base_foo_defs = Base_foo_def foo_sig_def foo_mhead_def
 lemmas Ext_foo_defs  = Ext_foo_def  foo_sig_def
-ML {* bind_thms ("wt_intros",map (rewrite_rule [id_def]) (thms "wt.intros")) *}
-lemmas wtIs = wt_Call wt_Super wt_FVar wt_StatRef wt_intros
+
+
+
 
 lemma wf_Base_foo: "wf_mdecl tprg Base Base_foo"
 apply (unfold Base_foo_defs )
-apply (auto intro!: wf_mdeclI wf_mheadI intro!: wtIs
+apply (auto intro!: wf_mdeclI wf_mheadI intro!: wtIs 
             simp add: mhead_resTy_simp)
+(* for definite assignment *)
+apply (rule exI)
+apply (simp add: parameters_def)
+apply (rule conjI)
+apply  (rule da.Comp)
+apply     (rule da.Expr)
+apply     (rule da.AssLVar)
+apply       (rule da.AccLVar)
+apply         (simp)
+apply        (rule assigned.select_convs)
+apply       (simp)
+apply      (rule assigned.select_convs)
+apply     (simp)
+apply    (simp)
+apply    (rule da.Jmp)
+apply      (simp)
+apply     (rule assigned.select_convs)
+apply    (simp)
+apply   (rule assigned.select_convs)
+apply  (simp)
+apply (simp)
 done
 
+
 lemma wf_Ext_foo: "wf_mdecl tprg Ext Ext_foo"
 apply (unfold Ext_foo_defs )
 apply (auto intro!: wf_mdeclI wf_mheadI intro!: wtIs 
@@ -903,6 +931,37 @@
 apply    simp
 apply   (rule_tac [2] narrow.subcls [THEN cast.narrow])
 apply   (auto intro!: wtIs)
+(* for definite assignment *)
+apply (rule exI)
+apply (simp add: parameters_def)
+apply (rule conjI)
+apply  (rule da.Comp)
+apply     (rule da.Expr)
+apply     (rule da.Ass)   
+apply       simp
+apply      (rule da.FVar)
+apply      (rule da.Cast)
+apply      (rule da.AccLVar)
+apply        simp
+apply       (rule assigned.select_convs)
+apply      simp
+apply     (rule da_Lit)
+apply     (simp)
+apply    (rule da.Comp)
+apply       (rule da.Expr)
+apply       (rule da.AssLVar)
+apply         (rule da.Lit)                  
+apply        (rule assigned.select_convs)
+apply       simp
+apply      (rule da.Jmp)
+apply        simp
+apply       (rule assigned.select_convs)
+apply      simp
+apply     (rule assigned.select_convs)
+apply    (simp)
+apply   (rule assigned.select_convs)
+apply  simp
+apply simp
 done
 
 declare mhead_resTy_simp [simp add]
@@ -911,23 +970,35 @@
 lemma wf_BaseC: "wf_cdecl tprg (Base,BaseCl)"
 apply (unfold wf_cdecl_def BaseCl_def arr_viewed_from_def)
 apply (auto intro!: wf_Base_foo)
-apply       (auto intro!: ws_cdecl_Base simp add: Base_foo_def foo_mhead_def)
-apply (auto intro!: wtIs)
+apply        (auto intro!: ws_cdecl_Base simp add: Base_foo_def foo_mhead_def)
+apply   (auto intro!: wtIs)
+(* for definite assignment *)
+apply   (rule exI)
+apply   (rule da.Expr)
+apply  (rule da.Ass)
+apply    (simp)
+apply   (rule da.FVar)
+apply   (rule da.Cast)
+apply   (rule da_Lit)
+apply   simp
+apply  (rule da.NewA)
+apply  (rule da.Lit)
 apply (auto simp add: Base_foo_defs entails_def Let_def)
 apply   (insert Base_foo_no_stat_override, simp add: Base_foo_def,blast)+
-apply   (insert Base_foo_no_hide         , simp add: Base_foo_def,blast)
+apply (insert Base_foo_no_hide         , simp add: Base_foo_def,blast)
 done
 
 
 lemma wf_ExtC: "wf_cdecl tprg (Ext,ExtCl)"
 apply (unfold wf_cdecl_def ExtCl_def)
 apply (auto intro!: wf_Ext_foo ws_cdecl_Ext)
-apply (auto simp add: entails_def snd_special_simp)
-apply (insert Ext_foo_stat_override)
-apply  (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
-apply  (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
-apply  (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
-apply (insert Ext_foo_no_hide)
+apply  (auto simp add: entails_def snd_special_simp)
+apply      (insert Ext_foo_stat_override)
+apply      (rule exI,rule da.Skip)      
+apply     (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
+apply    (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
+apply   (force simp add: qmdecl_def Ext_foo_def Base_foo_def)
+apply  (insert Ext_foo_no_hide)
 apply  (simp_all add: qmdecl_def)
 apply  blast+
 done
@@ -935,6 +1006,7 @@
 lemma wf_MainC: "wf_cdecl tprg (Main,MainCl)"
 apply (unfold wf_cdecl_def MainCl_def)
 apply (auto intro: ws_cdecl_Main)
+apply (rule exI,rule da.Skip)
 done
 
 lemma wf_idecl_all: "p=tprg \<Longrightarrow> Ball (set Ifaces) (wf_idecl p)"
@@ -948,7 +1020,8 @@
 apply (unfold standard_classes_def Let_def 
        ObjectC_def SXcptC_def Object_mdecls_def SXcpt_mdecls_def)
 apply (simp (no_asm) add: wf_cdecl_def ws_cdecls)
-apply (auto simp add:is_acc_class_def accessible_in_RefT_simp SXcpt_def)
+apply (auto simp add:is_acc_class_def accessible_in_RefT_simp SXcpt_def 
+            intro: da.Skip)
 apply (auto simp add: Object_def Classes_def standard_classes_def 
                       SXcptC_def SXcpt_def)
 done
@@ -970,14 +1043,13 @@
 apply (rule conjI)
 apply (simp add: Object_mdecls_def)
 apply safe
-apply  (cut_tac xn_cases_old)   (* FIXME (insert xn_cases) *)
+apply  (cut_tac xn_cases)   
 apply  (simp (no_asm_simp) add: Classes_def standard_classes_def)
 apply  (insert wf_idecl_all)
 apply  (insert wf_cdecl_all)
 apply  auto
 done
 
-
 section "max spec"
 
 lemma appl_methds_Base_foo: 
@@ -998,7 +1070,6 @@
 apply auto
 done
 
-
 section "well-typedness"
 
 lemma wt_test: "\<lparr>prg=tprg,cls=Main,lcl=empty(VName e\<mapsto>Class Base)\<rparr>\<turnstile>test ?pTs\<Colon>\<surd>"
@@ -1051,6 +1122,49 @@
 apply (rule wtIs (* Skip *))
 done
 
+section "definite assignment"
+
+lemma da_test: "\<lparr>prg=tprg,cls=Main,lcl=empty(VName e\<mapsto>Class Base)\<rparr>
+                  \<turnstile>{} \<guillemotright>\<langle>test ?pTs\<rangle>\<guillemotright> \<lparr>nrm={VName e},brk=\<lambda> l. UNIV\<rparr>"
+apply (unfold test_def arr_viewed_from_def)
+apply (rule da.Comp)
+apply    (rule da.Expr)
+apply    (rule da.AssLVar)
+apply      (rule da.NewC)
+apply     (rule assigned.select_convs)
+apply    (simp)
+apply   (rule da.Try)
+apply      (rule da.Expr)
+apply      (rule da.Call)
+apply       (rule da.AccLVar)
+apply         (simp)   
+apply        (rule assigned.select_convs)
+apply       (simp)
+apply      (rule da.Cons)
+apply       (rule da.Lit)
+apply      (rule da.Nil)
+apply     (rule da.Loop)
+apply        (rule da.Acc)
+apply         (simp)
+apply        (rule da.AVar)
+apply         (rule da.Acc)
+apply          simp
+apply         (rule da.FVar)
+apply         (rule da.Cast)
+apply         (rule da.Lit)
+apply        (rule da.Lit)
+apply       (rule da_Skip)
+apply       (simp)
+apply      (simp,rule assigned.select_convs)
+apply     (simp)
+apply    (simp,rule assigned.select_convs)
+apply   (simp)
+apply  simp
+apply  blast
+apply simp
+apply (simp add: intersect_ts_def)
+done
+
 
 section "execution"
 
@@ -1212,7 +1326,7 @@
 apply   (simp)
 apply   (rule halloc.New)
 apply    (simp (no_asm_simp))
-apply   (drule atleast_free_weaken,rotate_tac -1,drule atleast_free_weaken)
+apply   (drule atleast_free_weaken,drule atleast_free_weaken)
 apply   (simp (no_asm_simp))
 apply  (simp add: upd_gobj_def)
       (* end init Ext *)
@@ -1251,19 +1365,25 @@
 apply     (rule init_done, simp)
 apply       (simp add: invocation_declclass_def dynlookup_def dynmethd_Ext_foo)
 apply    (simp add: invocation_declclass_def dynlookup_def dynmethd_Ext_foo)
-apply    (rule eval_Is (* Expr *))
-apply    (rule eval_Is (* Ass *))
-apply     (rule eval_Is (* FVar *))
-apply       (rule init_done, simp)
-apply      (rule eval_Is (* Cast *))
-apply       (rule eval_Is (* Acc *))
-apply       (rule eval_Is (* LVar *))
-apply      (simp)
-apply      (simp split del: split_if)
+apply    (rule eval_Is (* Comp *))
+apply     (rule eval_Is (* Expr *))
+apply     (rule eval_Is (* Ass *))
+apply      (rule eval_Is (* FVar *))
+apply         (rule init_done, simp)
+apply        (rule eval_Is (* Cast *))
+apply         (rule eval_Is (* Acc *))
+apply         (rule eval_Is (* LVar *))
+apply        (simp)
+apply       (simp split del: split_if)
 apply      (simp add: check_field_access_def Let_def)
-apply    (rule eval_Is (* XcptE *))
+apply     (rule eval_Is (* XcptE *))
+apply    (simp)
+apply    (rule conjI)
+apply     (simp)
+apply    (rule eval_Is (* Comp *))
 apply   (simp)
       (* end method call *)
+apply  simp
 apply  (rule sxalloc.intros)
 apply  (rule halloc.New)
 apply   (erule alloc_one [THEN conjunct1])
--- a/src/HOL/Bali/Name.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Name.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -8,11 +8,11 @@
 theory Name = Basis:
 
 (* cf. 6.5 *) 
-typedecl tnam		(* ordinary type name, i.e. class or interface name *)
-typedecl pname          (* package name *)
-typedecl mname		(* method name *)
-typedecl vname          (* variable or field name *)
-typedecl label          (* label as destination of break or continue *)
+typedecl tnam	--{* ordinary type name, i.e. class or interface name *}
+typedecl pname  --{* package name *}
+typedecl mname  --{* method name *}
+typedecl vname  --{* variable or field name *}
+typedecl label  --{* label as destination of break or continue *}
 
 arities
   tnam  :: "type"
@@ -22,11 +22,11 @@
   label :: "type"
 
 
-datatype ename        (* expression name *) 
+datatype ename        --{* expression name *} 
         = VNam vname 
-        | Res         (* special name to model the return value of methods *)
+        | Res         --{* special name to model the return value of methods *}
 
-datatype lname        (* names for local variables and the This pointer *)
+datatype lname        --{* names for local variables and the This pointer *}
         = EName ename 
         | This
 syntax   
@@ -37,7 +37,7 @@
   "VName n" == "EName (VNam n)"
   "Result"  == "EName Res"
 
-datatype xname		(* names of standard exceptions *)
+datatype xname		--{* names of standard exceptions *}
 	= Throwable
 	| NullPointer | OutOfMemory | ClassCast   
 	| NegArrSize  | IndOutBound | ArrStore
@@ -50,22 +50,14 @@
 apply auto
 done
 
-lemma xn_cases_old:   (* FIXME cf. Example.thy *)
-  "ALL xn. xn = Throwable   \<or> xn = NullPointer \<or>  
-         xn = OutOfMemory \<or> xn = ClassCast \<or> 
-         xn = NegArrSize  \<or> xn = IndOutBound \<or> xn = ArrStore"
-apply (rule allI)
-apply (induct_tac xn)
-apply auto
-done
 
-datatype tname	(* type names for standard classes and other type names *)
+datatype tname	--{* type names for standard classes and other type names *}
 	= Object_
 	| SXcpt_   xname
 	| TName   tnam
 
-record   qtname = (* qualified tname cf. 6.5.3, 6.5.4*)
-          pid :: pname 
+record   qtname = --{* qualified tname cf. 6.5.3, 6.5.4*}
+          pid :: pname  
           tid :: tname
 
 axclass has_pname < "type"
@@ -102,7 +94,7 @@
   (type) "'a qtname_scheme" <= (type) "\<lparr>pid::pname,tid::tname,\<dots>::'a\<rparr>"
 
 
-consts java_lang::pname (* package java.lang *)
+consts java_lang::pname --{* package java.lang *}
 
 consts 
   Object :: qtname
--- a/src/HOL/Bali/State.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/State.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -19,17 +19,17 @@
 
 section "objects"
 
-datatype  obj_tag =     (* tag for generic object   *)
-	  CInst qtname   (* class instance           *)
-	| Arr  ty int   (* array with component type and length *)
-     (* | CStat qtname     the tag is irrelevant for a class object,
+datatype  obj_tag =     --{* tag for generic object   *}
+	  CInst qtname  --{* class instance           *}
+	| Arr  ty int   --{* array with component type and length *}
+    --{* | 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) *)
+                           it (see below) *}
 
-types	vn   = "fspec + int"                    (* variable name      *)
+types	vn   = "fspec + int"                    --{* variable name      *}
 record	obj  = 
-          tag :: "obj_tag"                      (* generalized object *)
+          tag :: "obj_tag"                      --{* generalized object *}
           values :: "(vn, val) table"      
 
 translations 
@@ -79,11 +79,7 @@
 lemma obj_ty_cong [simp]: 
   "obj_ty (obj \<lparr>values:=vs\<rparr>) = obj_ty obj" 
 by auto
-(*
-lemma obj_ty_cong [simp]: 
-  "obj_ty (obj \<lparr>values:=vs(n\<mapsto>v)\<rparr>) = obj_ty (obj \<lparr>values:=vs\<rparr>)" 
-by auto
-*)
+
 lemma obj_ty_CInst [simp]: 
  "obj_ty \<lparr>tag=CInst C,values=vs\<rparr> = Class C" 
 by (simp add: obj_ty_def)
@@ -137,7 +133,7 @@
 
 section "object references"
 
-types oref = "loc + qtname"          (* generalized object reference *)
+types oref = "loc + qtname"         --{* generalized object reference *}
 syntax
   Heap  :: "loc   \<Rightarrow> oref"
   Stat  :: "qtname \<Rightarrow> oref"
@@ -219,7 +215,7 @@
 
 section "stores"
 
-types	globs                  (* global variables: heap and static variables *)
+types	globs               --{* global variables: heap and static variables *}
 	= "(oref , obj) table"
 	heap
 	= "(loc  , obj) table"
@@ -433,14 +429,6 @@
 apply (simp (no_asm))
 done
 
-(*
-lemma heap_upd_gobj_Heap: "!!a. heap (upd_gobj (Heap a) n v s) = ?X"
-apply (rule ext)
-apply (simp (no_asm))
-apply (case_tac "globs s (Heap a)")
-apply  auto
-*)
-
 lemma heap_upd_gobj_Stat [simp]: "heap (upd_gobj (Stat C) n v s) = heap s"
 apply (rule ext)
 apply (simp (no_asm))
@@ -599,11 +587,17 @@
 lemma absorb_other [simp]: "a \<noteq> Some (Jump j) \<Longrightarrow> absorb j a = a"
 by (auto simp add: absorb_def)
 
+lemma absorb_Some_NoneD: "absorb j (Some abr) = None \<Longrightarrow> abr = Jump j"
+  by (simp add: absorb_def)
+
+lemma absorb_Some_JumpD: "absorb j s = Some (Jump j') \<Longrightarrow> j'\<noteq>j"
+  by (simp add: absorb_def)
+
 
 section "full program state"
 
 types
-  state = "abopt \<times> st"          (* state including exception information *)
+  state = "abopt \<times> st"          --{* state including abruption information *}
 
 syntax 
   Norm   :: "st \<Rightarrow> state"
@@ -690,6 +684,12 @@
 apply (simp (no_asm))
 done
 
+lemma abupd_store_invariant [simp]: "store (abupd f s) = store s"
+  by (cases s) simp
+
+lemma supd_abrupt_invariant [simp]: "abrupt (supd f s) = abrupt s"
+  by (cases s) simp
+
 syntax
 
   set_lvars     :: "locals \<Rightarrow> state \<Rightarrow> state"
@@ -842,5 +842,6 @@
        \<Longrightarrow> error_free (x, set_locals l s')"
 by (simp add: error_free_def)
 
+
 end
 
--- a/src/HOL/Bali/Table.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Table.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -31,16 +31,16 @@
 \end{itemize}
 *}
 
-types ('a, 'b) table    (* table with key type 'a and contents type 'b *)
+types ('a, 'b) table    --{* table with key type 'a and contents type 'b *}
       = "'a \<leadsto> 'b"
-      ('a, 'b) tables   (* non-unique table with key 'a and contents 'b *)
+      ('a, 'b) tables   --{* non-unique table with key 'a and contents 'b *}
       = "'a \<Rightarrow> 'b set"
 
 
 section "map of / table of"
 
 syntax
-  table_of      :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) table"    (* concrete table *)
+  table_of      :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) table"   --{* concrete table *}
   
 translations
   "table_of" == "map_of"
@@ -58,8 +58,8 @@
 cond_override:: 
   "('b \<Rightarrow>'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b) table"
 
-(* when merging tables old and new, only override an entry of table old when  
-   the condition cond holds *)
+--{* when merging tables old and new, only override an entry of table old when  
+   the condition cond holds *}
 "cond_override cond old new \<equiv>
  \<lambda> k.
   (case new k of
@@ -295,10 +295,10 @@
                      ('a, 'b) tables"             (infixl "\<oplus>\<oplus>" 100)
   hidings_entails:: "('a, 'b) tables \<Rightarrow> ('a, 'c) tables \<Rightarrow> 
                      ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"   ("_ hidings _ entails _" 20)
-  (* variant for unique table: *)
+  --{* variant for unique table: *}
   hiding_entails :: "('a, 'b) table  \<Rightarrow> ('a, 'c) table  \<Rightarrow> 
                      ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"   ("_ hiding _ entails _"  20)
-  (* variant for a unique table and conditional overriding: *)
+  --{* variant for a unique table and conditional overriding: *}
   cond_hiding_entails :: "('a, 'b) table  \<Rightarrow> ('a, 'c) table  
                           \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"  
                           ("_ hiding _ under _ entails _"  20)
--- a/src/HOL/Bali/Term.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Term.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -30,7 +30,6 @@
 simplifications:
 \begin{itemize}
 \item expression statement allowed for any expression
-\item no unary, binary, etc, operators
 \item This  is modeled as a special non-assignable local variable
 \item Super is modeled as a general expression with the same value as This
 \item access to field x in current class via This.x
@@ -73,7 +72,8 @@
 	| Std xname  --{* intermediate standard exception, see Eval.thy *}
 
 datatype error
-       = AccessViolation --{* Access to a member that isn't permitted *}
+       =  AccessViolation  --{* Access to a member that isn't permitted *}
+        | CrossMethodJump  --{* Method exits with a break or continue *}
 
 datatype abrupt       --{* abrupt completion *} 
         = Xcpt xcpt   --{* exception *}
@@ -206,7 +206,7 @@
 	| Comp  stmt stmt       ("_;; _"                  [      66,65]65)
 	| If_   expr stmt stmt  ("If'(_') _ Else _"       [   80,79,79]70)
 	| Loop  label expr stmt ("_\<bullet> While'(_') _"        [   99,80,79]70)
-        | Do jump               --{* break, continue, return *}
+        | Jmp jump              --{* break, continue, return *}
 	| Throw expr
         | TryC  stmt qtname vname stmt ("Try _ Catch'(_ _') _"  [79,99,80,79]70)
              --{* @{term "Try c1 Catch(C vn) c2"} *} 
@@ -258,8 +258,8 @@
  "this"       == "Acc (LVar This)"
  "!!v"        == "Acc (LVar (EName (VNam v)))"
  "v:==e"      == "Expr (Ass (LVar (EName (VNam  v))) e)"
- "Return e"   == "Expr (Ass (LVar (EName Res)) e);; Do Ret" 
-                  --{* \tt Res := e;; Do Ret *}
+ "Return e"   == "Expr (Ass (LVar (EName Res)) e);; Jmp Ret" 
+                  --{* \tt Res := e;; Jmp Ret *}
  "StatRef rt" == "Cast (RefT rt) (Lit Null)"
   
 constdefs
@@ -273,8 +273,39 @@
 
 declare is_stmt_rews [simp]
 
+text {*
+  Here is some syntactic stuff to handle the injections of statements,
+  expressions, variables and expression lists into general terms.
+*}
+
+syntax 
+  expr_inj_term:: "expr \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>e" 1000)
+  stmt_inj_term:: "stmt \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>s" 1000)
+  var_inj_term::  "var \<Rightarrow> term"  ("\<langle>_\<rangle>\<^sub>v" 1000)
+  lst_inj_term:: "expr list \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>l" 1000)
+
+translations 
+  "\<langle>e\<rangle>\<^sub>e" \<rightharpoonup> "In1l e"
+  "\<langle>c\<rangle>\<^sub>s" \<rightharpoonup> "In1r c"
+  "\<langle>v\<rangle>\<^sub>v" \<rightharpoonup> "In2 v"
+  "\<langle>es\<rangle>\<^sub>l" \<rightharpoonup> "In3 es"
+
+text {* It seems to be more elegant to have an overloaded injection like the
+following.
+*}
+
 axclass inj_term < "type"
-consts inj_term:: "'a::inj_term \<Rightarrow> term" ("\<langle>_\<rangle>" 83)
+consts inj_term:: "'a::inj_term \<Rightarrow> term" ("\<langle>_\<rangle>" 1000)
+
+text {* How this overloaded injections work can be seen in the theory 
+@{text DefiniteAssignment}. Other big inductive relations on
+terms defined in theories @{text WellType}, @{text Eval}, @{text Evaln} and
+@{text AxSem} don't follow this convention right now, but introduce subtle 
+syntactic sugar in the relations themselves to make a distinction on 
+expressions, statements and so on. So unfortunately you will encounter a 
+mixture of dealing with these injections. The translations above are used
+as bridge between the different conventions.  
+*}
 
 instance stmt::inj_term ..
 
@@ -284,6 +315,9 @@
 lemma stmt_inj_term_simp: "\<langle>c::stmt\<rangle> = In1r c"
 by (simp add: stmt_inj_term_def)
 
+lemma  stmt_inj_term [iff]: "\<langle>x::stmt\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
+  by (simp add: stmt_inj_term_simp)
+
 instance expr::inj_term ..
 
 defs (overloaded)
@@ -292,6 +326,9 @@
 lemma expr_inj_term_simp: "\<langle>e::expr\<rangle> = In1l e"
 by (simp add: expr_inj_term_def)
 
+lemma expr_inj_term [iff]: "\<langle>x::expr\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
+  by (simp add: expr_inj_term_simp)
+
 instance var::inj_term ..
 
 defs (overloaded)
@@ -300,6 +337,9 @@
 lemma var_inj_term_simp: "\<langle>v::var\<rangle> = In2 v"
 by (simp add: var_inj_term_def)
 
+lemma var_inj_term [iff]: "\<langle>x::var\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
+  by (simp add: var_inj_term_simp)
+
 instance "list":: (type) inj_term ..
 
 defs (overloaded)
@@ -308,4 +348,93 @@
 lemma expr_list_inj_term_simp: "\<langle>es::expr list\<rangle> = In3 es"
 by (simp add: expr_list_inj_term_def)
 
+lemma expr_list_inj_term [iff]: "\<langle>x::expr list\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
+  by (simp add: expr_list_inj_term_simp)
+
+lemmas inj_term_simps = stmt_inj_term_simp expr_inj_term_simp var_inj_term_simp
+                        expr_list_inj_term_simp
+lemmas inj_term_sym_simps = stmt_inj_term_simp [THEN sym] 
+                            expr_inj_term_simp [THEN sym]
+                            var_inj_term_simp [THEN sym]
+                            expr_list_inj_term_simp [THEN sym]
+
+lemma stmt_expr_inj_term [iff]: "\<langle>t::stmt\<rangle> \<noteq> \<langle>w::expr\<rangle>"
+  by (simp add: inj_term_simps)
+lemma expr_stmt_inj_term [iff]: "\<langle>t::expr\<rangle> \<noteq> \<langle>w::stmt\<rangle>"
+  by (simp add: inj_term_simps)
+lemma stmt_var_inj_term [iff]: "\<langle>t::stmt\<rangle> \<noteq> \<langle>w::var\<rangle>"
+  by (simp add: inj_term_simps)
+lemma var_stmt_inj_term [iff]: "\<langle>t::var\<rangle> \<noteq> \<langle>w::stmt\<rangle>"
+  by (simp add: inj_term_simps)
+lemma stmt_elist_inj_term [iff]: "\<langle>t::stmt\<rangle> \<noteq> \<langle>w::expr list\<rangle>"
+  by (simp add: inj_term_simps)
+lemma elist_stmt_inj_term [iff]: "\<langle>t::expr list\<rangle> \<noteq> \<langle>w::stmt\<rangle>"
+  by (simp add: inj_term_simps)
+lemma expr_var_inj_term [iff]: "\<langle>t::expr\<rangle> \<noteq> \<langle>w::var\<rangle>"
+  by (simp add: inj_term_simps)
+lemma var_expr_inj_term [iff]: "\<langle>t::var\<rangle> \<noteq> \<langle>w::expr\<rangle>"
+  by (simp add: inj_term_simps)
+lemma expr_elist_inj_term [iff]: "\<langle>t::expr\<rangle> \<noteq> \<langle>w::expr list\<rangle>"
+  by (simp add: inj_term_simps)
+lemma elist_expr_inj_term [iff]: "\<langle>t::expr list\<rangle> \<noteq> \<langle>w::expr\<rangle>"
+  by (simp add: inj_term_simps)
+lemma var_elist_inj_term [iff]: "\<langle>t::var\<rangle> \<noteq> \<langle>w::expr list\<rangle>"
+  by (simp add: inj_term_simps)
+lemma elist_var_inj_term [iff]: "\<langle>t::expr list\<rangle> \<noteq> \<langle>w::var\<rangle>"
+  by (simp add: inj_term_simps)
+
+section {* Evaluation of unary operations *}
+consts eval_unop :: "unop \<Rightarrow> val \<Rightarrow> val"
+primrec
+"eval_unop UPlus   v = Intg (the_Intg v)"
+"eval_unop UMinus  v = Intg (- (the_Intg v))"
+"eval_unop UBitNot v = Intg 42"                -- "FIXME: Not yet implemented"
+"eval_unop UNot    v = Bool (\<not> the_Bool v)"
+
+section {* Evaluation of binary operations *}
+consts eval_binop :: "binop \<Rightarrow> val \<Rightarrow> val \<Rightarrow> val"
+primrec
+"eval_binop Mul     v1 v2 = Intg ((the_Intg v1) * (the_Intg v2))" 
+"eval_binop Div     v1 v2 = Intg ((the_Intg v1) div (the_Intg v2))"
+"eval_binop Mod     v1 v2 = Intg ((the_Intg v1) mod (the_Intg v2))"
+"eval_binop Plus    v1 v2 = Intg ((the_Intg v1) + (the_Intg v2))"
+"eval_binop Minus   v1 v2 = Intg ((the_Intg v1) - (the_Intg v2))"
+
+-- "Be aware of the explicit coercion of the shift distance to nat"
+"eval_binop LShift  v1 v2 = Intg ((the_Intg v1) *   (2^(nat (the_Intg v2))))"
+"eval_binop RShift  v1 v2 = Intg ((the_Intg v1) div (2^(nat (the_Intg v2))))"
+"eval_binop RShiftU v1 v2 = Intg 42" --"FIXME: Not yet implemented"
+
+"eval_binop Less    v1 v2 = Bool ((the_Intg v1) < (the_Intg v2))" 
+"eval_binop Le      v1 v2 = Bool ((the_Intg v1) \<le> (the_Intg v2))"
+"eval_binop Greater v1 v2 = Bool ((the_Intg v2) < (the_Intg v1))"
+"eval_binop Ge      v1 v2 = Bool ((the_Intg v2) \<le> (the_Intg v1))"
+
+"eval_binop Eq      v1 v2 = Bool (v1=v2)"
+"eval_binop Neq     v1 v2 = Bool (v1\<noteq>v2)"
+"eval_binop BitAnd  v1 v2 = Intg 42" -- "FIXME: Not yet implemented"
+"eval_binop And     v1 v2 = Bool ((the_Bool v1) \<and> (the_Bool v2))"
+"eval_binop BitXor  v1 v2 = Intg 42" -- "FIXME: Not yet implemented"
+"eval_binop Xor     v1 v2 = Bool ((the_Bool v1) \<noteq> (the_Bool v2))"
+"eval_binop BitOr   v1 v2 = Intg 42" -- "FIXME: Not yet implemented"
+"eval_binop Or      v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
+"eval_binop CondAnd v1 v2 = Bool ((the_Bool v1) \<and> (the_Bool v2))"
+"eval_binop CondOr  v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
+
+constdefs need_second_arg :: "binop \<Rightarrow> val \<Rightarrow> bool"
+"need_second_arg binop v1 \<equiv> \<not> ((binop=CondAnd \<and>  \<not> the_Bool v1) \<or>
+                               (binop=CondOr  \<and> the_Bool v1))"
+text {* @{term CondAnd} and @{term CondOr} only evalulate the second argument
+ if the value isn't already determined by the first argument*}
+
+lemma need_second_arg_CondAnd [simp]: "need_second_arg CondAnd (Bool b) = b" 
+by (simp add: need_second_arg_def)
+
+lemma need_second_arg_CondOr [simp]: "need_second_arg CondOr (Bool b) = (\<not> b)" 
+by (simp add: need_second_arg_def)
+
+lemma need_second_arg_strict[simp]: 
+ "\<lbrakk>binop\<noteq>CondAnd; binop\<noteq>CondOr\<rbrakk> \<Longrightarrow> need_second_arg binop b"
+by (cases binop) 
+   (simp_all add: need_second_arg_def)
 end
\ No newline at end of file
--- a/src/HOL/Bali/Trans.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Trans.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Bali/Trans.thy
     ID:         $Id$
-    Author:     David von Oheimb
+    Author:     David von Oheimb and Norbert Schirmer
     License:    GPL (GNU GENERAL PUBLIC LICENSE)
 
 Operational transition (small-step) semantics of the 
@@ -323,7 +323,7 @@
 Loop: "G\<turnstile>(\<langle>l\<bullet> While(e) c\<rangle>,Norm s) 
         \<mapsto>1 (\<langle>If(e) (Cont l\<bullet>c;; l\<bullet> While(e) c) Else Skip\<rangle>,Norm s)"
 
-Do: "G\<turnstile>(\<langle>Do j\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,(Some (Jump j), s))"
+Jmp: "G\<turnstile>(\<langle>Jmp j\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,(Some (Jump j), s))"
 
 ThrowE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
          \<Longrightarrow>
--- a/src/HOL/Bali/Type.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Type.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -16,21 +16,21 @@
 \end{itemize}
 *}
 
-datatype prim_ty       	(* primitive type, cf. 4.2 *)
-	= Void    	(* 'result type' of void methods *)
+datatype prim_ty       	--{* primitive type, cf. 4.2 *}
+	= Void    	--{* result type of void methods *}
 	| Boolean
 	| Integer
 
 
-datatype ref_ty		(* reference type, cf. 4.3 *)
-	= NullT		(* null type, cf. 4.1 *)
-	| IfaceT qtname	(* interface type *)
-	| ClassT qtname	(* class type *)
-	| ArrayT ty	(* array type *)
+datatype ref_ty		--{* reference type, cf. 4.3 *}
+	= NullT		--{* null type, cf. 4.1 *}
+	| IfaceT qtname	--{* interface type *}
+	| ClassT qtname	--{* class type *}
+	| ArrayT ty	--{* array type *}
 
-and ty	    		(* any type, cf. 4.1 *)
-	= PrimT prim_ty	(* primitive type *)
-	| RefT  ref_ty	(* reference type *)
+and ty	    		--{* any type, cf. 4.1 *}
+	= PrimT prim_ty	--{* primitive type *}
+	| RefT  ref_ty	--{* reference type *}
 
 translations
   "prim_ty" <= (type) "Type.prim_ty"
--- a/src/HOL/Bali/TypeRel.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/TypeRel.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -34,11 +34,11 @@
 (*subcls1, in Decl.thy*)                     (* direct subclass           *)
 (*subcls , by translation*)                  (*        subclass           *)
 (*subclseq, by translation*)                 (* subclass + identity       *)
-  implmt1   :: "prog \<Rightarrow> (qtname \<times> qtname) set" (* direct implementation *)
-  implmt    :: "prog \<Rightarrow> (qtname \<times> qtname) set" (*        implementation *)
-  widen     :: "prog \<Rightarrow> (ty    \<times> ty   ) set" (*        widening       *)
-  narrow    :: "prog \<Rightarrow> (ty    \<times> ty   ) set" (*        narrowing      *)
-  cast     :: "prog \<Rightarrow> (ty    \<times> ty   ) set"  (*        casting        *)
+  implmt1   :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{* direct implementation *}
+  implmt    :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{*        implementation *}
+  widen     :: "prog \<Rightarrow> (ty    \<times> ty   ) set" --{*        widening       *}
+  narrow    :: "prog \<Rightarrow> (ty    \<times> ty   ) set" --{*        narrowing      *}
+  cast     :: "prog \<Rightarrow> (ty    \<times> ty   ) set"  --{*        casting        *}
 
 syntax
 
@@ -73,7 +73,7 @@
 translations
 
 	"G\<turnstile>I \<prec>I1 J" == "(I,J) \<in> subint1 G"
-	"G\<turnstile>I \<preceq>I  J" == "(I,J) \<in>(subint1 G)^*" (* cf. 9.1.3 *)
+	"G\<turnstile>I \<preceq>I  J" == "(I,J) \<in>(subint1 G)^*" --{* cf. 9.1.3 *}
   	(* Defined in Decl.thy:
         "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" == "(C,D) \<in> subcls1 G"
 	"G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" 
@@ -359,7 +359,7 @@
 section "implementation relation"
 
 defs
-  (* direct implementation, cf. 8.1.3 *)
+  --{* direct implementation, cf. 8.1.3 *}
 implmt1_def:"implmt1 G\<equiv>{(C,I). C\<noteq>Object \<and> (\<exists>c\<in>class G C: I\<in>set (superIfs c))}"
 
 lemma implmt1D: "G\<turnstile>C\<leadsto>1I \<Longrightarrow> C\<noteq>Object \<and> (\<exists>c\<in>class G C: I\<in>set (superIfs c))"
@@ -368,7 +368,7 @@
 done
 
 
-inductive "implmt G" intros                    (* cf. 8.1.4 *)
+inductive "implmt G" intros                    --{* cf. 8.1.4 *}
 
   direct:         "G\<turnstile>C\<leadsto>1J     \<spacespace>\<spacespace>     \<Longrightarrow> G\<turnstile>C\<leadsto>J"
   subint:        "\<lbrakk>G\<turnstile>C\<leadsto>1I; G\<turnstile>I\<preceq>I J\<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
@@ -400,10 +400,11 @@
 
 section "widening relation"
 
-inductive "widen G" intros(*widening, viz. method invocation conversion, cf. 5.3
-			    i.e. kind of syntactic subtyping *)
-  refl:    "G\<turnstile>T\<preceq>T"(*identity conversion, cf. 5.1.1 *)
-  subint:  "G\<turnstile>I\<preceq>I J  \<Longrightarrow> G\<turnstile>Iface I\<preceq> Iface J"(*wid.ref.conv.,cf. 5.1.4 *)
+inductive "widen G" intros
+ --{*widening, viz. method invocation conversion, cf. 5.3
+			    i.e. kind of syntactic subtyping *}
+  refl:    "G\<turnstile>T\<preceq>T" --{*identity conversion, cf. 5.1.1 *}
+  subint:  "G\<turnstile>I\<preceq>I J  \<Longrightarrow> G\<turnstile>Iface I\<preceq> Iface J" --{*wid.ref.conv.,cf. 5.1.4 *}
   int_obj: "G\<turnstile>Iface I\<preceq> Class Object"
   subcls:  "G\<turnstile>C\<preceq>\<^sub>C D  \<Longrightarrow> G\<turnstile>Class C\<preceq> Class D"
   implmt:  "G\<turnstile>C\<leadsto>I   \<Longrightarrow> G\<turnstile>Class C\<preceq> Iface I"
@@ -428,6 +429,26 @@
 apply (ind_cases "G\<turnstile>S\<preceq>T")
 by auto
 
+text {* 
+   These widening lemmata hold in Bali but are to strong for ordinary
+   Java. They  would not work for real Java Integral Types, like short, 
+   long, int. These lemmata are just for documentation and are not used.
+ *}
+
+lemma widen_PrimT_strong: "G\<turnstile>PrimT x\<preceq>T \<Longrightarrow> T = PrimT x"
+by (ind_cases "G\<turnstile>S\<preceq>T") simp_all
+
+lemma widen_PrimT2_strong: "G\<turnstile>S\<preceq>PrimT x \<Longrightarrow> S = PrimT x"
+by (ind_cases "G\<turnstile>S\<preceq>T") simp_all
+
+text {* Specialized versions for booleans also would work for real Java *}
+
+lemma widen_Boolean: "G\<turnstile>PrimT Boolean\<preceq>T \<Longrightarrow> T = PrimT Boolean"
+by (ind_cases "G\<turnstile>S\<preceq>T") simp_all
+
+lemma widen_Boolean2: "G\<turnstile>S\<preceq>PrimT Boolean \<Longrightarrow> S = PrimT Boolean"
+by (ind_cases "G\<turnstile>S\<preceq>T") simp_all
+
 lemma widen_RefT: "G\<turnstile>RefT R\<preceq>T \<Longrightarrow> \<exists>t. T=RefT t"
 apply (ind_cases "G\<turnstile>S\<preceq>T")
 by auto
@@ -516,10 +537,7 @@
 lemma widen_Array_Class: "G\<turnstile>S.[] \<preceq> Class C \<Longrightarrow> C=Object"
 by (auto dest: widen_Array)
 
-(*
-qed_typerel "widen_NT2" "G\<turnstile>S\<preceq>NT \<Longrightarrow> S = NT"
- [prove_widen_lemma "G\<turnstile>S\<preceq>T \<Longrightarrow> T = NT \<longrightarrow> S = NT"]
-*)
+
 lemma widen_NT2: "G\<turnstile>S\<preceq>NT \<Longrightarrow> S = NT"
 apply (ind_cases "G\<turnstile>S\<preceq>T")
 by auto
@@ -634,18 +652,32 @@
 apply (ind_cases "G\<turnstile>S\<succ>T")
 by auto
 
+text {* 
+   These narrowing lemmata hold in Bali but are to strong for ordinary
+   Java. They  would not work for real Java Integral Types, like short, 
+   long, int. These lemmata are just for documentation and are not used.
+ *}
+lemma narrow_PrimT_strong: "G\<turnstile>PrimT pt\<succ>T \<Longrightarrow> T=PrimT pt"
+by (ind_cases "G\<turnstile>S\<succ>T") simp_all
+
+lemma narrow_PrimT2_strong: "G\<turnstile>S\<succ>PrimT pt \<Longrightarrow> S=PrimT pt"
+by (ind_cases "G\<turnstile>S\<succ>T") simp_all
+
+text {* Specialized versions for booleans also would work for real Java *}
+
+lemma narrow_Boolean: "G\<turnstile>PrimT Boolean\<succ>T \<Longrightarrow> T=PrimT Boolean"
+by (ind_cases "G\<turnstile>S\<succ>T") simp_all
+
+lemma narrow_Boolean2: "G\<turnstile>S\<succ>PrimT Boolean \<Longrightarrow> S=PrimT Boolean"
+by (ind_cases "G\<turnstile>S\<succ>T") simp_all
 
 section "casting relation"
 
-inductive "cast G" intros (* casting conversion, cf. 5.5 *)
+inductive "cast G" intros --{* casting conversion, cf. 5.5 *}
 
   widen:   "G\<turnstile>S\<preceq>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
   narrow:  "G\<turnstile>S\<succ>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
 
-(*
-lemma ??unknown??: "\<lbrakk>G\<turnstile>S\<preceq>T; G\<turnstile>S\<succ>T\<rbrakk> \<Longrightarrow> R"
- deferred *)
-
 (*unused*)
 lemma cast_RefT: "G\<turnstile>RefT R\<preceq>? T \<Longrightarrow> \<exists>t. T=RefT t"
 apply (ind_cases "G\<turnstile>S\<preceq>? T")
@@ -664,4 +696,36 @@
 apply (ind_cases "G\<turnstile>S\<preceq>? T")
 by (auto dest: widen_PrimT2 narrow_PrimT2)
 
+lemma cast_Boolean: 
+  assumes bool_cast: "G\<turnstile>PrimT Boolean\<preceq>? T" 
+  shows "T=PrimT Boolean"
+using bool_cast 
+proof (cases)
+  case widen 
+  hence "G\<turnstile>PrimT Boolean\<preceq> T"
+    by simp
+  thus ?thesis by  (rule widen_Boolean)
+next
+  case narrow
+  hence "G\<turnstile>PrimT Boolean\<succ>T"
+    by simp
+  thus ?thesis by (rule narrow_Boolean)
+qed
+
+lemma cast_Boolean2: 
+  assumes bool_cast: "G\<turnstile>S\<preceq>? PrimT Boolean" 
+  shows "S = PrimT Boolean"
+using bool_cast 
+proof (cases)
+  case widen 
+  hence "G\<turnstile>S\<preceq> PrimT Boolean"
+    by simp
+  thus ?thesis by  (rule widen_Boolean2)
+next
+  case narrow
+  hence "G\<turnstile>S\<succ>PrimT Boolean"
+    by simp
+  thus ?thesis by (rule narrow_Boolean2)
+qed
+
 end
--- a/src/HOL/Bali/TypeSafe.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/TypeSafe.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -5,7 +5,7 @@
 *)
 header {* The type soundness proof for Java *}
 
-theory TypeSafe = Eval + WellForm + Conform:
+theory TypeSafe = DefiniteAssignmentCorrect + Conform:
 
 section "error free"
  
@@ -24,8 +24,7 @@
   have "error_free (abrupt1,store1)"
   proof (induct)
     case Abrupt 
-    then show ?case
-      .
+    then show ?case .
   next
     case New
     then show ?case
@@ -101,27 +100,64 @@
  (\<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')))"      
 
+
+constdefs
   rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool"
           ("_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_"                               [71,71,71,71,71,71] 70)
   "G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T 
     \<equiv> case T of
-        Inl T  \<Rightarrow> if (\<exists>vf. t=In2 vf)
-                  then G,s\<turnstile>fst (the_In2 v)\<Colon>\<preceq>T \<and> s\<le>|snd (the_In2 v)\<preceq>T\<Colon>\<preceq>(G,L)
+        Inl T  \<Rightarrow> if (\<exists> var. t=In2 var)
+                  then (\<forall> n. (the_In2 t) = LVar n 
+                         \<longrightarrow> (fst (the_In2 v) = the (locals s n)) \<and>
+                             (locals s n \<noteq> None \<longrightarrow> G,s\<turnstile>fst (the_In2 v)\<Colon>\<preceq>T)) \<and>
+                      (\<not> (\<exists> n. the_In2 t=LVar n) \<longrightarrow> (G,s\<turnstile>fst (the_In2 v)\<Colon>\<preceq>T))\<and>
+                      (s\<le>|snd (the_In2 v)\<preceq>T\<Colon>\<preceq>(G,L))
                   else G,s\<turnstile>the_In1 v\<Colon>\<preceq>T
       | Inr Ts \<Rightarrow> list_all2 (conf G s) (the_In3 v) Ts"
 
+text {*
+ With @{term rconf} we describe the conformance of the result value of a term.
+ This definition gets rather complicated because of the relations between the
+ injections of the different terms, types and values. The main case distinction
+ is between single values and value lists. In case of value lists, every 
+ value has to conform to its type. For single values we have to do a further
+ case distinction, between values of variables @{term "\<exists>var. t=In2 var" } and
+ ordinary values. Values of variables are modelled as pairs consisting of the
+ current value and an update function which will perform an assignment to the
+ variable. This stems form the decision, that we only have one evaluation rule
+ for each kind of variable. The decision if we read or write to the 
+ variable is made by syntactic enclosing rules. So conformance of 
+ variable-values must ensure that both the current value and an update will 
+ conform to the type. With the introduction of definite assignment of local
+ variables we have to do another case distinction. For the notion of conformance
+ local variables are allowed to be @{term None}, since the definedness is not 
+ ensured by conformance but by definite assignment. Field and array variables 
+ must contain a value. 
+*}
+ 
+
+
 lemma rconf_In1 [simp]: 
  "G,L,s\<turnstile>In1 ec\<succ>In1 v \<Colon>\<preceq>Inl T  =  G,s\<turnstile>v\<Colon>\<preceq>T"
 apply (unfold rconf_def)
 apply (simp (no_asm))
 done
 
-lemma rconf_In2 [simp]: 
- "G,L,s\<turnstile>In2 va\<succ>In2 vf\<Colon>\<preceq>Inl T  = (G,s\<turnstile>fst vf\<Colon>\<preceq>T \<and> s\<le>|snd vf\<preceq>T\<Colon>\<preceq>(G,L))"
+lemma rconf_In2_no_LVar [simp]: 
+ "\<forall> n. va\<noteq>LVar n \<Longrightarrow> 
+   G,L,s\<turnstile>In2 va\<succ>In2 vf\<Colon>\<preceq>Inl T  = (G,s\<turnstile>fst vf\<Colon>\<preceq>T \<and> s\<le>|snd vf\<preceq>T\<Colon>\<preceq>(G,L))"
 apply (unfold rconf_def)
-apply (simp (no_asm))
+apply auto
 done
 
+lemma rconf_In2_LVar [simp]: 
+ "va=LVar n \<Longrightarrow> 
+   G,L,s\<turnstile>In2 va\<succ>In2 vf\<Colon>\<preceq>Inl T  
+    = ((fst vf = the (locals s n)) \<and>
+       (locals s n \<noteq> None \<longrightarrow> G,s\<turnstile>fst vf\<Colon>\<preceq>T) \<and> s\<le>|snd vf\<preceq>T\<Colon>\<preceq>(G,L))"
+apply (unfold rconf_def)
+by simp
+
 lemma rconf_In3 [simp]: 
  "G,L,s\<turnstile>In3 es\<succ>In3 vs\<Colon>\<preceq>Inr Ts = list_all2 (\<lambda>v T. G,s\<turnstile>v\<Colon>\<preceq>T) vs Ts"
 apply (unfold rconf_def)
@@ -222,13 +258,6 @@
 apply (auto intro!: oconf_init_obj_lemma unique_fields)
 done
 
-(*
-lemma obj_split: "P obj = (\<forall> oi vs. obj = \<lparr>tag=oi,values=vs\<rparr> \<longrightarrow> ?P \<lparr>tag=oi,values=vs\<rparr>)"
-apply auto
-apply (case_tac "obj")
-apply auto
-*)
-
 lemma conforms_newG: "\<lbrakk>globs s oref = None; (x, s)\<Colon>\<preceq>(G,L);   
   wf_prog G; case oref of Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=vs\<rparr>)  
                         | Stat C \<Rightarrow> is_class G C\<rbrakk> \<Longrightarrow>  
@@ -272,9 +301,14 @@
 done
 
 lemma sxalloc_type_sound: 
- "\<And>s1 s2. \<lbrakk>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; wf_prog G\<rbrakk> \<Longrightarrow> case fst s1 of  
-  None \<Rightarrow> s2 = s1 | Some x \<Rightarrow>  
-  (\<exists>a. fst s2 = Some(Xcpt (Loc a)) \<and> (\<forall>L. s1\<Colon>\<preceq>(G,L) \<longrightarrow> s2\<Colon>\<preceq>(G,L)))"
+ "\<And>s1 s2. \<lbrakk>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; wf_prog G\<rbrakk> \<Longrightarrow> 
+  case fst s1 of  
+    None \<Rightarrow> s2 = s1 
+  | Some abr \<Rightarrow> (case abr of
+                   Xcpt x \<Rightarrow> (\<exists>a. fst s2 = Some(Xcpt (Loc a)) \<and> 
+                                  (\<forall>L. s1\<Colon>\<preceq>(G,L) \<longrightarrow> s2\<Colon>\<preceq>(G,L)))
+                 | Jump j \<Rightarrow> s2 = s1
+                 | Error e \<Rightarrow> s2 = s1)"
 apply (simp (no_asm_simp) only: split_tupled_all)
 apply (erule sxalloc.induct)
 apply   auto
@@ -555,7 +589,43 @@
          eq_declC_sm_dm eq_mheads static
     show ?thesis by auto
   qed
-qed   
+qed
+
+corollary DynT_mheadsE [consumes 7]: 
+--{* Same as @{text DynT_mheadsD} but better suited for application in 
+typesafety proof   *}
+ assumes invC_compatible: "G\<turnstile>mode\<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 mheads: "(statDeclT,sm) \<in> mheads G C statT sig"
+     and mode: "mode=invmode sm e" 
+     and invC: "invC = invocation_class mode s a' statT"
+     and declC: "declC =invocation_declclass G mode s a' statT sig"
+     and dm: "\<And> dm. \<lbrakk>methd G declC sig = Some dm; 
+                      dynlookup G statT invC sig = Some dm; 
+                      G\<turnstile>resTy (mthd dm)\<preceq>resTy sm; 
+                      wf_mdecl G declC (sig, mthd dm);
+                      declC = declclass dm;
+                      is_static dm = is_static sm;  
+                      is_class G invC; is_class G declC; G\<turnstile>invC\<preceq>\<^sub>C declC;  
+                      (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>
+                             statDeclT = ClassT (declclass dm))\<rbrakk> \<Longrightarrow> P"
+   shows "P"
+proof -
+    from invC_compatible mode have "G\<turnstile>invmode sm e\<rightarrow>invC\<preceq>statT" by simp 
+    moreover note wf wt_e mheads
+    moreover from invC mode 
+    have "invC = invocation_class (invmode sm e) s a' statT" by simp
+    moreover from declC mode 
+    have "declC =invocation_declclass G (invmode sm e) s a' statT sig" by simp
+    ultimately show ?thesis
+      by (rule DynT_mheadsD [THEN exE,rule_format])
+         (elim conjE,rule dm)
+qed
+   
 
 lemma DynT_conf: "\<lbrakk>G\<turnstile>invocation_class mode s a' statT \<preceq>\<^sub>C declC; wf_prog G;
  isrtype G (statT);
@@ -605,9 +675,10 @@
 done
 
 lemma Fin_lemma: 
-"\<lbrakk>G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> (x2,s2); wf_prog G; (Some a, s1)\<Colon>\<preceq>(G, L); (x2,s2)\<Colon>\<preceq>(G, L)\<rbrakk> 
+"\<lbrakk>G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> (x2,s2); wf_prog G; (Some a, s1)\<Colon>\<preceq>(G, L); (x2,s2)\<Colon>\<preceq>(G, L);
+  dom (locals s1) \<subseteq> dom (locals s2)\<rbrakk> 
 \<Longrightarrow>  (abrupt_if True (Some a) x2, s2)\<Colon>\<preceq>(G, L)"
-apply (force elim: eval_gext' conforms_xgext split add: split_abrupt_if)
+apply (auto elim: eval_gext' conforms_xgext split add: split_abrupt_if)
 done
 
 lemma FVar_lemma1: 
@@ -634,8 +705,6 @@
 apply clarsimp
 apply (drule conf_RefTD, clarsimp, rule var_tys_Some_eq [THEN iffD2])
 apply (auto dest!: widen_Array split add: obj_tag.split)
-apply (rotate_tac -1) (* to front: tag obja = CInst pid_field_type to enable
-                         conditional rewrite *)
 apply (rule fields_table_SomeI)
 apply (auto elim!: fields_mono subcls_is_class2)
 done
@@ -750,17 +819,17 @@
          intro: var_tys_Some_eq [THEN iffD2] conforms_upd_gobj)
 done
 
+
 section "Call"
 
 lemma conforms_init_lvars_lemma: "\<lbrakk>wf_prog G;  
-  wf_mhead G P sig mh; 
-  Ball (set lvars) (split (\<lambda>vn. is_type G)); 
+  wf_mhead G P sig mh;
   list_all2 (conf G s) pvs pTsa; G\<turnstile>pTsa[\<preceq>](parTs sig)\<rbrakk> \<Longrightarrow>  
-  G,s\<turnstile>init_vals (table_of lvars)(pars mh[\<mapsto>]pvs)
-      [\<Colon>\<preceq>]table_of lvars(pars mh[\<mapsto>]parTs sig)"
+  G,s\<turnstile>empty (pars mh[\<mapsto>]pvs)
+      [\<sim>\<Colon>\<preceq>]table_of lvars(pars mh[\<mapsto>]parTs sig)"
 apply (unfold wf_mhead_def)
 apply clarify
-apply (erule (2) Ball_set_table [THEN lconf_init_vals, THEN lconf_ext_list])
+apply (erule (1) wlconf_empty_vals [THEN wlconf_ext_list])
 apply (drule wf_ws_prog)
 apply (erule (2) conf_list_widen)
 done
@@ -771,9 +840,15 @@
    =
   (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit . l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
 apply (unfold lconf_def)
-apply safe
-apply (case_tac "n")
-apply (force split add: lname.split)+
+apply (auto split add: lname.splits)
+done
+
+lemma wlconf_map_lname [simp]: 
+  "G,s\<turnstile>(lname_case l1 l2)[\<sim>\<Colon>\<preceq>](lname_case L1 L2)
+   =
+  (G,s\<turnstile>l1[\<sim>\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit . l2)[\<sim>\<Colon>\<preceq>](\<lambda>x::unit. L2))"
+apply (unfold wlconf_def)
+apply (auto split add: lname.splits)
 done
 
 lemma lconf_map_ename [simp]:
@@ -781,10 +856,18 @@
    =
   (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit. l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
 apply (unfold lconf_def)
-apply safe
-apply (force split add: ename.split)+
+apply (auto split add: ename.splits)
 done
 
+lemma wlconf_map_ename [simp]:
+  "G,s\<turnstile>(ename_case l1 l2)[\<sim>\<Colon>\<preceq>](ename_case L1 L2)
+   =
+  (G,s\<turnstile>l1[\<sim>\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit. l2)[\<sim>\<Colon>\<preceq>](\<lambda>x::unit. L2))"
+apply (unfold wlconf_def)
+apply (auto split add: ename.splits)
+done
+
+
 
 lemma defval_conf1 [rule_format (no_asm), elim]: 
   "is_type G T \<longrightarrow> (\<exists>v\<in>Some (default_val T): G,s\<turnstile>v\<Colon>\<preceq>T)"
@@ -793,6 +876,9 @@
 apply (auto intro: prim_ty.induct)
 done
 
+lemma np_no_jump: "x\<noteq>Some (Jump j) \<Longrightarrow> (np a') x \<noteq> Some (Jump j)"
+by (auto simp add: abrupt_if_def)
+
 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]
@@ -814,8 +900,7 @@
     \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object);
   invC  = invocation_class (invmode (mhd sm) e) s a' statT;
   declC = invocation_declclass G (invmode (mhd sm) e) s a' statT sig;
-  Ball (set (lcls (mbody (mthd dm)))) 
-       (split (\<lambda>vn. is_type G)) 
+  x\<noteq>Some (Jump Ret) 
  \<rbrakk> \<Longrightarrow>  
   init_lvars G declC sig (invmode (mhd sm) e) a'  
   pvs (x,s)\<Colon>\<preceq>(G,\<lambda> k. 
@@ -833,14 +918,20 @@
 apply (drule  (4) DynT_conf)
 apply clarsimp
 (* apply intro *)
-apply (drule (4) conforms_init_lvars_lemma)
+apply (drule (3) conforms_init_lvars_lemma 
+                 [where ?lvars="(lcls (mbody (mthd dm)))"])
 apply (case_tac "dm",simp)
 apply (rule conjI)
-apply (unfold lconf_def, clarify)
-apply (rule defval_conf1)
-apply (clarsimp simp add: wf_mhead_def is_acc_type_def)
-apply (case_tac "is_static sm")
-apply simp_all
+apply (unfold wlconf_def, clarify)
+apply   (clarsimp simp add: wf_mhead_def is_acc_type_def)
+apply   (case_tac "is_static sm")
+apply     simp
+apply     simp
+
+apply   simp
+apply   (case_tac "is_static sm")
+apply     simp
+apply     (simp add: np_no_jump)
 done
 declare split_paired_All [simp] split_paired_Ex [simp] 
 declare split_if     [split] split_if_asm     [split] 
@@ -851,7 +942,6 @@
 *}
 
 
-
 subsection "accessibility"
 
 text {* 
@@ -859,7 +949,6 @@
 *} (* dummy text command to break paragraph for latex;
               large paragraphs exhaust memory of debian pdflatex *)
 
-(* #### stat raus und gleich is_static f schreiben *) 
 theorem dynamic_field_access_ok:
   assumes wf: "wf_prog G" and
     not_Null: "\<not> stat \<longrightarrow> a\<noteq>Null" and
@@ -871,8 +960,8 @@
         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)"
-  shows "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f) \<and> 
-     G\<turnstile>Field fn f in dynC dyn_accessible_from accC"
+  shows "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 stat have static: "(is_static f)" by simp
@@ -914,112 +1003,6 @@
     by blast
 qed
 
-(*
-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
@@ -1199,26 +1182,581 @@
   qed
 qed
 
+lemma map_upds_eq_length_append_simp:
+  "\<And> tab qs. length ps = length qs \<Longrightarrow>  tab(ps[\<mapsto>]qs@zs) = tab(ps[\<mapsto>]qs)"
+proof (induct ps) 
+  case Nil thus ?case by simp
+next
+  case (Cons p ps tab qs)
+  have "length (p#ps) = length qs" .
+  then obtain q qs' where qs: "qs=q#qs'" and eq_length: "length ps=length qs'"
+    by (cases qs) auto
+  from eq_length have "(tab(p\<mapsto>q))(ps[\<mapsto>]qs'@zs)=(tab(p\<mapsto>q))(ps[\<mapsto>]qs')"
+    by (rule Cons.hyps)
+  with qs show ?case 
+    by simp
+qed
+
+lemma map_upds_upd_eq_length_simp:
+  "\<And> tab qs x y. length ps = length qs 
+                  \<Longrightarrow> tab(ps[\<mapsto>]qs)(x\<mapsto>y) = tab(ps@[x][\<mapsto>]qs@[y])"
+proof (induct "ps")
+  case Nil thus ?case by simp
+next
+  case (Cons p ps tab qs x y)
+  have "length (p#ps) = length qs" .
+  then obtain q qs' where qs: "qs=q#qs'" and eq_length: "length ps=length qs'"
+    by (cases qs) auto
+  from eq_length 
+  have "(tab(p\<mapsto>q))(ps[\<mapsto>]qs')(x\<mapsto>y) = (tab(p\<mapsto>q))(ps@[x][\<mapsto>]qs'@[y])"
+    by (rule Cons.hyps)
+  with qs show ?case
+    by simp
+qed
+
+
+lemma map_upd_cong: "tab=tab'\<Longrightarrow> tab(x\<mapsto>y) = tab'(x\<mapsto>y)"
+by simp
+
+lemma map_upd_cong_ext: "tab z=tab' z\<Longrightarrow> (tab(x\<mapsto>y)) z = (tab'(x\<mapsto>y)) z"
+by (simp add: fun_upd_def)
+
+lemma map_upds_cong: "tab=tab'\<Longrightarrow> tab(xs[\<mapsto>]ys) = tab'(xs[\<mapsto>]ys)"
+by (cases xs) simp+
+
+lemma map_upds_cong_ext: 
+ "\<And> tab tab' ys. tab z=tab' z \<Longrightarrow> (tab(xs[\<mapsto>]ys)) z = (tab'(xs[\<mapsto>]ys)) z"
+proof (induct xs)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs tab tab' ys)
+  have "(tab(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) z = (tab'(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) z"
+    by (rule Cons.hyps) (rule map_upd_cong_ext)
+  thus ?case
+    by simp
+qed
+   
+lemma map_upd_override: "(tab(x\<mapsto>y)) x = (tab'(x\<mapsto>y)) x"
+  by simp
+
+
+lemma map_upds_override_cong:
+"\<And> tab tab' zs. x\<in> set ws \<Longrightarrow> 
+ (tab(ws[\<mapsto>]zs)) x = (tab'(ws[\<mapsto>]zs)) x"
+proof (induct ws)
+  case Nil thus ?case by simp
+next
+  case (Cons w ws tab tab' zs)
+  have x: "x \<in> set (w#ws)" .
+  show ?case
+  proof (cases "x=w")
+    case True thus ?thesis
+      by simp (rule map_upds_cong_ext, rule map_upd_override)
+  next
+    case False
+    with x have "x \<in> set ws"
+      by simp
+    with Cons.hyps show ?thesis
+      by simp
+  qed
+qed
+
+lemma map_upds_in_suffix: assumes x: "x \<in> set xs" 
+ shows "\<And> tab qs. (tab(ps @ xs[\<mapsto>]qs)) x = (tab(xs[\<mapsto>]drop (length ps) qs)) x"
+proof (induct ps)
+  case Nil thus ?case by simp
+next
+  case (Cons p ps tab qs)
+  have "(tab(p\<mapsto>hd qs)(ps @ xs[\<mapsto>](tl qs))) x
+          =(tab(p\<mapsto>hd qs)(xs[\<mapsto>]drop (length ps) (tl qs))) x"
+    by (rule Cons.hyps)
+  moreover
+  have "drop (Suc (length ps)) qs=drop (length ps) (tl qs)"
+    by (cases qs) simp+
+  ultimately show ?case
+    by simp (rule map_upds_override_cong)
+qed
+ 
+
+lemma map_upds_eq_length_suffix: "\<And> tab qs. 
+        length ps = length qs \<Longrightarrow> tab(ps@xs[\<mapsto>]qs) = tab(ps[\<mapsto>]qs)(xs[\<mapsto>][])"
+proof (induct ps)
+  case Nil thus ?case by simp
+next
+  case (Cons p ps tab qs)
+  then obtain q qs' where qs: "qs=q#qs'" and eq_length: "length ps=length qs'"
+    by (cases qs) auto
+  from eq_length
+  have "tab(p\<mapsto>q)(ps @ xs[\<mapsto>]qs') = tab(p\<mapsto>q)(ps[\<mapsto>]qs')(xs[\<mapsto>][])"
+    by (rule Cons.hyps)
+  with qs show ?case 
+    by simp
+qed
+  
+  
+lemma map_upds_upds_eq_length_prefix_simp:
+  "\<And> tab qs. length ps = length qs
+              \<Longrightarrow> tab(ps[\<mapsto>]qs)(xs[\<mapsto>]ys) = tab(ps@xs[\<mapsto>]qs@ys)"
+proof (induct ps)
+  case Nil thus ?case by simp
+next
+  case (Cons p ps tab qs)
+  then obtain q qs' where qs: "qs=q#qs'" and eq_length: "length ps=length qs'"
+    by (cases qs) auto
+  from eq_length 
+  have "tab(p\<mapsto>q)(ps[\<mapsto>]qs')(xs[\<mapsto>]ys) = tab(p\<mapsto>q)(ps @ xs[\<mapsto>](qs' @ ys))"
+    by (rule Cons.hyps)
+  with qs 
+  show ?case by simp
+qed
+
+lemma map_upd_cut_irrelevant:
+"\<lbrakk>(tab(x\<mapsto>y)) vn = Some el; (tab'(x\<mapsto>y)) vn = None\<rbrakk>
+    \<Longrightarrow> tab vn = Some el"
+by (cases "tab' vn = None") (simp add: fun_upd_def)+
+
+lemma map_upd_Some_expand:
+"\<lbrakk>tab vn = Some z\<rbrakk>
+    \<Longrightarrow> \<exists> z. (tab(x\<mapsto>y)) vn = Some z"
+by (simp add: fun_upd_def)
+
+lemma map_upds_Some_expand:
+"\<And> tab ys z. \<lbrakk>tab vn = Some z\<rbrakk>
+    \<Longrightarrow> \<exists> z. (tab(xs[\<mapsto>]ys)) vn = Some z"
+proof (induct xs)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs tab ys z)
+  have "tab vn = Some z" .
+  then obtain z' where "(tab(x\<mapsto>hd ys)) vn = Some z'" 
+    by (rule map_upd_Some_expand [of tab,elim_format]) blast
+  hence "\<exists> z. (tab (x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) vn = Some z"
+    by (rule Cons.hyps)
+  thus ?case
+    by simp
+qed
+
+
+lemma map_upd_Some_swap:
+ "(tab(r\<mapsto>w)(u\<mapsto>v)) vn = Some z \<Longrightarrow> \<exists> z. (tab(u\<mapsto>v)(r\<mapsto>w)) vn = Some z"
+by (simp add: fun_upd_def)
+
+lemma map_upd_None_swap:
+ "(tab(r\<mapsto>w)(u\<mapsto>v)) vn = None \<Longrightarrow> (tab(u\<mapsto>v)(r\<mapsto>w)) vn = None"
+by (simp add: fun_upd_def)
+
+
+lemma map_eq_upd_eq: "tab vn = tab' vn \<Longrightarrow> (tab(x\<mapsto>y)) vn = (tab'(x\<mapsto>y)) vn"
+by (simp add: fun_upd_def)
+
+lemma map_eq_upds_eq: 
+  "\<And> tab tab' ys. 
+   tab vn = tab' vn \<Longrightarrow> (tab(xs[\<mapsto>]ys)) vn = (tab'(xs[\<mapsto>]ys)) vn"
+proof (induct xs)
+  case Nil thus ?case by simp 
+next
+  case (Cons x xs tab tab' ys)
+  have "tab vn = tab' vn" .
+  hence "(tab(x\<mapsto>hd ys)) vn = (tab'(x\<mapsto>hd ys)) vn"
+    by (rule map_eq_upd_eq)
+  hence "(tab(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) vn = (tab'(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) vn"
+    by (rule Cons.hyps)
+  thus ?case 
+    by simp
+qed
+
+lemma map_upd_in_expansion_map_swap:
+ "\<lbrakk>(tab(x\<mapsto>y)) vn = Some z;tab vn \<noteq> Some z\<rbrakk> 
+                 \<Longrightarrow>  (tab'(x\<mapsto>y)) vn = Some z"
+by (simp add: fun_upd_def)
+
+lemma map_upds_in_expansion_map_swap:
+ "\<And>tab tab' ys z. \<lbrakk>(tab(xs[\<mapsto>]ys)) vn = Some z;tab vn \<noteq> Some z\<rbrakk> 
+                 \<Longrightarrow>  (tab'(xs[\<mapsto>]ys)) vn = Some z"
+proof (induct xs)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs tab tab' ys z)
+  from Cons.prems obtain 
+    some: "(tab(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) vn = Some z" and
+    tab_not_z: "tab vn \<noteq> Some z"
+    by simp
+  show ?case
+  proof (cases "(tab(x\<mapsto>hd ys)) vn \<noteq> Some z")
+    case True
+    with some have "(tab'(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) vn = Some z"
+      by (rule Cons.hyps)
+    thus ?thesis 
+      by simp
+  next
+    case False
+    hence tabx_z: "(tab(x\<mapsto>hd ys)) vn = Some z" by blast
+    moreover
+    from tabx_z tab_not_z
+    have "(tab'(x\<mapsto>hd ys)) vn = Some z" 
+      by (rule map_upd_in_expansion_map_swap)
+    ultimately
+    have "(tab(x\<mapsto>hd ys)) vn =(tab'(x\<mapsto>hd ys)) vn"
+      by simp
+    hence "(tab(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) vn = (tab'(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) vn"
+      by (rule map_eq_upds_eq)
+    with some 
+    show ?thesis 
+      by simp
+  qed
+qed
+   
+lemma map_upds_Some_swap: 
+ assumes r_u: "(tab(r\<mapsto>w)(u\<mapsto>v)(xs[\<mapsto>]ys)) vn = Some z"
+    shows "\<exists> z. (tab(u\<mapsto>v)(r\<mapsto>w)(xs[\<mapsto>]ys)) vn = Some z"
+proof (cases "(tab(r\<mapsto>w)(u\<mapsto>v)) vn = Some z")
+  case True
+  then obtain z' where "(tab(u\<mapsto>v)(r\<mapsto>w)) vn = Some z'"
+    by (rule map_upd_Some_swap [elim_format]) blast
+  thus "\<exists> z. (tab(u\<mapsto>v)(r\<mapsto>w)(xs[\<mapsto>]ys)) vn = Some z"
+    by (rule map_upds_Some_expand)
+next
+  case False
+  with r_u
+  have "(tab(u\<mapsto>v)(r\<mapsto>w)(xs[\<mapsto>]ys)) vn = Some z"
+    by (rule map_upds_in_expansion_map_swap)
+  thus ?thesis
+    by simp
+qed
+ 
+lemma map_upds_Some_insert:
+  assumes z: "(tab(xs[\<mapsto>]ys)) vn = Some z"
+    shows "\<exists> z. (tab(u\<mapsto>v)(xs[\<mapsto>]ys)) vn = Some z"
+proof (cases "\<exists> z. tab vn = Some z")
+  case True
+  then obtain z' where "tab vn = Some z'" by blast
+  then obtain z'' where "(tab(u\<mapsto>v)) vn = Some z''"
+    by (rule map_upd_Some_expand [elim_format]) blast
+  thus ?thesis
+    by (rule map_upds_Some_expand)
+next
+  case False
+  hence "tab vn \<noteq> Some z" by simp
+  with z
+  have "(tab(u\<mapsto>v)(xs[\<mapsto>]ys)) vn = Some z"
+    by (rule map_upds_in_expansion_map_swap)
+  thus ?thesis ..
+qed
+   
+lemma map_upds_None_cut:
+assumes expand_None: "(tab(xs[\<mapsto>]ys)) vn = None"
+  shows "tab vn = None"
+proof (cases "tab vn = None")
+  case True thus ?thesis by simp
+next
+  case False then obtain z where "tab vn = Some z" by blast
+  then obtain z' where "(tab(xs[\<mapsto>]ys)) vn = Some z'"
+    by (rule map_upds_Some_expand [where  ?tab="tab",elim_format]) blast
+  with expand_None show ?thesis
+    by simp
+qed
+    
+
+lemma map_upds_cut_irrelevant:
+"\<And> tab tab' ys. \<lbrakk>(tab(xs[\<mapsto>]ys)) vn = Some el; (tab'(xs[\<mapsto>]ys)) vn = None\<rbrakk>
+                  \<Longrightarrow> tab vn = Some el"
+proof  (induct "xs")
+  case Nil thus ?case by simp
+next
+  case (Cons x xs tab tab' ys)
+  from Cons.prems
+  have "(tab(x\<mapsto>hd ys)) vn = Some el"
+    by - (rule Cons.hyps,auto)
+  moreover from Cons.prems 
+  have "(tab'(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) vn = None" 
+    by simp
+  hence "(tab'(x\<mapsto>hd ys)) vn = None"
+    by (rule map_upds_None_cut)
+  ultimately show "tab vn = Some el" 
+    by (rule map_upd_cut_irrelevant)
+qed
+   
+lemma dom_vname_split:
+ "dom (lname_case (ename_case (tab(x\<mapsto>y)(xs[\<mapsto>]ys)) a) b)
+   = dom (lname_case (ename_case (tab(x\<mapsto>y)) a) b) \<union> 
+     dom (lname_case (ename_case (tab(xs[\<mapsto>]ys)) a) b)"
+  (is "?List x xs y ys = ?Hd x y \<union> ?Tl xs ys")
+proof 
+  show "?List x xs y ys \<subseteq> ?Hd x y \<union> ?Tl xs ys"
+  proof 
+    fix el 
+    assume el_in_list: "el \<in> ?List x xs y ys"
+    show "el \<in>  ?Hd x y \<union> ?Tl xs ys"
+    proof (cases el)
+      case This
+      with el_in_list show ?thesis by (simp add: dom_def)
+    next
+      case (EName en)
+      show ?thesis
+      proof (cases en)
+	case Res
+	with EName el_in_list show ?thesis by (simp add: dom_def)
+      next
+	case (VNam vn)
+	with EName el_in_list show ?thesis 
+	  by (auto simp add: dom_def dest: map_upds_cut_irrelevant)
+      qed
+    qed
+  qed
+next
+  show "?Hd x y \<union> ?Tl xs ys  \<subseteq> ?List x xs y ys" 
+  proof 
+    fix el 
+    assume  el_in_hd_tl: "el \<in>  ?Hd x y \<union> ?Tl xs ys"
+    show "el \<in> ?List x xs y ys"
+    proof (cases el)
+      case This
+      with el_in_hd_tl show ?thesis by (simp add: dom_def)
+    next
+      case (EName en)
+      show ?thesis
+      proof (cases en)
+	case Res
+	with EName el_in_hd_tl show ?thesis by (simp add: dom_def)
+      next
+	case (VNam vn)
+	with EName el_in_hd_tl show ?thesis 
+	  by (auto simp add: dom_def intro: map_upds_Some_expand 
+                                            map_upds_Some_insert)
+      qed
+    qed
+  qed
+qed
+
+lemma dom_map_upd: "\<And> tab. dom (tab(x\<mapsto>y)) = dom tab \<union> {x}"
+by (auto simp add: dom_def fun_upd_def)
+
+lemma dom_map_upds: "\<And> tab ys. dom (tab(xs[\<mapsto>]ys)) = dom tab \<union> set xs"
+proof (induct xs)
+  case Nil thus ?case by (simp add: dom_def)
+next
+  case (Cons x xs tab ys)
+  have "dom (tab(x\<mapsto>hd ys)(xs[\<mapsto>]tl ys)) = dom (tab(x\<mapsto>hd ys)) \<union> set xs" .
+  moreover 
+  have "dom (tab(x\<mapsto>hd ys)) = dom tab \<union> {x}"
+    by (rule dom_map_upd)
+  ultimately
+  show ?case
+    by simp
+qed
+ 
+
+
+lemma dom_ename_case_None_simp:
+ "dom (ename_case vname_tab None) = VNam ` (dom vname_tab)"
+  apply (auto simp add: dom_def image_def )
+  apply (case_tac "x")
+  apply auto
+  done
+
+lemma dom_ename_case_Some_simp:
+ "dom (ename_case vname_tab (Some a)) = VNam ` (dom vname_tab) \<union> {Res}"
+  apply (auto simp add: dom_def image_def )
+  apply (case_tac "x")
+  apply auto
+  done
+
+lemma dom_lname_case_None_simp:
+  "dom (lname_case ename_tab None) = EName ` (dom ename_tab)"
+  apply (auto simp add: dom_def image_def )
+  apply (case_tac "x")
+  apply auto
+  done
+
+lemma dom_lname_case_Some_simp:
+ "dom (lname_case ename_tab (Some a)) = EName ` (dom ename_tab) \<union> {This}"
+  apply (auto simp add: dom_def image_def)
+  apply (case_tac "x")
+  apply auto
+  done
+
+lemmas dom_lname_ename_case_simps =  
+     dom_ename_case_None_simp dom_ename_case_Some_simp 
+     dom_lname_case_None_simp dom_lname_case_Some_simp
+
+lemma image_comp: 
+ "f ` g ` A = (f \<circ> g) ` A"
+by (auto simp add: image_def)
+
+lemma dom_locals_init_lvars: 
+  assumes m: "m=(mthd (the (methd G C sig)))"  
+  shows "dom (locals (store (init_lvars G C sig (invmode m e) a pvs s)))  
+           = parameters m"
+proof -
+  from m
+  have static_m': "is_static m = static m"
+    by simp
+  have dom_vnames: "dom (empty(pars m[\<mapsto>]pvs))=set (pars m)"
+    by (simp add: dom_map_upds)
+  show ?thesis
+  proof (cases "static m")
+    case True
+    with static_m' dom_vnames m
+    show ?thesis
+      by (cases s) (simp add: init_lvars_def Let_def parameters_def
+                              dom_lname_ename_case_simps image_comp)
+  next
+    case False
+    with static_m' dom_vnames m
+    show ?thesis
+      by (cases s) (simp add: init_lvars_def Let_def parameters_def
+                              dom_lname_ename_case_simps image_comp)
+  qed
+qed
+
+lemma da_e2_BinOp:
+  assumes da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>BinOp binop e1 e2\<rangle>\<^sub>e\<guillemotright> A"
+    and wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-e1T"
+    and wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-e2T" 
+    and wt_binop: "wt_binop G binop e1T e2T" 
+    and conf_s0: "s0\<Colon>\<preceq>(G,L)"  
+    and normal_s1: "normal s1"
+    and	eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
+    and conf_v1: "G,store s1\<turnstile>v1\<Colon>\<preceq>e1T"
+    and wf: "wf_prog G"
+  shows "\<exists> E2. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) 
+         \<guillemotright>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<guillemotright> E2"
+proof -
+  note inj_term_simps [simp]
+  from da obtain E1 where
+    da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1"
+    by cases simp+
+  obtain E2 where
+    "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) 
+      \<guillemotright>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<guillemotright> E2"
+  proof (cases "need_second_arg binop v1")
+    case False
+    obtain S where
+      daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>Skip\<rangle>\<^sub>s\<guillemotright> S"
+      by (auto intro: da_Skip [simplified] assigned.select_convs)
+    thus ?thesis
+      using that by (simp add: False)
+  next
+    case True
+    from eval_e1 have 
+      s0_s1:"dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+      by (rule dom_locals_eval_mono_elim)
+    {
+      assume condAnd: "binop=CondAnd"
+      have ?thesis
+      proof -
+	from da obtain E2' where
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+             \<turnstile> dom (locals (store s0)) \<union> assigns_if True e1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> E2'"
+	  by cases (simp add: condAnd)+
+	moreover
+	have "dom (locals (store s0)) 
+          \<union> assigns_if True e1 \<subseteq> dom (locals (store s1))"
+	proof -
+	  from condAnd wt_binop have e1T: "e1T=PrimT Boolean"
+	    by simp
+	  with normal_s1 conf_v1 obtain b where "v1=Bool b"
+	    by (auto dest: conf_Boolean)
+	  with True condAnd
+	  have v1: "v1=Bool True"
+	    by simp
+	  from eval_e1 normal_s1 
+	  have "assigns_if True e1 \<subseteq> dom (locals (store s1))"
+	    by (rule assigns_if_good_approx' [elim_format])
+	       (insert wt_e1, simp_all add: e1T v1)
+	  with s0_s1 show ?thesis by (rule Un_least)
+	qed
+	ultimately
+	show ?thesis
+	  using that by (cases rule: da_weakenE) (simp add: True)
+      qed
+    }
+    moreover
+    { 
+      assume condOr: "binop=CondOr"
+      have ?thesis
+	(* Beweis durch Analogie/Example/Pattern?, True\<rightarrow>False; And\<rightarrow>Or *)
+      proof -
+	from da obtain E2' where
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+              \<turnstile> dom (locals (store s0)) \<union> assigns_if False e1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> E2'"
+	  by cases (simp add: condOr)+
+	moreover
+	have "dom (locals (store s0)) 
+                     \<union> assigns_if False e1 \<subseteq> dom (locals (store s1))"
+	proof -
+	  from condOr wt_binop have e1T: "e1T=PrimT Boolean"
+	    by simp
+	  with normal_s1 conf_v1 obtain b where "v1=Bool b"
+	    by (auto dest: conf_Boolean)
+	  with True condOr
+	  have v1: "v1=Bool False"
+	    by simp
+	  from eval_e1 normal_s1 
+	  have "assigns_if False e1 \<subseteq> dom (locals (store s1))"
+	    by (rule assigns_if_good_approx' [elim_format])
+	       (insert wt_e1, simp_all add: e1T v1)
+	  with s0_s1 show ?thesis by (rule Un_least)
+	qed
+	ultimately
+	show ?thesis
+	  using that by (rule da_weakenE) (simp add: True)
+      qed
+    }
+    moreover
+    {
+      assume notAndOr: "binop\<noteq>CondAnd" "binop\<noteq>CondOr"
+      have ?thesis
+      proof -
+	from da notAndOr obtain E1' where
+          da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1'"
+	  and da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1' \<guillemotright>In1l e2\<guillemotright> A"
+	  by cases simp+
+	from eval_e1 wt_e1 da_e1 wf normal_s1 
+	have "nrm E1' \<subseteq> dom (locals (store s1))"
+	  by (cases rule: da_good_approxE') rules
+	with da_e2 show ?thesis
+	  using that by (rule da_weakenE) (simp add: True)
+      qed
+    }
+    ultimately show ?thesis
+      by (cases binop) auto
+  qed
+  thus ?thesis ..
+qed
+
 section "main proof of type safety"
-
+    
 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)"           
+  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      da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>t\<guillemotright>A"
+   and      wf: "wf_prog G" 
+   and conf_s0: "s0\<Colon>\<preceq>(G,L)"           
   shows "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 -
+  note inj_term_simps [simp]
+  let ?TypeSafeObj = "\<lambda> s0 s1 t v. 
+          \<forall>  L accC T A. s0\<Colon>\<preceq>(G,L) \<longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T
+                      \<longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>t\<guillemotright>A  
+                      \<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)"
   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>  
+  have "\<And> L accC T A. \<lbrakk>s0\<Colon>\<preceq>(G,L);\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T;
+                      \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>t\<guillemotright>A\<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  
+    is "\<And> L accC T A. ?Conform L s0 \<Longrightarrow> ?WellTyped L accC T t  
+                 \<Longrightarrow> ?DefAss L accC s0 t A 
                  \<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) 
+    case (Abrupt s t xc L accC T A) 
     have "(Some xc, s)\<Colon>\<preceq>(G,L)" .
     then show "(Some xc, s)\<Colon>\<preceq>(G,L) \<and> 
       (normal (Some xc, s) 
@@ -1226,7 +1764,7 @@
       (error_free (Some xc, s) = error_free (Some xc, s))"
       by (simp)
   next
-    case (Skip s L accC T)
+    case (Skip s L accC T A)
     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>
@@ -1234,33 +1772,42 @@
               (error_free (Norm s) = error_free (Norm s))"
       by (simp)
   next
-    case (Expr e s0 s1 v L accC T)
+    case (Expr e s0 s1 v L accC T A)
     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)" .
+    moreover
     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 
+    moreover
+    from Expr.prems obtain E where
+      "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store ((Norm s0)::state)))\<guillemotright>In1l e\<guillemotright>E"
+      by (elim da_elim_cases) simp
+    ultimately 
     obtain "s1\<Colon>\<preceq>(G, L)" and "error_free s1"
-      by (blast)
+      by (rule hyp [elim_format]) simp
     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)
+    case (Lab c l s0 s1 L accC T A)
     have     hyp: "PROP ?TypeSafe (Norm s0) s1 (In1r c) \<diamondsuit>" .
     have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    moreover
     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
+    moreover from Lab.prems obtain C where
+     "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store ((Norm s0)::state)))\<guillemotright>In1r c\<guillemotright>C"
+      by (elim da_elim_cases) simp
+    ultimately
     obtain       conf_s1: "s1\<Colon>\<preceq>(G, L)" and 
            error_free_s1: "error_free s1" 
-      by (blast)
+      by (rule hyp [elim_format]) simp
     from conf_s1 have "abupd (absorb l) s1\<Colon>\<preceq>(G, L)"
       by (cases s1) (auto intro: conforms_absorb)
     with wt error_free_s1
@@ -1270,9 +1817,9 @@
           (error_free (Norm s0) = error_free (abupd (absorb 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" .
+    case (Comp c1 c2 s0 s1 s2 L accC T A)
+    have eval_c1: "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1" .
+    have eval_c2: "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)" .
@@ -1280,91 +1827,263 @@
     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
+    from Comp.prems
+    obtain C1 C2
+      where da_c1: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile> 
+                      dom (locals (store ((Norm s0)::state))) \<guillemotright>In1r c1\<guillemotright> C1" and 
+            da_c2: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>  nrm C1 \<guillemotright>In1r c2\<guillemotright> C2" 
+      by (elim da_elim_cases) simp
+    from conf_s0 wt_c1 da_c1
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and 
+           error_free_s1: "error_free s1"
+      by (rule hyp_c1 [elim_format]) simp
     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)
+    proof (cases "normal s1")
+      case False
+      with eval_c2 have "s2=s1" by auto
+      with conf_s1 error_free_s1 False wt show ?thesis
+	by simp
+    next
+      case True
+      obtain C2' where 
+	"\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In1r c2\<guillemotright> C2'"
+      proof -
+	from eval_c1 wt_c1 da_c1 wf True
+	have "nrm C1 \<subseteq> dom (locals (store s1))"
+	  by (cases rule: da_good_approxE') rules
+	with da_c2 show ?thesis
+	  by (rule da_weakenE)
+      qed
+      with conf_s1 wt_c2 
+      obtain "s2\<Colon>\<preceq>(G, L)" and "error_free s2"
+	by (rule hyp_c2 [elim_format]) (simp add: error_free_s1)
+      thus ?thesis
+	using wt by simp
+    qed
   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 eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1" .
+    have eval_then_else: "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>"
+    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>"
+      (*
+                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) (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
+    from If.prems obtain E C where
+      da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store ((Norm s0)::state))) 
+                                       \<guillemotright>In1l e\<guillemotright> E" and
+      da_then_else: 
+      "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+         (dom (locals (store ((Norm s0)::state))) \<union> assigns_if (the_Bool b) e)
+          \<guillemotright>In1r (if the_Bool b then c1 else c2)\<guillemotright> C"
+     (*
+     da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                                      \<union> assigns_if True e) \<guillemotright>In1r c1\<guillemotright> C1" and
+     da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                                       \<union> assigns_if False e) \<guillemotright>In1r c2\<guillemotright> C2" *)
+      by (elim da_elim_cases) (cases "the_Bool b",auto)
+    from conf_s0 wt_e da_e  
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
+      by (rule hyp_e [elim_format]) simp
     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)
+    proof (cases "normal s1")
+      case False
+      with eval_then_else have "s2=s1" by auto
+      with conf_s1 error_free_s1 False wt show ?thesis
+	by simp
+    next
+      case True
+      obtain C' where
+	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+          (dom (locals (store s1)))\<guillemotright>In1r (if the_Bool b then c1 else c2)\<guillemotright> C'"
+      proof -
+	from eval_e have 
+	  "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+	  by (rule dom_locals_eval_mono_elim)
+        moreover
+	from eval_e True wt_e 
+	have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+	  by (rule assigns_if_good_approx')
+	ultimately 
+	have "dom (locals (store ((Norm s0)::state))) 
+                \<union> assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+	  by (rule Un_least)
+	with da_then_else show ?thesis
+	  by (rule da_weakenE)
+      qed
+      with conf_s1 wt_then_else  
+      obtain "s2\<Colon>\<preceq>(G, L)" and "error_free s2"
+	by (rule hyp_then_else [elim_format]) (simp add: error_free_s1)
+      with wt show ?thesis
+	by simp
+    qed
+    -- {* Note that we don't have to show that @{term b} really is a boolean 
+          value. With @{term the_Bool} we enforce to get a value of boolean 
+          type. So execution will be type safe, even if b would be
+          a string, for example. We might not expect such a behaviour to be
+          called type safe. To remedy the situation we would have to change
+          the evaulation rule, so that it only has a type safe evaluation if
+          we actually get a boolean value for the condition. That b is actually
+          a boolean value is part of @{term hyp_e}. See also Loop 
+       *}
   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" .
+    case (Loop b c e l s0 s1 s2 s3 L accC T A)
+    have eval_e: "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
+    have da:"\<lparr>prg=G, cls=accC, lcl=L\<rparr>
+            \<turnstile> dom (locals(store ((Norm s0)::state))) \<guillemotright>In1r (l\<bullet> While(e) c)\<guillemotright> A".
+    then
+    obtain E C where
+      da_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>
+              \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e\<guillemotright> E" and
+      da_c: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>
+              \<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                   \<union> assigns_if True e) \<guillemotright>In1r c\<guillemotright> C" 
+      by (rule da_elim_cases) simp
+    from conf_s0 wt_e da_e 
     obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
-      by blast
+      by (rule hyp_e [elim_format]) simp
     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")
+    proof (cases "normal s1")
       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 
+      note normal_s1 = this
       show ?thesis
-	by (simp)
+      proof (cases "the_Bool b")
+	case True
+	with Loop.hyps  obtain
+          eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" and 
+          eval_while: "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
+	  by simp 
+	have "?TypeSafeObj s1 s2 (In1r c) \<diamondsuit>"
+	  using Loop.hyps True by simp
+	note hyp_c = this [rule_format]
+	have "?TypeSafeObj (abupd (absorb (Cont l)) s2)
+          s3 (In1r (l\<bullet> While(e) c)) \<diamondsuit>"
+	  using Loop.hyps True by simp
+	note hyp_w = this [rule_format]
+	from eval_e have 
+	  s0_s1: "dom (locals (store ((Norm s0)::state)))
+                    \<subseteq> dom (locals (store s1))"
+	  by (rule dom_locals_eval_mono_elim)
+	obtain C' where
+	  "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(dom (locals (store s1)))\<guillemotright>In1r c\<guillemotright> C'" 
+	proof -
+	  note s0_s1
+          moreover
+	  from eval_e normal_s1 wt_e 
+	  have "assigns_if True e \<subseteq> dom (locals (store s1))"
+	    by (rule assigns_if_good_approx' [elim_format]) (simp add: True)
+	  ultimately 
+	  have "dom (locals (store ((Norm s0)::state))) 
+                 \<union> assigns_if True e \<subseteq> dom (locals (store s1))"
+	    by (rule Un_least)
+	  with da_c show ?thesis
+	    by (rule da_weakenE)
+	qed
+	with conf_s1 wt_c  
+	obtain conf_s2:  "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
+	  by (rule hyp_c [elim_format]) (simp add: error_free_s1)
+	from error_free_s2 
+	have error_free_ab_s2: "error_free (abupd (absorb (Cont l)) s2)"
+	  by simp
+	from conf_s2 have "abupd (absorb (Cont l)) s2 \<Colon>\<preceq>(G, L)"
+	  by (cases s2) (auto intro: conforms_absorb)
+	moreover note wt
+	moreover
+	obtain A' where 
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+              dom (locals(store (abupd (absorb (Cont l)) s2)))
+                \<guillemotright>In1r (l\<bullet> While(e) c)\<guillemotright> A'"
+	proof -
+	  note s0_s1
+	  also from eval_c 
+	  have "dom (locals (store s1)) \<subseteq> dom (locals (store s2))"
+	    by (rule dom_locals_eval_mono_elim)
+	  also have "\<dots> \<subseteq> dom (locals (store (abupd (absorb (Cont l)) s2)))"
+	    by simp
+	  finally
+          have "dom (locals (store ((Norm s0)::state))) \<subseteq> \<dots>" .
+	  with da show ?thesis
+	    by (rule da_weakenE)
+	qed
+	ultimately obtain "s3\<Colon>\<preceq>(G, L)" and "error_free s3"
+	  by (rule hyp_w [elim_format]) (simp add: error_free_ab_s2)
+	with wt show ?thesis
+	  by simp
+      next
+	case False
+	with Loop.hyps have "s3=s1" by simp
+	with conf_s1 error_free_s1 wt
+	show ?thesis
+	  by simp
+      qed
     next
       case False
-      with Loop have "s3=s1" by simp
+      have "s3=s1"
+      proof -
+	from False obtain abr where abr: "abrupt s1 = Some abr"
+	  by (cases s1) auto
+	from eval_e _ wt_e have no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+	  by (rule eval_expression_no_jump 
+               [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified]) 
+             (simp_all add: wf)
+	    
+	show ?thesis
+	proof (cases "the_Bool b")
+	  case True  
+	  with Loop.hyps obtain
+            eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" and 
+            eval_while: "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
+	    by simp
+	  from eval_c abr have "s2=s1" by auto
+	  moreover from calculation no_jmp have "abupd (absorb (Cont l)) s2=s2"
+	    by (cases s1) (simp add: absorb_def)
+	  ultimately show ?thesis
+	    using eval_while abr
+	    by auto
+	next
+	  case False
+	  with Loop.hyps show ?thesis by simp
+	qed
+      qed
       with conf_s1 error_free_s1 wt
       show ?thesis
-	by (simp)
+	by simp
     qed
   next
-    case (Do j s L accC T)
+    case (Jmp j s L accC T A)
     have "Norm s\<Colon>\<preceq>(G, L)" .
+    moreover
+    from Jmp.prems 
+    have "j=Ret \<longrightarrow> Result \<in> dom (locals (store ((Norm s)::state)))"
+      by (elim da_elim_cases)
+    ultimately have "(Some (Jump j), s)\<Colon>\<preceq>(G, L)" by auto
     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>
+           \<longrightarrow> G,L,store (Some (Jump j), s)\<turnstile>In1r (Jmp 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)
+    case (Throw a e s0 s1 L accC T A)
     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)" .
@@ -1373,11 +2092,15 @@
       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
+    from Throw.prems obtain E where 
+      da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+             \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e\<guillemotright> E"
+      by (elim da_elim_cases) simp
+    from conf_s0 wt_e da_e 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
+      by (rule hyp [elim_format]) simp
     with wf throwable
     have "abupd (throw a) s1\<Colon>\<preceq>(G, L)" 
       by (cases s1) (auto dest: Throw_lemma)
@@ -1388,35 +2111,43 @@
             (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" .
+    case (Try catchC c1 c2 s0 s1 s2 s3 vn L accC T A)
+    have eval_c1: "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
+      wt_c2: "\<lparr>prg=G,cls=accC,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
+      by (rule wt_elim_cases) simp
+    from Try.prems obtain C1 C2 where
+      da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1r c1\<guillemotright> C1"  and
+      da_c2:
+       "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class catchC)\<rparr>
+        \<turnstile> (dom (locals (store ((Norm s0)::state))) \<union> {VName vn}) \<guillemotright>In1r c2\<guillemotright> C2"
+      by (elim da_elim_cases) simp
+    from conf_s0 wt_c1 da_c1
     obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
-      by blast
+      by (rule hyp_c1 [elim_format]) simp
     from conf_s1 sx_alloc wf 
     have conf_s2: "s2\<Colon>\<preceq>(G, L)" 
-      by (auto dest: sxalloc_type_sound split: option.splits)
+      by (auto dest: sxalloc_type_sound split: option.splits abrupt.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 
+    proof (cases "\<exists> x. abrupt s1 = Some (Xcpt x)")
+      case False
+      from sx_alloc wf
       have eq_s2_s1: "s2=s1"
-	by (auto dest: sxalloc_type_sound split: option.splits)
-      with True 
+	by (rule sxalloc_type_sound [elim_format])
+	   (insert False, auto split: option.splits abrupt.splits )
+      with False 
       have "\<not>  G,s2\<turnstile>catch catchC"
 	by (simp add: catch_def)
       with Try
@@ -1426,7 +2157,7 @@
       show ?thesis
 	by simp
     next
-      case False
+      case True
       note exception_s1 = this
       show ?thesis
       proof (cases "G,s2\<turnstile>catch catchC") 
@@ -1440,24 +2171,54 @@
       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 True Try.hyps
+	have "?TypeSafeObj (new_xcpt_var vn s2) s3 (In1r c2) \<diamondsuit>"
+	  by simp
+	note hyp_c2 = this [rule_format]
 	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)
+	  by (auto dest!: sxalloc_type_sound split: option.splits abrupt.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))"
+	have "new_xcpt_var vn 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
+	moreover note wt_c2
+	moreover
+	obtain C2' where
+	  "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class catchC)\<rparr>
+          \<turnstile> (dom (locals (store (new_xcpt_var vn s2)))) \<guillemotright>In1r c2\<guillemotright> C2'"
+	proof -
+	  have "(dom (locals (store ((Norm s0)::state))) \<union> {VName vn}) 
+                  \<subseteq> dom (locals (store (new_xcpt_var vn s2)))"
+          proof -
+            have "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1" .
+            hence "dom (locals (store ((Norm s0)::state))) 
+                    \<subseteq> dom (locals (store s1))"
+              by (rule dom_locals_eval_mono_elim)
+            also
+            from sx_alloc
+            have "\<dots> \<subseteq> dom (locals (store s2))"
+              by (rule dom_locals_sxalloc_mono)
+            also 
+            have "\<dots> \<subseteq> dom (locals (store (new_xcpt_var vn s2)))" 
+              by (cases s2) (simp add: new_xcpt_var_def, blast) 
+            also
+            have "{VName vn} \<subseteq> \<dots>"
+              by (cases s2) simp
+            ultimately show ?thesis
+              by (rule Un_least)
+          qed
+	  with da_c2 show ?thesis
+	    by (rule da_weakenE)
+	qed
+	ultimately
 	obtain       conf_s3: "s3\<Colon>\<preceq>(G, L(VName vn\<mapsto>Class catchC))" and
                error_free_s3: "error_free s3"
-	  by (cases s2) auto
+	  by (rule hyp_c2 [elim_format])
+             (cases s2, simp add: xcpt_s2 error_free_s2) 
 	from conf_s3 fresh_vn 
 	have "s3\<Colon>\<preceq>(G,L)"
 	  by (blast intro: conforms_deallocL)
@@ -1467,9 +2228,9 @@
       qed
     qed
   next
-    case (Fin c1 c2 s0 s1 s2 s3 x1 L accC T)
-    have "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1, s1)" .
-    have c2: "G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2" .
+    case (Fin c1 c2 s0 s1 s2 s3 x1 L accC T A)
+    have eval_c1: "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1, s1)" .
+    have eval_c2: "G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2" .
     have s3: "s3= (if \<exists>err. x1 = Some (Error err) 
                      then (x1, s1)
                      else abupd (abrupt_if (x1 \<noteq> None) x1) s2)" .
@@ -1478,17 +2239,38 @@
     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>"
+      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  
+    from Fin.prems obtain C1 C2 where 
+      da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+               \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1r c1\<guillemotright> C1" and
+      da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+               \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1r c2\<guillemotright> C2"
+      by (elim da_elim_cases) simp 
+    from conf_s0 wt_c1 da_c1   
     obtain conf_s1: "(x1,s1)\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free (x1,s1)"
-      by blast
+      by (rule hyp_c1 [elim_format]) simp
     from conf_s1 have "Norm s1\<Colon>\<preceq>(G, L)"
       by (rule conforms_NormI)
-    with wt_c2 hyp_c2
+    moreover note wt_c2
+    moreover obtain C2'
+      where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+               \<turnstile> dom (locals (store ((Norm s1)::state))) \<guillemotright>In1r c2\<guillemotright> C2'"
+    proof -
+      from eval_c1
+      have "dom (locals (store ((Norm s0)::state))) 
+             \<subseteq> dom (locals (store (x1,s1)))"
+        by (rule dom_locals_eval_mono_elim)
+      hence "dom (locals (store ((Norm s0)::state))) 
+              \<subseteq> dom (locals (store ((Norm s1)::state)))"
+	by simp
+      with da_c2 show ?thesis
+	by (rule da_weakenE)
+    qed
+    ultimately
     obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
-      by blast
+      by (rule hyp_c2 [elim_format]) simp
     from error_free_s1 s3 
     have s3': "s3=abupd (abrupt_if (x1 \<noteq> None) x1) s2"
       by simp
@@ -1499,7 +2281,10 @@
       case None with conf_s2 s3' wt show ?thesis by auto
     next
       case (Some x) 
-      with c2 wf conf_s1 conf_s2
+      from eval_c2 have 
+	"dom (locals (store ((Norm s1)::state))) \<subseteq> dom (locals (store s2))"
+	by (rule dom_locals_eval_mono_elim)
+      with Some eval_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
@@ -1523,50 +2308,91 @@
           (error_free (Norm s0) = error_free s3)"
     proof (cases "inited C (globs s0)")
       case True
-      with Init have "s3 = Norm s0"
+      with Init.hyps 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) 
+      with Init.hyps obtain 
+           eval_init_super: 
+           "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
+	by simp
+      have "?TypeSafeObj (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
+	using False Init.hyps by simp
+      note hyp_init_super = this [rule_format] 
+      have "?TypeSafeObj ((set_lvars empty) s1) s2 (In1r (init c)) \<diamondsuit>"
+	using False Init.hyps by simp
+      note hyp_init_c = this [rule_format]
       from conf_s0 wf cls_C False
-      have conf_s0': "(Norm ((init_class_obj G C) s0))\<Colon>\<preceq>(G, L)"
+      have "(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>"
+      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)
-      with conf_s0' hyp_init_super
+      moreover
+      obtain S where
+	da_init_super:
+	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+          \<turnstile> dom (locals (store ((Norm ((init_class_obj G C) s0))::state))) 
+               \<guillemotright>In1r (if C = Object then Skip else Init (super c))\<guillemotright> S"
+      proof (cases "C=Object")
+	case True 
+	with da_Skip show ?thesis
+	  using that by (auto intro: assigned.select_convs)
+      next
+	case False 
+	with da_Init show ?thesis
+	  by - (rule that, auto intro: assigned.select_convs)
+      qed
+      ultimately 
       obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
-	by blast 
-      then
+	by (rule hyp_init_super [elim_format]) simp
+      from eval_init_super wt_init_super wf
+      have s1_no_ret: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+	by - (rule eval_statement_no_jump [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>"],
+              auto)
+      with conf_s1
       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 (cases s1) (auto intro: conforms_set_locals)
+      moreover 
+      from error_free_s1
+      have error_free_empty: "error_free ((set_lvars empty) s1)"
 	by simp
-      moreover note hyp_init_c wf cls_C 
+      from cls_C wf have wt_init_c: "\<lparr>prg=G, cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd>"
+	by (rule wf_prog_cdecl [THEN wf_cdecl_wt_init])
+      moreover from cls_C wf obtain I
+	where "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>In1r (init c)\<guillemotright> I"
+	by (rule wf_prog_cdecl [THEN wf_cdeclE,simplified]) blast
+       (*  simplified: to rewrite \<langle>init c\<rangle> to In1r (init c) *) 
+      then obtain I' where
+	"\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>dom (locals (store ((set_lvars empty) s1))) 
+            \<guillemotright>In1r (init c)\<guillemotright> I'"
+	  by (rule da_weakenE) simp
       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
+	by (rule hyp_init_c [elim_format]) (simp add: error_free_empty)
+      have "abrupt s2 \<noteq> Some (Jump Ret)"
+      proof -
+	from s1_no_ret 
+	have "\<And> j. abrupt ((set_lvars empty) s1) \<noteq> Some (Jump j)"
+	  by simp
+	moreover
+	from cls_C wf have "jumpNestingOkS {} (init c)"
+	  by (rule wf_prog_cdecl [THEN wf_cdeclE])
+	ultimately 
+	show ?thesis
+	  using eval_init wt_init_c wf
+	  by - (rule eval_statement_no_jump 
+                     [where ?Env="\<lparr>prg=G,cls=C,lcl=empty\<rparr>"],simp+)
+      qed
+      with conf_s2 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
@@ -1577,18 +2403,25 @@
 	by simp
     qed
   next
-    case (NewC C a s0 s1 s2 L accC T)
+    case (NewC C a s0 s1 s2 L accC T A)
     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" .
+    moreover
+    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
+    hence "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>Init C\<Colon>\<surd>" by auto
+    moreover obtain I where 
+      "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+          \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1r (Init C)\<guillemotright> I"
+      by (auto intro: da_Init [simplified] assigned.select_convs)
+     (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
+    ultimately 
     obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
-      by auto
+      by (rule hyp [elim_format]) simp 
     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)"
@@ -1602,42 +2435,78 @@
           (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>" .
+    case (NewA elT a e i s0 s1 s2 s3 L accC T A)
+    have eval_init: "G\<turnstile>Norm s0 \<midarrow>init_comp_ty elT\<rightarrow> s1" .
+    have eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>i\<rightarrow> s2" .
+    have halloc: "G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr elT (the_Intg i)\<succ>a\<rightarrow> s3".
+    have hyp_init: "PROP ?TypeSafe (Norm s0) s1 (In1r (init_comp_ty elT)) \<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" .
+    have     wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (New elT[e])\<Colon>T" .
     then obtain
-      wt_init: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and
+      wt_init: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>init_comp_ty elT\<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.[])"
+            elT: "is_type G elT" and
+           T: "T=Inl (elT.[])"
       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
+    from NewA.prems 
+    have da_e:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                 \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e\<guillemotright> A"
+      by (elim da_elim_cases) simp
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
+    proof -
+      note conf_s0 wt_init
+      moreover obtain I where
+	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+         \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1r (init_comp_ty elT)\<guillemotright> I"
+      proof (cases "\<exists>C. elT = Class C")
+	case True
+	thus ?thesis
+	  by - (rule that, (auto intro: da_Init [simplified] 
+                                        assigned.select_convs
+                              simp add: init_comp_ty_def))
+	 (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
+      next
+	case False
+	thus ?thesis
+	by - (rule that, (auto intro: da_Skip [simplified] 
+                                      assigned.select_convs
+                           simp add: init_comp_ty_def))
+         (* simplified: to rewrite \<langle>Skip\<rangle> to In1r (Skip) *)
+      qed
+      ultimately show ?thesis
+	by (rule hyp_init [elim_format]) auto
+    qed 
     obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
-      by blast
+    proof -
+      from eval_init 
+      have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+	by (rule dom_locals_eval_mono_elim)
+      with da_e 
+      obtain A' where
+       "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+            \<turnstile> dom (locals (store s1)) \<guillemotright>In1l e\<guillemotright> A'"
+	by (rule da_weakenE)
+      with conf_s1 wt_size
+      show ?thesis
+	by (rule hyp_size [elim_format]) (simp add: that error_free_s1) 
+    qed
     from conf_s2 have "abupd (check_neg i) s2\<Colon>\<preceq>(G, L)"
       by (cases s2) auto
-    with halloc wf T 
+    with halloc wf elT 
     have halloc_type_safe:
-          "s3\<Colon>\<preceq>(G, L) \<and> (normal s3 \<longrightarrow> G,store s3\<turnstile>Addr a\<Colon>\<preceq>T.[])"
+          "s3\<Colon>\<preceq>(G, L) \<and> (normal s3 \<longrightarrow> G,store s3\<turnstile>Addr a\<Colon>\<preceq>elT.[])"
       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
+    with halloc_type_safe T
     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>
+          (normal s3 \<longrightarrow> G,L,store s3\<turnstile>In1l (New elT[e])\<succ>In1 (Addr a)\<Colon>\<preceq>T) \<and>
           (error_free (Norm s0) = error_free s3) "
       by simp
   next
-    case (Cast castT e s0 s1 s2 v L accC T)
+    case (Cast castT e s0 s1 s2 v L accC T A)
     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)" .
@@ -1648,9 +2517,15 @@
               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 Cast.prems 
+    have "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                 \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e\<guillemotright> A"
+      by (elim da_elim_cases) simp
+    with conf_s0 wt_e
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and
+           v_ok: "normal s1 \<longrightarrow> G,store s1\<turnstile>v\<Colon>\<preceq>eT" and
+      error_free_s1: "error_free s1"
+      by (rule hyp [elim_format]) simp
     from conf_s1 s2 
     have conf_s2: "s2\<Colon>\<preceq>(G, L)"
       by (cases s1) simp
@@ -1663,9 +2538,9 @@
       proof -
 	from s2 norm_s2 have "normal s1"
 	  by (cases s1) simp
-	with wt_e conf_s0 hyp 
+	with v_ok 
 	have "G,store s1\<turnstile>v\<Colon>\<preceq>eT"
-	  by force
+	  by simp
 	with eT wf s2 T norm_s2
 	show ?thesis
 	  by (cases s1) (auto dest: fits_conf)
@@ -1677,16 +2552,31 @@
            (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)
+    case (Inst instT b e s0 s1 v L accC T A)
+    have hyp: "PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    from Inst.prems obtain eT
+    where wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-RefT eT"  and
+             T: "T=Inl (PrimT Boolean)" 
+      by (elim wt_elim_cases) simp
+    from Inst.prems 
+    have da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                 \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e\<guillemotright> A"
+      by (elim da_elim_cases) simp
+    from conf_s0 wt_e da_e
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and
+              v_ok: "normal s1 \<longrightarrow> G,store s1\<turnstile>v\<Colon>\<preceq>RefT eT" and
+      error_free_s1: "error_free s1"
+      by (rule hyp [elim_format]) simp
+    with T show ?case
+      by simp
   next
-    case (Lit s v L accC T)
+    case (Lit s v L accC T A)
     then show ?case
       by (auto elim!: wt_elim_cases 
                intro: conf_litval simp add: empty_dt_def)
   next
-    case (UnOp e s0 s1 unop v L accC T)
+    case (UnOp e s0 s1 unop v L accC T A)
     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 (UnOp unop e)\<Colon>T" .
@@ -1695,11 +2585,15 @@
             wt_unop: "wt_unop unop eT" and
                   T: "T=Inl (PrimT (unop_type unop))" 
       by (auto elim!: wt_elim_cases)
-    from conf_s0 wt_e 
+    from UnOp.prems obtain A where
+       da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e\<guillemotright> A"
+      by (elim da_elim_cases) simp
+    from conf_s0 wt_e da_e
     obtain     conf_s1: "s1\<Colon>\<preceq>(G, L)"  and
                   wt_v: "normal s1 \<longrightarrow> G,store s1\<turnstile>v\<Colon>\<preceq>eT" and
          error_free_s1: "error_free s1"
-      by (auto dest!: hyp)
+      by (rule hyp [elim_format]) simp
     from wt_v T wt_unop
     have "normal s1\<longrightarrow>G,L,snd s1\<turnstile>In1l (UnOp unop e)\<succ>In1 (eval_unop unop v)\<Colon>\<preceq>T"
       by (cases unop) auto
@@ -1709,7 +2603,10 @@
      error_free (Norm s0) = error_free s1"
       by simp
   next
-    case (BinOp binop e1 e2 s0 s1 s2 v1 v2 L accC T)
+    case (BinOp binop e1 e2 s0 s1 s2 v1 v2 L accC T A)
+    have eval_e1: "G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1" .
+    have eval_e2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
+                             else In1r Skip)\<succ>\<rightarrow> (In1 v2, s2)" .
     have hyp_e1: "PROP ?TypeSafe (Norm s0) s1 (In1l e1) (In1 v1)" .
     have hyp_e2: "PROP ?TypeSafe       s1  s2 
                    (if need_second_arg binop v1 then In1l e2 else In1r Skip) 
@@ -1721,30 +2618,83 @@
          wt_e2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e2\<Colon>-e2T" and
       wt_binop: "wt_binop G binop e1T e2T" and
              T: "T=Inl (PrimT (binop_type binop))"
-      by (auto elim!: wt_elim_cases)
+      by (elim wt_elim_cases) simp
     have wt_Skip: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>Skip\<Colon>\<surd>"
       by simp
-    from conf_s0 wt_e1 
+    obtain S where
+      daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> dom (locals (store s1)) \<guillemotright>In1r Skip\<guillemotright> S"
+      by (auto intro: da_Skip [simplified] assigned.select_convs)
+    have da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store ((Norm s0::state)))) 
+                  \<guillemotright>\<langle>BinOp binop e1 e2\<rangle>\<^sub>e\<guillemotright> A".
+    then obtain E1 where
+      da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e1\<guillemotright> E1"
+      by (elim da_elim_cases) simp+
+    from conf_s0 wt_e1 da_e1
     obtain      conf_s1: "s1\<Colon>\<preceq>(G, L)"  and
                   wt_v1: "normal s1 \<longrightarrow> G,store s1\<turnstile>v1\<Colon>\<preceq>e1T" and
           error_free_s1: "error_free s1"
-      by (auto dest!: hyp_e1)
-    from conf_s1 wt_e2 wt_Skip error_free_s1
-    obtain      conf_s2: "s2\<Colon>\<preceq>(G, L)"  and
-          error_free_s2: "error_free s2"
-      by (cases "need_second_arg binop v1")
-         (auto dest!: hyp_e2 )
+      by (rule hyp_e1 [elim_format]) simp
     from wt_binop T
-    have "G,L,snd s2\<turnstile>In1l (BinOp binop e1 e2)\<succ>In1 (eval_binop binop v1 v2)\<Colon>\<preceq>T"
+    have conf_v:
+      "G,L,snd s2\<turnstile>In1l (BinOp binop e1 e2)\<succ>In1 (eval_binop binop v1 v2)\<Colon>\<preceq>T"
       by (cases binop) auto
-    with conf_s2 conf_s1 error_free_s2 error_free_s1
+    -- {* Note that we don't use the information that v1 really is compatible 
+          with the expected type e1T and v2 is compatible with e2T, 
+          because @{text eval_binop} will anyway produce an output of 
+          the right type.
+          So evaluating the addition of an integer with a string is type
+          safe. This is a little bit annoying since we may regard such a
+          behaviour as not type safe.
+          If we want to avoid this we can redefine @{text eval_binop} so that
+          it only produces a output of proper type if it is assigned to 
+          values of the expected types, and arbitrary if the inputs have 
+          unexpected types. The proof can easily be adapted since we
+          have the hypothesis that the values have a proper type.
+          This also applies to unary operations.
+       *}
+    from eval_e1 have 
+      s0_s1:"dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+      by (rule dom_locals_eval_mono_elim)
     show "s2\<Colon>\<preceq>(G, L) \<and>
           (normal s2 \<longrightarrow>
         G,L,snd s2\<turnstile>In1l (BinOp binop e1 e2)\<succ>In1 (eval_binop binop v1 v2)\<Colon>\<preceq>T) \<and>
           error_free (Norm s0) = error_free s2"
-      by simp
+    proof (cases "normal s1")
+      case False
+      with eval_e2 have "s2=s1" by auto
+      with conf_s1 error_free_s1 False show ?thesis
+	by auto
+    next
+      case True
+      note normal_s1 = this
+      show ?thesis 
+      proof (cases "need_second_arg binop v1")
+	case False
+	with normal_s1 eval_e2 have "s2=s1"
+	  by (cases s1) (simp, elim eval_elim_cases,simp)
+	with conf_s1 conf_v error_free_s1
+	show ?thesis by simp
+      next
+	case True
+	note need_second_arg = this
+	with hyp_e2 
+	have hyp_e2': "PROP ?TypeSafe s1 s2 (In1l e2) (In1 v2)" by simp
+	from da wt_e1 wt_e2 wt_binop conf_s0 normal_s1 eval_e1 
+          wt_v1 [rule_format,OF normal_s1] wf
+	obtain E2 where
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In1l e2\<guillemotright> E2"
+	  by (rule da_e2_BinOp [elim_format]) 
+             (auto simp add: need_second_arg )
+	with conf_s1 wt_e2 
+	obtain "s2\<Colon>\<preceq>(G, L)" and "error_free s2"
+	  by (rule hyp_e2' [elim_format]) (simp add: error_free_s1)
+	with conf_v show ?thesis by simp
+      qed
+    qed
   next
-    case (Super s L accC T)
+    case (Super s L accC T A)
     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
@@ -1753,8 +2703,11 @@
          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])
+    from Super.prems
+    obtain "This \<in> dom (locals s)"
+      by (elim da_elim_cases) simp
+    with conf_s C  have "G,s\<turnstile>val_this s\<Colon>\<preceq>Class C"
+      by (auto dest: conforms_localD [THEN wlconfD])
     with neq_Obj cls_C wf
     have "G,s\<turnstile>val_this s\<Colon>\<preceq>Class (super c)"
       by (auto intro: conf_widen
@@ -1766,14 +2719,50 @@
            (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)
+    case (Acc upd s0 s1 w v L accC T A) 
+    have hyp: "PROP ?TypeSafe (Norm s0) s1 (In2 v) (In2 (w,upd))" .
+    have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
+    from Acc.prems obtain vT where
+      wt_v: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>v\<Colon>=vT" and
+         T: "T=Inl vT"
+      by (elim wt_elim_cases) simp
+    from Acc.prems obtain V where
+      da_v: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In2 v\<guillemotright> V"
+      by (cases "\<exists> n. v=LVar n") (insert da.LVar,auto elim!: da_elim_cases)
+    {
+      fix n assume lvar: "v=LVar n"
+      have "locals (store s1) n \<noteq> None"
+      proof -
+	from Acc.prems lvar have 
+	  "n \<in> dom (locals s0)"
+	  by (cases "\<exists> n. v=LVar n") (auto elim!: da_elim_cases)
+	also
+	have "dom (locals s0) \<subseteq> dom (locals (store s1))"
+	proof -
+	  have "G\<turnstile>Norm s0 \<midarrow>v=\<succ>(w, upd)\<rightarrow> s1" .
+	  thus ?thesis
+	    by (rule dom_locals_eval_mono_elim) simp
+	qed
+	finally show ?thesis
+	  by blast
+      qed
+    } note lvar_in_locals = this 
+    from conf_s0 wt_v da_v
+    obtain conf_s1: "s1\<Colon>\<preceq>(G, L)"
+      and  conf_var: "(normal s1 \<longrightarrow> G,L,store s1\<turnstile>In2 v\<succ>In2 (w, upd)\<Colon>\<preceq>Inl vT)"
+      and  error_free_s1: "error_free s1"
+      by (rule hyp [elim_format]) simp
+    from lvar_in_locals conf_var T
+    have "(normal s1 \<longrightarrow> G,L,store s1\<turnstile>In1l (Acc v)\<succ>In1 w\<Colon>\<preceq>T)"
+      by (cases "\<exists> n. v=LVar n") auto
+    with conf_s1 error_free_s1 show ?case
+      by simp
   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" .
+    case (Ass e upd s0 s1 s2 v var w L accC T A)
+    have eval_var: "G\<turnstile>Norm s0 \<midarrow>var=\<succ>(w, upd)\<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_var: "PROP ?TypeSafe (Norm s0) s1 (In2 var) (In2 (w,upd))" .
     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" .
@@ -1783,58 +2772,157 @@
 	  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")
+    show "assign upd v s2\<Colon>\<preceq>(G, L) \<and>
+           (normal (assign upd v s2) \<longrightarrow>
+            G,L,store (assign upd v s2)\<turnstile>In1l (var:=e)\<succ>In1 v\<Colon>\<preceq>T) \<and>
+      (error_free (Norm s0) = error_free (assign upd v s2))"
+    proof (cases "\<exists> vn. var=LVar vn")
       case False
-      with eval_e 
-      have "s2=s1"
-	by auto
-      with False conf_s1 error_free_s1
+      with Ass.prems
+      obtain V E where
+	da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                   \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In2 var\<guillemotright> V" and
+	da_e:   "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>In1l e\<guillemotright> E"
+	by (elim da_elim_cases) simp+
+      from conf_s0 wt_var da_var 
+      obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" 
+	and  conf_var: "normal s1 
+                         \<longrightarrow> G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
+	and  error_free_s1: "error_free s1"
+	by (rule hyp_var [elim_format]) simp
       show ?thesis
-	by auto
-    next
-      case True
-      note normal_s1=this
-      show ?thesis 
-      proof (cases "normal s2")
+      proof (cases "normal s1")
 	case False
-	with conf_s2 error_free_s2 
-	show ?thesis
+	with eval_e have "s2=s1" by auto
+	with False have "assign upd v s2=s1"
+	  by simp
+	with conf_s1 error_free_s1 False 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)" 
+	note normal_s1=this
+	obtain A' where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                         \<turnstile> dom (locals (store s1)) \<guillemotright>In1l e\<guillemotright> A'"
+	proof -
+	  from eval_var wt_var da_var wf normal_s1
+	  have "nrm V \<subseteq>  dom (locals (store s1))"
+	    by (cases rule: da_good_approxE') rules
+	  with da_e show ?thesis
+	    by (rule da_weakenE) 
+	qed
+	with conf_s1 wt_e 
+	obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and 
+          conf_v: "normal s2 \<longrightarrow> G,store s2\<turnstile>v\<Colon>\<preceq>eT" and
+          error_free_s2: "error_free s2"
+	  by (rule hyp_e [elim_format]) (simp add: error_free_s1)
+	show ?thesis 
+	proof (cases "normal s2")
+	  case False
+	  with conf_s2 error_free_s2 
+	  show ?thesis
+	    by auto
+	next
+	  case True
+	  from True conf_v
+	  have conf_v_eT: "G,store s2\<turnstile>v\<Colon>\<preceq>eT"
+	    by simp
+	  with widen wf
+	  have conf_v_varT: "G,store s2\<turnstile>v\<Colon>\<preceq>varT"
+	    by (auto intro: conf_widen)
+	  from normal_s1 conf_var
+	  have "G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
+	    by simp
+	  then 
+	  have conf_assign:  "store s1\<le>|upd\<preceq>varT\<Colon>\<preceq>(G, L)" 
+	    by (simp add: rconf_def)
+	  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 True
+      then obtain vn where vn: "var=LVar vn"
+	by blast
+      with Ass.prems
+      obtain E where
+	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> 
+	           \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e\<guillemotright> E"
+	by (elim da_elim_cases) simp+
+      from da.LVar vn obtain V where
+	da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                   \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In2 var\<guillemotright> V"
+	by auto
+      obtain E' where
+	da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                   \<turnstile> dom (locals (store s1)) \<guillemotright>In1l e\<guillemotright> E'"
+      proof -
+	have "dom (locals (store ((Norm s0)::state))) 
+                \<subseteq> dom (locals (store (s1)))"
+	  by (rule dom_locals_eval_mono_elim)
+	with da_e show ?thesis
+	  by (rule da_weakenE)
+      qed
+      from conf_s0 wt_var da_var 
+      obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" 
+	and  conf_var: "normal s1 
+                         \<longrightarrow> G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
+	and  error_free_s1: "error_free s1"
+	by (rule hyp_var [elim_format]) simp
+      show ?thesis
+      proof (cases "normal s1")
+	case False
+	with eval_e have "s2=s1" by auto
+	with False have "assign upd v s2=s1"
+	  by simp
+	with conf_s1 error_free_s1 False show ?thesis
 	  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)
+      next
+	case True
+	note normal_s1 = this
+	from conf_s1 wt_e da_e'
+	obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and 
+          conf_v: "normal s2 \<longrightarrow> G,store s2\<turnstile>v\<Colon>\<preceq>eT" and
+          error_free_s2: "error_free s2"
+	  by (rule hyp_e [elim_format]) (simp add: error_free_s1)
+	show ?thesis 
+	proof (cases "normal s2")
+	  case False
+	  with conf_s2 error_free_s2 
+	  show ?thesis
+	    by auto
+	next
+	  case True
+	  from True conf_v
+	  have conf_v_eT: "G,store s2\<turnstile>v\<Colon>\<preceq>eT"
+	    by simp
+	  with widen wf
+	  have conf_v_varT: "G,store s2\<turnstile>v\<Colon>\<preceq>varT"
+	    by (auto intro: conf_widen)
+	  from normal_s1 conf_var
+	  have "G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
+	    by simp
+	  then 
+	  have conf_assign:  "store s1\<le>|upd\<preceq>varT\<Colon>\<preceq>(G, L)" 
+	    by (simp add: rconf_def)
+	  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
     qed
   next
-    case (Cond b e0 e1 e2 s0 s1 s2 v L accC T)
+-- {* 
+\par
+*} (* dummy text command to break paragraph for latex;
+              large paragraphs exhaust memory of debian pdflatex *)
+    case (Cond b e0 e1 e2 s0 s1 s2 v L accC T A)
     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 eval_e1_e2: "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)" .
@@ -1847,40 +2935,98 @@
       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
+    with Cond.prems obtain E0 E1 E2 where
+         da_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> dom (locals (store ((Norm s0)::state))) 
+                      \<guillemotright>In1l e0\<guillemotright> E0" and
+         da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                         \<union> assigns_if True e0) \<guillemotright>In1l e1\<guillemotright> E1" and
+         da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                  \<turnstile> (dom (locals (store ((Norm s0)::state))) 
+                        \<union> assigns_if False e0) \<guillemotright>In1l e2\<guillemotright> E2"
+       by (elim da_elim_cases) simp+
+    from conf_s0 wt_e0 da_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
+      by (rule hyp_e0 [elim_format]) simp
     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)
+    proof (cases "normal s1")
+      case False
+      with eval_e1_e2 have "s2=s1" by auto
+      with conf_s1 error_free_s1 False show ?thesis
+	by auto
+    next
+      case True
+      have s0_s1: "dom (locals (store ((Norm s0)::state))) 
+                    \<union> assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
+      proof -
+	from eval_e0 have 
+	  "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+	  by (rule dom_locals_eval_mono_elim)
+        moreover
+	from eval_e0 True wt_e0 
+	have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
+	  by (rule assigns_if_good_approx') 
+	ultimately show ?thesis by (rule Un_least)
+      qed 
+      show ?thesis
+      proof (cases "the_Bool b")
+	case True
+	with hyp_if have hyp_e1: "PROP ?TypeSafe s1 s2 (In1l e1) (In1 v)" 
+	  by simp
+	from da_e1 s0_s1 True obtain E1' where
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>In1l e1\<guillemotright> E1'"
+	  by - (rule da_weakenE,auto)
+	with conf_s1 wt_e1
+	obtain 
+	  "s2\<Colon>\<preceq>(G, L)"
+	  "(normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1l e1\<succ>In1 v\<Colon>\<preceq>Inl T1)"
+	  "error_free s2"            
+	  by (rule hyp_e1 [elim_format]) (simp add: error_free_s1)
+	moreover
+	from statT  
+	have "G\<turnstile>T1\<preceq>statT"
+	  by auto
+	ultimately show ?thesis
+	  using T wf by auto
+      next
+	case False
+	with hyp_if have hyp_e2: "PROP ?TypeSafe s1 s2 (In1l e2) (In1 v)" 
+	  by simp
+	from da_e2 s0_s1 False obtain E2' where
+	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>In1l e2\<guillemotright> E2'"
+	  by - (rule da_weakenE,auto)
+	with conf_s1 wt_e2
+	obtain 
+	  "s2\<Colon>\<preceq>(G, L)"
+	  "(normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1l e2\<succ>In1 v\<Colon>\<preceq>Inl T2)"
+	  "error_free s2"            
+	  by (rule hyp_e2 [elim_format]) (simp add: error_free_s1)
+	moreover
+	from statT  
+	have "G\<turnstile>T2\<preceq>statT"
+	  by auto
+	ultimately show ?thesis
+	  using T wf by auto
+      qed
+    qed
   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" .
+    case (Call invDeclC a accC' args e mn mode pTs' s0 s1 s2 s3 s3' s4 statT 
+           v vs L accC T A)
+    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 
+                      = 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".
+           "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" .
+       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_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)".
@@ -1896,128 +3042,155 @@
                     T: "T =Inl (resTy statM)" and
         eq_accC_accC': "accC=accC'"
       by (rule wt_elim_cases) fastsimp+
-    from conf_s0 wt_e hyp_e 
+    from Call.prems obtain E where
+      da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+               \<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>In1l e\<guillemotright> E" and
+      da_args: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E \<guillemotright>In3 args\<guillemotright> A" 
+      by (elim da_elim_cases) simp
+    from conf_s0 wt_e da_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
+           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 )
+      by (rule hyp_e [elim_format]) simp
+    { 
+      assume abnormal_s2: "\<not> normal s2"
+      have "set_lvars (locals (store s2)) s4 = s2"
+      proof -
+	from abnormal_s2 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 abnormal_s2 check
+	have eq_s3'_s3: "s3'=s3" 
+	  by (auto simp add: check_method_access_def Let_def)
+	moreover
+	from eq_s3'_s3 abnormal_s2 keep_abrupt eval_methd
+	have "s4=s3'"
+	  by auto
+	ultimately show
+	  "set_lvars (locals (store s2)) s4 = s2"
+	  by (cases s2,cases s3) (simp add: init_lvars_def2)
+      qed
+    } note propagate_abnormal_s2 = this
     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)
+           (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 (Norm s0) =
                 error_free ((set_lvars (locals (store s2))) s4))"
-    proof (cases "normal s2")
+    proof (cases "normal s1")
       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, fastsimp simp add: init_lvars_def2)
-      with False conf_s2 error_free_s2
+      with eval_args have "s2=s1" by auto
+      with False propagate_abnormal_s2 conf_s1 error_free_s1 
       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)
+      note normal_s1 = this
+      obtain A' where 
+	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In3 args\<guillemotright> A'"
+      proof -
+	from eval_e wt_e da_e wf normal_s1
+	have "nrm E \<subseteq>  dom (locals (store s1))"
+	  by (cases rule: da_good_approxE') rules
+	with da_args show ?thesis
+	  by (rule da_weakenE) 
+      qed
+      with conf_s1 wt_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 (rule hyp_args [elim_format]) (simp add: error_free_s1)
+      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 ?thesis
-      proof (cases "a'=Null \<longrightarrow> is_static statM")
+      proof (cases "normal s2")
 	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'"
+	with propagate_abnormal_s2 conf_s2 error_free_s2
+	show ?thesis
 	  by auto
-	ultimately have
-	  "set_lvars (locals (store s2)) s4 
-           = (Some (Xcpt (Std NullPointer)),store s2)"
-	  by (cases s2) (cases s4, fastsimp 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:
+	note normal_s2 = True
+	with normal_s1 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 s3) (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 
+	    by (cases rule: DynT_mheadsE) simp
+	  obtain L' where 
 	   L':"L'=(\<lambda> k. 
                  (case k of
                     EName e
@@ -2028,73 +3201,148 @@
                         | 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,
-                auto 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)
+	    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 (simp add: normal_s2)
+	    apply (cases s2, simp add: L' init_lvars
+	                     cong add: lname.case_cong ename.case_cong)
+	    done
+	  with eq_s3'_s3 
+	  have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
+	  moreover
+	  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,
+                 auto simp add: callee_lcl_def  
+                      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)
+	  moreover
+	  obtain M where 
+	    "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr> 
+	       \<turnstile> dom (locals (store s3')) 
+	          \<guillemotright>In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<guillemotright> M"
+	  proof -
+	    from wf_dynM
+	    obtain M' where
+	      da_body: 
+	      "\<lparr>prg=G, cls=invDeclC
+               ,lcl=callee_lcl invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> (mthd dynM)
+               \<rparr> \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'" and
+              res: "Result \<in> nrm M'"
+	      by (rule wf_mdeclE) rules
+	    from da_body is_static_eq L' have
+	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
+                 \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'"
+	      by (simp add: callee_lcl_def  
+                  cong add: lname.case_cong ename.case_cong)
+	    moreover have "parameters (mthd dynM) \<subseteq>  dom (locals (store s3'))"
+	    proof -
+	      from is_static_eq 
+	      have "(invmode (mthd dynM) e) = (invmode statM e)"
+		by (simp add: invmode_def)
+	      with init_lvars dynM' is_static_eq normal_s2 mode 
+	      have "parameters (mthd dynM) = dom (locals (store s3))"
+		using dom_locals_init_lvars 
+                  [of "mthd dynM" G invDeclC "\<lparr>name=mn,parTs=pTs'\<rparr>" e a vs s2]
+		by simp
+	      also from check
+	      have "dom (locals (store s3)) \<subseteq>  dom (locals (store s3'))"
+		by (simp add:  eq_s3'_s3)
+	      finally show ?thesis .
+	    qed
+	    ultimately obtain M2 where
+	      da:
+	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
+                \<turnstile> dom (locals (store s3')) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M2" and
+              M2: "nrm M' \<subseteq> nrm M2"
+	      by (rule da_weakenE)
+	    from res M2 have "Result \<in> nrm M2"
+	      by blast
+	    moreover from wf_dynM
+	    have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
+	      by (rule wf_mdeclE)
+	    ultimately
+	    obtain M3 where
+	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3')) 
+                     \<guillemotright>\<langle>Body (declclass dynM) (stmt (mbody (mthd dynM)))\<rangle>\<guillemotright> M3"
+	      using da
+	      by (rules intro: da.Body assigned.select_convs)
+	    from _ this [simplified]
+	    show ?thesis
+	      by (rule da.Methd [simplified,elim_format])
+	         (auto intro: dynM')
+	  qed
+	  ultimately 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 (rule hyp_methd [elim_format]) 
+               (simp add: error_free_s3 eq_s3'_s3)
+	  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 )
+	  moreover
+	  have "abrupt s4 \<noteq> Some (Jump Ret)"
+	  proof -
+	    from normal_s2 init_lvars
+	    have "abrupt s3 \<noteq> Some (Jump Ret)"
+	      by (cases s2) (simp add: init_lvars_def2 abrupt_if_def)
+	    with check
+	    have "abrupt s3' \<noteq> Some (Jump Ret)"
+	      by (cases s3) (auto simp add: check_method_access_def Let_def)
+	    with eval_methd
+	    show ?thesis
+	      by (rule Methd_no_jump)
+	  qed
+	  ultimately 
+	  have "(set_lvars (locals (store s2))) s4\<Colon>\<preceq>(G, L)"
+	    using conf_s2 conf_s4
+	    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
+	    by (cases s4) auto
+	  moreover note error_free_s4 T
+	  ultimately 
+	  show ?thesis
+	    by simp
+	qed
       qed
     qed
   next
-    case (Methd D s0 s1 sig v L accC T)
+    case (Methd D s0 s1 sig v L accC T A)
     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)" .
@@ -2106,15 +3354,22 @@
                   \<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 
+    moreover
+    from Methd.prems m have 
+       da_body: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                   \<turnstile> (dom (locals (store ((Norm s0)::state))))
+                       \<guillemotright>In1l (Body (declclass m) (stmt (mbody (mthd m))))\<guillemotright> A"
+      by - (erule da_elim_cases,simp)
+    ultimately
     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)"
+      using hyp [of _ _ "(Inl bodyT)"] conf_s0 
       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" .
+    case (Body D c s0 s1 s2 s3 L accC T A)
+    have eval_init: "G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1" .
+    have eval_c: "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)" .
@@ -2126,12 +3381,35 @@
       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
+    from Body.prems obtain C where
+      da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                   \<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>In1r c\<guillemotright> C" and
+      jmpOk: "jumpNestingOkS {Ret} c" and
+      res: "Result \<in> nrm C"
+      by (elim da_elim_cases) simp
+    note conf_s0
+    moreover from iscls_D 
+    have "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>Init D\<Colon>\<surd>" by auto
+    moreover obtain I where 
+      "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+          \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1r (Init D)\<guillemotright> I"
+      by (auto intro: da_Init [simplified] assigned.select_convs)
+    ultimately obtain
+      conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1:  "error_free s1"
+       by (rule hyp_init [elim_format]) simp
+    obtain C' where da_C': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                             \<turnstile> (dom (locals (store s1)))\<guillemotright>In1r c\<guillemotright> C'"
+               and nrm_C': "nrm C \<subseteq> nrm C'"
+    proof -
+      from eval_init 
+      have "(dom (locals (store ((Norm s0)::state)))) 
+                     \<subseteq> (dom (locals (store s1)))"
+	by (rule dom_locals_eval_mono_elim)
+      with da_c show ?thesis by (rule da_weakenE)
+    qed
+    from conf_s1 wt_c da_C' 
     obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
-      by blast
+      by (rule hyp_c [elim_format]) (simp add: error_free_s1)
     from conf_s2
     have "abupd (absorb Ret) s2\<Colon>\<preceq>(G, L)"
       by (cases s2) (auto intro: conforms_absorb)
@@ -2139,13 +3417,55 @@
     from error_free_s2
     have "error_free (abupd (absorb Ret) s2)"
       by simp
+    moreover have "abrupt (abupd (absorb Ret) s3) \<noteq> Some (Jump Ret)"
+      by (cases s3) (simp add: absorb_def)
+    moreover have "s3=s2"
+    proof -
+      from iscls_D
+      have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init D)\<Colon>\<surd>"
+	by auto
+      from eval_init wf
+      have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+	by - (rule eval_statement_no_jump [OF _ _ _ wt_init],auto)
+      from eval_c _ wt_c wf
+      have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
+	by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
+      moreover 
+      have "s3 =
+                (if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or> 
+                        abrupt s2 = Some (Jump (Cont l))
+                 then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)" .
+      ultimately show ?thesis 
+	by force
+    qed
+    moreover
+    {
+      assume normal_upd_s2:  "normal (abupd (absorb Ret) s2)"
+      have "Result \<in> dom (locals (store s2))"
+      proof -
+	from normal_upd_s2
+	have "normal s2 \<or> abrupt s2 = Some (Jump Ret)"
+	  by (cases s2) (simp add: absorb_def)
+	thus ?thesis
+	proof 
+	  assume "normal s2"
+	  with eval_c wt_c da_C' wf res nrm_C'
+	  show ?thesis
+	    by (cases rule: da_good_approxE') blast
+	next
+	  assume "abrupt s2 = Some (Jump Ret)"
+	  with conf_s2 show ?thesis
+	    by (cases s2) (auto dest: conforms_RetD simp add: dom_def)
+	qed 
+      qed
+    }
     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)
+    show "abupd (absorb Ret) s3\<Colon>\<preceq>(G, L) \<and>
+           (normal (abupd (absorb Ret) s3) \<longrightarrow>
+             G,L,store (abupd (absorb Ret) s3)
              \<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)) "
+          (error_free (Norm s0) = error_free (abupd (absorb Ret) s3)) "
       by (cases s2) (auto intro: conforms_locals)
   next
     case (LVar s vn L accC T)
@@ -2156,9 +3476,9 @@
         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)
+    have conf_fst: "locals s vn \<noteq> None \<longrightarrow> G,s\<turnstile>fst (lvar vn s)\<Colon>\<preceq>vnT"  
+     by (auto elim: conforms_localD [THEN wlconfD]  
+              simp add: lvar_def) 
     moreover
     from conf_s conf_fst vnT 
     have "s\<le>|snd (lvar vn s)\<preceq>vnT\<Colon>\<preceq>(G, L)"
@@ -2169,9 +3489,9 @@
                  (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 
+      by (simp add: lvar_def) 
   next
-    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v L accC' T)
+    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v L accC' T A)
     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" .
@@ -2187,22 +3507,42 @@
                 stat: "stat=is_static f" and
 	           T: "T=(Inl (type f))"
       by (rule wt_elim_cases) (auto simp add: member_is_static_simp)
+    from FVar.prems eq_accC_accC'
+    have da_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>
+                 \<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>In1l e\<guillemotright> A"
+      by (elim da_elim_cases) simp
+    note conf_s0
+    moreover
     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
+    hence "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
+      by simp
+    moreover obtain I where 
+      "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1r (Init statDeclC)\<guillemotright> I"
+      by (auto intro: da_Init [simplified] assigned.select_convs)
+    ultimately 
     obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
-      by auto
-    from conf_s1 wt_e hyp_e
+      by (rule hyp_init [elim_format]) simp
+    obtain A' where 
+      "\<lparr>prg=G, cls=accC, lcl=L\<rparr> \<turnstile> (dom (locals (store s1)))\<guillemotright>In1l e\<guillemotright> A'"
+    proof -
+      from eval_init
+      have "(dom (locals (store ((Norm s0)::state)))) 
+	       \<subseteq> (dom (locals (store s1)))"
+	by (rule dom_locals_eval_mono_elim)
+      with da_e show ?thesis
+	by (rule da_weakenE)
+    qed
+    with 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 force
-    from conf_s1 wt_e error_free_s1 hyp_e
-    have error_free_s2: "error_free s2"
-      by auto
+                  conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC" and
+           error_free_s2: "error_free s2"
+      by (rule hyp_e [elim_format]) (simp add: error_free_s1)
     from fvar 
     have store_s2': "store s2'=store s2"
       by (cases s2) (simp add: fvar_def2)
@@ -2267,7 +3607,7 @@
           (error_free (Norm s0) = error_free s3)"
       by auto
   next
-    case (AVar a e1 e2 i s0 s1 s2 s2' v L accC T)
+    case (AVar a e1 e2 i s0 s1 s2 s2' v L accC T A)
     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)" .
@@ -2276,67 +3616,98 @@
     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
+       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 
+    from AVar.prems obtain E1 where
+      da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                \<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>In1l e1\<guillemotright> E1" and
+      da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>In1l e2\<guillemotright> A" 
+      by (elim da_elim_cases) simp
+    from conf_s0 wt_e1 da_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 
+      by (rule hyp_e1 [elim_format]) simp
     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
+    proof (cases "normal s1")
+      case False
+      moreover
+      from False eval_e2 have eq_s2_s1: "s2=s1" by auto
+      moreover
+      from eq_s2_s1 False have  "\<not> normal s2" by simp
+      then have "snd (avar G i a s2) = s2"
+	by (cases s2) (simp add: avar_def2)
+      with avar have "s2'=s2"
+	by (cases "(avar G i a s2)") simp
+      ultimately show ?thesis
+	using conf_s1 error_free_s1
+	by auto
+    next
+      case True
+      obtain A' where 
+	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In1l e2\<guillemotright> A'"
+      proof -
+	from eval_e1 wt_e1 da_e1 wf True
+	have "nrm E1 \<subseteq> dom (locals (store s1))"
+	  by (cases rule: da_good_approxE') rules
+	with da_e2 show ?thesis
+	  by (rule da_weakenE)
+      qed
+      with conf_s1 wt_e2
+      obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
+	by (rule hyp_e2 [elim_format]) (simp add: error_free_s1)
+      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 ?thesis 
+	by auto
+    qed
   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)
+    case (Cons e es s0 s1 s2 v vs L accC T A)
     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)" .
@@ -2348,27 +3719,50 @@
        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]
+    from Cons.prems obtain E where
+      da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                \<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>In1l e\<guillemotright> E" and
+      da_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E \<guillemotright>In3 es\<guillemotright> A" 
+      by (elim da_elim_cases) simp
+    from conf_s0 wt_e da_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
+      by (rule hyp_e [elim_format]) simp
     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
+      (error_free (Norm s0) = error_free s2)"
+    proof (cases "normal s1")
+      case False
+      with eval_es have "s2=s1" by auto
+      with False conf_s1 error_free_s1
+      show ?thesis
+	by auto
+    next
+      case True
+      obtain A' where 
+	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In3 es\<guillemotright> A'"
+      proof -
+	from eval_e wt_e da_e wf True
+	have "nrm E \<subseteq> dom (locals (store s1))"
+	  by (cases rule: da_good_approxE') rules
+	with da_es show ?thesis
+	  by (rule da_weakenE)
+      qed
+      with conf_s1 wt_es
+      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 (rule hyp_es [elim_format]) (simp add: error_free_s1)
+      moreover
+      from True eval_es conf_v 
+      have conf_v': "G,store s2\<turnstile>v\<Colon>\<preceq>eT"
+	apply clarify
+	apply (rule conf_gext)
+	apply (auto dest: eval_gext)
+	done
+      ultimately show ?thesis using T by simp
+    qed
   qed
   then show ?thesis .
 qed
@@ -2378,48 +3772,66 @@
 
 *} (* dummy text command to break paragraph for latex;
               large paragraphs exhaust memory of debian pdflatex *)
+
+corollary eval_type_soundE [consumes 5]:
+  assumes eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v, s1)"
+  and     conf: "s0\<Colon>\<preceq>(G, L)"
+  and       wt: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>t\<Colon>T"
+  and       da: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile> dom (locals (snd s0)) \<guillemotright>t\<guillemotright> A"
+  and       wf: "wf_prog G"
+  and     elim: "\<lbrakk>s1\<Colon>\<preceq>(G, L); normal s1 \<Longrightarrow> G,L,snd s1\<turnstile>t\<succ>v\<Colon>\<preceq>T; 
+                  error_free s0 = error_free s1\<rbrakk> \<Longrightarrow> P"
+  shows "P"
+using eval wt da wf conf
+by (rule eval_type_sound [elim_format]) (rules intro: elim) 
+
  
 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> 
+ "\<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;
+   \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>In1l e\<guillemotright>A\<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 (drule (4) eval_type_sound)
 apply clarsimp
 done
 
 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> 
+"\<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;
+  \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>In3 es\<guillemotright>A\<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 (drule (4) eval_type_sound)
 apply clarsimp
 done
 
 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>  
+"\<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;
+ \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>In2 v\<guillemotright>A\<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 (drule (4) eval_type_sound)
 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> 
+"\<lbrakk>G\<turnstile>s \<midarrow>c\<rightarrow> s'; wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>;
+ \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>In1r c\<guillemotright>A\<rbrakk> 
  \<Longrightarrow> s'\<Colon>\<preceq>(G,L) \<and> (error_free s \<longrightarrow> error_free s')"
-apply (drule (3) eval_type_sound)
+apply (drule (4) eval_type_sound)
 apply clarsimp
 done
 
 lemma wf_eval_Fin: 
-  assumes wf:    "wf_prog G" and
-          wt_c1: "\<lparr>prg = G, cls = C, lcl = L\<rparr>\<turnstile>In1r c1\<Colon>Inl (PrimT Void)" and
-        conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" and
-        eval_c1: "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1,s1)" and
-        eval_c2: "G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2" and
-            s3: "s3=abupd (abrupt_if (x1\<noteq>None) x1) s2"
+  assumes wf:    "wf_prog G" 
+    and   wt_c1: "\<lparr>prg = G, cls = C, lcl = L\<rparr>\<turnstile>In1r c1\<Colon>Inl (PrimT Void)"
+    and   da_c1: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store (Norm s0)))\<guillemotright>In1r c1\<guillemotright>A"
+    and conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" 
+    and eval_c1: "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1,s1)" 
+    and eval_c2: "G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2" 
+    and      s3: "s3=abupd (abrupt_if (x1\<noteq>None) x1) s2"
   shows "G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<rightarrow> s3"
 proof -
-  from eval_c1 wt_c1 wf conf_s0
+  from eval_c1 wt_c1 da_c1 wf conf_s0
   have "error_free (x1,s1)"
     by (auto dest: eval_type_sound)
   with eval_c1 eval_c2 s3
@@ -2427,4 +3839,209 @@
     by - (rule eval.Fin, auto simp add: error_free_def)
 qed
 
+subsection "Ideas for the future"
+
+text {* In the type soundness proof and the correctness proof of 
+definite assignment we perform induction on the evaluation relation with the 
+further preconditions that the term is welltyped and definitely assigned. During
+the proofs we have to establish the welltypedness and definite assignment of 
+the subterms to be able to apply the induction hypothesis. So large parts of
+both proofs are the same work in propagating welltypedness and definite 
+assignment. So we can derive a new induction rule for induction on the 
+evaluation of a wellformed term, were these propagations is already done, once
+and forever. 
+Then we can do the proofs with this rule and can enjoy the time we have saved.
+Here is a first and incomplete sketch of such a rule.*}
+theorem wellformed_eval_induct [consumes 4, case_names Abrupt Skip Expr Lab 
+                                Comp If]:
+  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      da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>t\<guillemotright>A"
+   and      wf: "wf_prog G" 
+   and  abrupt: "\<And> s t abr L accC T A. 
+                  \<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T; 
+                   \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store (Some abr,s)))\<guillemotright>t\<guillemotright>A
+                  \<rbrakk> \<Longrightarrow> P L accC (Some abr, s) t (arbitrary3 t) (Some abr, s)"
+   and    skip: "\<And> s L accC. P L accC (Norm s) \<langle>Skip\<rangle>\<^sub>s \<diamondsuit> (Norm s)"
+   and    expr: "\<And> e s0 s1 v L accC eT E.
+                 \<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-eT;
+                  \<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                     \<turnstile>dom (locals (store ((Norm s0)::state)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright>E;
+                  P L accC (Norm s0) \<langle>e\<rangle>\<^sub>e \<lfloor>v\<rfloor>\<^sub>e s1\<rbrakk> 
+                 \<Longrightarrow>  P L accC (Norm s0) \<langle>Expr e\<rangle>\<^sub>s \<diamondsuit> s1"
+   and     lab: "\<And> c l s0 s1 L accC C.
+                 \<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>;
+                  \<lparr>prg=G,cls=accC, lcl=L\<rparr>
+                     \<turnstile>dom (locals (store ((Norm s0)::state)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>C;
+                  P L accC (Norm s0) \<langle>c\<rangle>\<^sub>s \<diamondsuit> s1\<rbrakk>
+                  \<Longrightarrow> P L accC (Norm s0) \<langle>l\<bullet> c\<rangle>\<^sub>s \<diamondsuit> (abupd (absorb l) s1)"
+   and    comp: "\<And> c1 c2 s0 s1 s2 L accC C1.
+                 \<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<rightarrow> s1;G\<turnstile>s1 \<midarrow>c2 \<rightarrow> s2;
+                  \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>;
+                  \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c2\<Colon>\<surd>;
+                  \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+                     dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1;
+                  P L accC (Norm s0) \<langle>c1\<rangle>\<^sub>s \<diamondsuit> s1;
+                  \<And> Q. \<lbrakk>normal s1; 
+                         \<And> C2.\<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+                                  \<turnstile>dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2;
+                               P L accC s1 \<langle>c2\<rangle>\<^sub>s \<diamondsuit> s2\<rbrakk> \<Longrightarrow> Q
+                        \<rbrakk> \<Longrightarrow> Q 
+                  \<rbrakk>\<Longrightarrow> P L accC (Norm s0) \<langle>c1;; c2\<rangle>\<^sub>s \<diamondsuit> s2" 
+    and    if: "\<And> b c1 c2 e s0 s1 s2 L accC E.
+                \<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
+                 G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2;
+                 \<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>;
+                 \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+                     dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E;
+                 P L accC (Norm s0) \<langle>e\<rangle>\<^sub>e \<lfloor>b\<rfloor>\<^sub>e s1;
+                 \<And> Q. \<lbrakk>normal s1;
+                        \<And> C. \<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))
+                                   \<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> C;
+                              P L accC s1 \<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s \<diamondsuit> s2
+                              \<rbrakk> \<Longrightarrow> Q
+                       \<rbrakk> \<Longrightarrow> Q
+                \<rbrakk> \<Longrightarrow> P L accC (Norm s0) \<langle>If(e) c1 Else c2\<rangle>\<^sub>s \<diamondsuit> s2" 
+   shows "P L accC s0 t v s1"  
+proof -
+  note inj_term_simps [simp]
+  from eval 
+  show "\<And> L accC T A. \<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T;
+                       \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>t\<guillemotright>A\<rbrakk>  
+        \<Longrightarrow> P L accC s0 t v s1" (is "PROP ?Hyp s0 t v s1")
+  proof (induct)
+    case Abrupt with abrupt show ?case .
+  next
+    case Skip from skip show ?case by simp
+  next
+    case (Expr e s0 s1 v L accC T A) 
+    from Expr.prems obtain eT where 
+      "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
+      by (elim wt_elim_cases) 
+    moreover
+    from Expr.prems obtain E where
+      "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store ((Norm s0)::state)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright>E"
+      by (elim da_elim_cases) simp
+    moreover from calculation
+    have "P L accC (Norm s0) \<langle>e\<rangle>\<^sub>e \<lfloor>v\<rfloor>\<^sub>e s1"
+      by (rule Expr.hyps) 
+    ultimately show ?case
+      by (rule expr)
+  next
+    case (Lab c l s0 s1 L accC T A)
+    from Lab.prems 
+    have "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
+      by (elim wt_elim_cases)
+    moreover
+    from Lab.prems obtain C where
+      "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store ((Norm s0)::state)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>C"
+      by (elim da_elim_cases) simp
+    moreover from calculation
+    have "P L accC (Norm s0) \<langle>c\<rangle>\<^sub>s \<diamondsuit> s1"
+      by (rule  Lab.hyps)  
+    ultimately show ?case
+      by (rule lab)
+  next
+    case (Comp c1 c2 s0 s1 s2 L accC T A) 
+    have eval_c1: "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1" .
+    have eval_c2: "G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2" .
+    from Comp.prems 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 (elim wt_elim_cases) 
+    from Comp.prems
+    obtain C1 C2
+      where da_c1: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile> 
+                      dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and 
+            da_c2: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>  nrm C1 \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2" 
+      by (elim da_elim_cases) simp
+    from wt_c1 da_c1
+    have P_c1: "P L accC (Norm s0) \<langle>c1\<rangle>\<^sub>s \<diamondsuit> s1"
+      by (rule Comp.hyps)
+    {
+      fix Q
+      assume normal_s1: "normal s1"
+      assume elim: "\<And> C2'. 
+                    \<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s1))\<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright>C2';
+                       P L accC s1 \<langle>c2\<rangle>\<^sub>s \<diamondsuit> s2\<rbrakk> \<Longrightarrow> Q"
+      have Q
+      proof -
+	obtain C2' where 
+	  da: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
+	proof -
+	  from eval_c1 wt_c1 da_c1 wf normal_s1
+	  have "nrm C1 \<subseteq> dom (locals (store s1))"
+	    by (cases rule: da_good_approxE') rules
+	  with da_c2 show ?thesis
+	    by (rule da_weakenE)
+	qed
+	with wt_c2 have "P L accC s1 \<langle>c2\<rangle>\<^sub>s \<diamondsuit> s2"
+	  by (rule Comp.hyps)
+	with da show ?thesis
+	  using elim by rules
+      qed
+    }
+    with eval_c1 eval_c2 wt_c1 wt_c2 da_c1 P_c1 
+    show ?case
+      by (rule comp) rules+
+  next
+    case (If b c1 c2 e s0 s1 s2 L accC T A)
+    have eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1" .
+    have eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2" .
+    from If.prems
+    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 (elim wt_elim_cases) (auto split add: split_if)
+    from If.prems obtain E C where
+      da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store ((Norm s0)::state))) 
+                                       \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
+      da_then_else: 
+      "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+         (dom (locals (store ((Norm s0)::state))) \<union> assigns_if (the_Bool b) e)
+          \<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> C"
+      by (elim da_elim_cases) (cases "the_Bool b",auto)
+    from wt_e da_e
+    have P_e: "P L accC (Norm s0) \<langle>e\<rangle>\<^sub>e \<lfloor>b\<rfloor>\<^sub>e s1"
+      by (rule If.hyps)
+    {
+      fix Q
+      assume normal_s1: "normal s1"
+      assume elim: "\<And> C. \<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))
+                                   \<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> C;
+                              P L accC s1 \<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s \<diamondsuit> s2
+                              \<rbrakk> \<Longrightarrow> Q"
+      have Q
+      proof -
+	obtain C' where
+	  da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+                (dom (locals (store s1)))\<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s \<guillemotright> C'"
+	proof -
+	  from eval_e have 
+	    "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+	    by (rule dom_locals_eval_mono_elim)
+          moreover
+	  from eval_e normal_s1 wt_e 
+	  have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+	    by (rule assigns_if_good_approx')
+	  ultimately 
+	  have "dom (locals (store ((Norm s0)::state))) 
+            \<union> assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+	    by (rule Un_least)
+	  with da_then_else show ?thesis
+	    by (rule da_weakenE)
+	qed
+	with wt_then_else
+	have "P L accC s1 \<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s \<diamondsuit> s2"
+	  by (rule If.hyps)
+	with da show ?thesis using elim by rules
+      qed
+    }
+    with eval_e eval_then_else wt_e wt_then_else da_e P_e
+    show ?case
+      by (rule if) rules+
+  next
+    oops
+
 end
--- a/src/HOL/Bali/Value.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/Value.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -9,15 +9,15 @@
 
 theory Value = Type:
 
-typedecl loc            (* locations, i.e. abstract references on objects *)
+typedecl loc            --{* locations, i.e. abstract references on objects *}
 arities	 loc :: "type"
 
 datatype val
-	= Unit          (* dummy result value of void methods *)
-	| Bool bool     (* Boolean value *)
-	| Intg int      (* integer value *)
-	| Null          (* null reference *)
-	| Addr loc      (* addresses, i.e. locations of objects *)
+	= Unit          --{* dummy result value of void methods *}
+	| Bool bool     --{* Boolean value *}
+	| Intg int      --{* integer value *}
+	| Null          --{* null reference *}
+	| Addr loc      --{* addresses, i.e. locations of objects *}
 
 
 translations "val" <= (type) "Term.val"
@@ -33,8 +33,8 @@
 types	dyn_ty      = "loc \<Rightarrow> ty option"
 consts
   typeof        :: "dyn_ty \<Rightarrow> val \<Rightarrow> ty option"
-  defpval       :: "prim_ty \<Rightarrow> val"      (* default value for primitive types *)
-  default_val   :: "     ty \<Rightarrow> val"      (* default value for all types *)
+  defpval       :: "prim_ty \<Rightarrow> val"  --{* default value for primitive types *}
+  default_val   :: "     ty \<Rightarrow> val"  --{* default value for all types *}
 
 primrec "typeof dt  Unit    = Some (PrimT Void)"
 	"typeof dt (Bool b) = Some (PrimT Boolean)"
--- a/src/HOL/Bali/WellForm.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/WellForm.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -5,8 +5,7 @@
 *)
 
 header {* Well-formedness of Java programs *}
-
-theory WellForm = WellType:
+theory WellForm = DefiniteAssignment:
 
 text {*
 For static checks on expressions and statements, see WellType.thy
@@ -81,6 +80,20 @@
 \end{itemize}
 *}
 
+constdefs callee_lcl:: "qtname \<Rightarrow> sig \<Rightarrow> methd \<Rightarrow> lenv"
+"callee_lcl C sig m 
+ \<equiv> \<lambda> k. (case k of
+            EName e 
+            \<Rightarrow> (case e of 
+                  VNam v 
+                  \<Rightarrow>(table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
+                | Res \<Rightarrow> Some (resTy m))
+	  | This \<Rightarrow> if is_static m then None else Some (Class C))"
+
+constdefs parameters :: "methd \<Rightarrow> lname set"
+"parameters m \<equiv>  set (map (EName \<circ> VNam) (pars m)) 
+                  \<union> (if (static m) then {} else {This})"
+
 constdefs
   wf_mdecl :: "prog \<Rightarrow> qtname \<Rightarrow> mdecl \<Rightarrow> bool"
  "wf_mdecl G C \<equiv> 
@@ -89,17 +102,33 @@
           unique (lcls (mbody m)) \<and> 
           (\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T) \<and> 
 	  (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and>
+          jumpNestingOkS {Ret} (stmt (mbody m)) \<and> 
           is_class G C \<and>
-          \<lparr>prg=G,cls=C,
-           lcl=\<lambda> k. 
-               (case k of
-                  EName e 
-                  \<Rightarrow> (case e of 
-                        VNam v 
-                        \<Rightarrow>(table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
-                      | Res \<Rightarrow> Some (resTy m))
-	        | This \<Rightarrow> if is_static m then None else Some (Class C))
-          \<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>"
+          \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd> \<and>
+          (\<exists> A. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr> 
+                \<turnstile> parameters m \<guillemotright>\<langle>stmt (mbody m)\<rangle>\<guillemotright> A 
+               \<and> Result \<in> nrm A)"
+
+lemma callee_lcl_VNam_simp [simp]:
+"callee_lcl C sig m (EName (VNam v)) 
+  = (table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v"
+by (simp add: callee_lcl_def)
+ 
+lemma callee_lcl_Res_simp [simp]:
+"callee_lcl C sig m (EName Res) = Some (resTy m)" 
+by (simp add: callee_lcl_def)
+
+lemma callee_lcl_This_simp [simp]:
+"callee_lcl C sig m (This) = (if is_static m then None else Some (Class C))" 
+by (simp add: callee_lcl_def)
+
+lemma callee_lcl_This_static_simp:
+"is_static m \<Longrightarrow> callee_lcl C sig m (This) = None"
+by simp
+
+lemma callee_lcl_This_not_static_simp:
+"\<not> is_static m \<Longrightarrow> callee_lcl C sig m (This) = Some (Class C)"
+by simp
 
 lemma wf_mheadI: 
 "\<lbrakk>length (parTs sig) = length (pars m); \<forall>T\<in>set (parTs sig). is_acc_type G P T;
@@ -113,23 +142,31 @@
   wf_mhead G (pid C) sig (mhead m); unique (lcls (mbody m));  
   (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None); 
   \<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T;
+  jumpNestingOkS {Ret} (stmt (mbody m));
   is_class G C;
-  \<lparr>prg=G,cls=C,
-   lcl=\<lambda> k. 
-       (case k of
-          EName e 
-          \<Rightarrow> (case e of 
-                VNam v 
-                \<Rightarrow> (table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
-              | Res \<Rightarrow> Some (resTy m))
-        | This \<Rightarrow> if is_static m then None else Some (Class C))
-  \<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>
+  \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>;
+  (\<exists> A. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr> \<turnstile> parameters m \<guillemotright>\<langle>stmt (mbody m)\<rangle>\<guillemotright> A
+        \<and> Result \<in> nrm A)
   \<rbrakk> \<Longrightarrow>  
   wf_mdecl G C (sig,m)"
 apply (unfold wf_mdecl_def)
 apply simp
 done
 
+lemma wf_mdeclE [consumes 1]:  
+  "\<lbrakk>wf_mdecl G C (sig,m); 
+    \<lbrakk>wf_mhead G (pid C) sig (mhead m); unique (lcls (mbody m));  
+     \<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None; 
+     \<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T;
+     jumpNestingOkS {Ret} (stmt (mbody m));
+     is_class G C;
+     \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>;
+   (\<exists> A. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile> parameters m \<guillemotright>\<langle>stmt (mbody m)\<rangle>\<guillemotright> A
+        \<and> Result \<in> nrm A)
+    \<rbrakk> \<Longrightarrow> P
+  \<rbrakk> \<Longrightarrow> P"
+by (unfold wf_mdecl_def) simp
+
 
 lemma wf_mdeclD1: 
 "wf_mdecl G C (sig,m) \<Longrightarrow>  
@@ -137,20 +174,13 @@
   (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and> 
   (\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T)"
 apply (unfold wf_mdecl_def)
-apply auto
+apply simp
 done
 
 lemma wf_mdecl_bodyD: 
 "wf_mdecl G C (sig,m) \<Longrightarrow>  
- (\<exists>T. \<lparr>prg=G,cls=C,
-       lcl = \<lambda> k. 
-         (case k of
-            EName e 
-            \<Rightarrow> (case e of 
-                VNam v \<Rightarrow> (table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
-                | Res  \<Rightarrow> Some (resTy m))
-          | This \<Rightarrow> if 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))"
+ (\<exists>T. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>Body C (stmt (mbody m))\<Colon>-T \<and> 
+      G\<turnstile>T\<preceq>(resTy m))"
 apply (unfold wf_mdecl_def)
 apply clarify
 apply (rule_tac x="(resTy m)" in exI)
@@ -315,7 +345,9 @@
       	                             \<not> is_static cm \<and>
                                      accmodi im \<le> accmodi cm))) \<and>
       (\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and> 
-      (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and> 
+      (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and>
+      jumpNestingOkS {} (init c) \<and>
+      (\<exists> A. \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<guillemotright> A) \<and>
       \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd> \<and> ws_cdecl G C (super c) \<and>
       (C \<noteq> Object \<longrightarrow> 
             (is_acc_class G (pid C) (super c) \<and>
@@ -362,6 +394,35 @@
               entails (\<lambda> newF oldF. accmodi oldF \<le> access newF))))"
 *)
 
+lemma wf_cdeclE [consumes 1]: 
+ "\<lbrakk>wf_cdecl G (C,c);
+   \<lbrakk>\<not>is_iface G C;
+    (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
+        (\<forall>s. \<forall> im \<in> imethds G I s.
+      	    (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
+      	                             \<not> is_static cm \<and>
+                                     accmodi im \<le> accmodi cm))); 
+      \<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f; unique (cfields c); 
+      \<forall>m\<in>set (methods c). wf_mdecl G C m; unique (methods c);
+      jumpNestingOkS {} (init c);
+      \<exists> A. \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<guillemotright> A;
+      \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd>; 
+      ws_cdecl G C (super c); 
+      (C \<noteq> Object \<longrightarrow> 
+            (is_acc_class G (pid C) (super c) \<and>
+            (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
+             entails (\<lambda> new. \<forall> old sig. 
+                       (G,sig\<turnstile>new overrides\<^sub>S old 
+                        \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
+                             accmodi old \<le> accmodi new \<and>
+      	                     \<not>is_static old)) \<and>
+                       (G,sig\<turnstile>new hides old 
+                         \<longrightarrow> (accmodi old \<le> accmodi new \<and>
+      	                      is_static old)))) 
+            ))\<rbrakk> \<Longrightarrow> P
+  \<rbrakk> \<Longrightarrow> P"
+by (unfold wf_cdecl_def) simp
+
 lemma wf_cdecl_unique: 
 "wf_cdecl G (C,c) \<Longrightarrow> unique (cfields c) \<and> unique (methods c)"
 apply (unfold wf_cdecl_def)
@@ -535,7 +596,7 @@
 	"wf_cdecl G ObjectC = (\<not>is_iface G Object \<and> Ball (set Object_mdecls)
   (wf_mdecl G Object) \<and> unique Object_mdecls)"
 apply (unfold wf_cdecl_def ws_cdecl_def ObjectC_def)
-apply (simp (no_asm))
+apply (auto intro: da.Skip)
 done
 
 lemma Object_is_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_class G Object"
@@ -1691,19 +1752,7 @@
                intro: order_trans)
   qed
 qed
- 	  
-(* ### Probleme: Die tollen methd_subcls_cases Lemma wird warscheinlich
-  kaum gebraucht: 
-Redundanz: stat_overrides.Direct old declared in declclass old folgt aus
-           member of 
-   Problem: Predikate wie overrides, sind global üper die Hierarchie hinweg
-            definiert, aber oft barucht man eben zusätlich Induktion
-            : von super c auf C; Dann ist aber auss dem Kontext
-            die Unterscheidung in die 5 fälle overkill,
-            da man dann warscheinlich meistens eh in einem speziellen
-            Fall kommt (durch die Hypothesen)
-*)
-    
+ 	      
 (* local lemma *)
 ML {* bind_thm("bexI'",permute_prems 0 1 bexI) *}
 ML {* bind_thm("ballE'",permute_prems 1 1 ballE) *}
@@ -2419,7 +2468,7 @@
 apply (auto split del: split_if_asm simp del: snd_conv 
             simp add: is_acc_class_def is_acc_type_def)
 apply    (erule typeof_empty_is_type)
-apply   (frule (1) wf_prog_cdecl [THEN wf_cdecl_supD], rotate_tac -1, 
+apply   (frule (1) wf_prog_cdecl [THEN wf_cdecl_supD], 
         force simp del: snd_conv, clarsimp simp add: is_acc_class_def)
 apply  (drule (1) max_spec2mheads [THEN conjunct1, THEN mheadsD])
 apply  (drule_tac [2] accfield_fields) 
@@ -2476,16 +2525,7 @@
 "\<lbrakk>methd G C sig = Some m; wf_prog G;  
   class G C = Some c\<rbrakk> \<Longrightarrow>  
  \<exists>T. \<lparr>prg=G,cls=(declclass m),
-      lcl=\<lambda> k. 
-          (case k of
-             EName e 
-             \<Rightarrow> (case e of 
-                   VNam v 
-                   \<Rightarrow> (table_of (lcls (mbody (mthd m)))
-                                ((pars (mthd m))[\<mapsto>](parTs sig))) v
-                 | Res \<Rightarrow> Some (resTy (mthd m)))
-           | This \<Rightarrow> if is_static m then None else Some (Class (declclass m)))
-     \<rparr>\<turnstile> Methd C sig\<Colon>-T \<and> G\<turnstile>T\<preceq>resTy m"
+      lcl=callee_lcl (declclass m) sig (mthd m)\<rparr>\<turnstile> Methd C sig\<Colon>-T \<and> G\<turnstile>T\<preceq>resTy m"
 apply (frule (2) methd_wf_mdecl, clarify)
 apply (force dest!: wf_mdecl_bodyD intro!: wt.Methd)
 done
--- a/src/HOL/Bali/WellType.thy	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/WellType.thy	Thu Oct 31 18:27:10 2002 +0100
@@ -29,12 +29,12 @@
 *}
 
 types	lenv
-	= "(lname, ty) table"   (* local variables, including This and Result *)
+	= "(lname, ty) table"  --{* local variables, including This and Result*}
 
 record env = 
-         prg:: "prog"    (* program *)
-         cls:: "qtname"  (* current package and class name *)
-         lcl:: "lenv"    (* local environment *)     
+         prg:: "prog"    --{* program *}
+         cls:: "qtname"  --{* current package and class name *}
+         lcl:: "lenv"    --{* local environment *}     
   
 translations
   "lenv" <= (type) "(lname, ty) table"
@@ -45,8 +45,8 @@
 
 
 syntax
-  pkg :: "env \<Rightarrow> pname" (* select the current package from an environment *)
-translations
+  pkg :: "env \<Rightarrow> pname" --{* select the current package from an environment *}
+translations 
   "pkg e" == "pid (cls e)"
 
 section "Static overloading: maximally specific methods "
@@ -54,7 +54,7 @@
 types
   emhead = "ref_ty \<times> mhead"
 
-(* Some mnemotic selectors for emhead *)
+--{* Some mnemotic selectors for emhead *}
 constdefs 
   "declrefT" :: "emhead \<Rightarrow> ref_ty"
   "declrefT \<equiv> fst"
@@ -106,32 +106,24 @@
 "mheads G S (ClassT C) = cmheads G S C"
 "mheads G S (ArrayT T) = accObjectmheads G S (ArrayT T)"
 
-
-(* more detailed than necessary for type-safety, see below. *)
 constdefs
-  (* applicable methods, cf. 15.11.2.1 *)
+  --{* applicable methods, cf. 15.11.2.1 *}
   appl_methds   :: "prog \<Rightarrow> qtname \<Rightarrow>  ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
  "appl_methds G S rt \<equiv> \<lambda> sig. 
       {(mh,pTs') |mh pTs'. mh \<in> mheads G S rt \<lparr>name=name sig,parTs=pTs'\<rparr> \<and> 
                            G\<turnstile>(parTs sig)[\<preceq>]pTs'}"
 
-  (* more specific methods, cf. 15.11.2.2 *)
+  --{* more specific methods, cf. 15.11.2.2 *}
   more_spec     :: "prog \<Rightarrow> emhead \<times> ty list \<Rightarrow> emhead \<times> ty list \<Rightarrow> bool"
  "more_spec G \<equiv> \<lambda>(mh,pTs). \<lambda>(mh',pTs'). G\<turnstile>pTs[\<preceq>]pTs'"
 (*more_spec G \<equiv>\<lambda>((d,h),pTs). \<lambda>((d',h'),pTs'). G\<turnstile>RefT d\<preceq>RefT d'\<and>G\<turnstile>pTs[\<preceq>]pTs'*)
 
-  (* maximally specific methods, cf. 15.11.2.2 *)
+  --{* maximally specific methods, cf. 15.11.2.2 *}
    max_spec     :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
 
  "max_spec G S rt sig \<equiv>{m. m \<in>appl_methds G S rt sig \<and>
 		      (\<forall>m'\<in>appl_methds G S rt sig. more_spec G m' m \<longrightarrow> m'=m)}"
-(*
-rules (* all properties of more_spec, appl_methods and max_spec we actually need
-         these can easily be proven from the above definitions *)
 
-max_spec2mheads "max_spec G S rt (mn, pTs) = insert (mh, pTs') A \<Longrightarrow>
-      mh\<in>mheads G S rt (mn, pTs') \<and> G\<turnstile>pTs[\<preceq>]pTs'"
-*)
 
 lemma max_spec2appl_meths: 
   "x \<in> max_spec G S T sig \<Longrightarrow> x \<in> appl_methds G S T sig" 
@@ -180,6 +172,8 @@
 apply (clarsimp simp add: invmode_IntVir_eq)
 done
 
+section "Typing for unary operations"
+
 consts unop_type :: "unop \<Rightarrow> prim_ty"
 primrec 
 "unop_type UPlus   = Integer"
@@ -194,6 +188,8 @@
 "wt_unop UBitNot t = (t = PrimT Integer)"
 "wt_unop UNot    t = (t = PrimT Boolean)"
 
+section "Typing for binary operations"
+
 consts binop_type :: "binop \<Rightarrow> prim_ty"
 primrec
 "binop_type Mul      = Integer"     
@@ -244,6 +240,7 @@
 "wt_binop G CondAnd t1 t2 = ((t1 = PrimT Boolean) \<and> (t2 = PrimT Boolean))"
 "wt_binop G CondOr  t1 t2 = ((t1 = PrimT Boolean) \<and> (t2 = PrimT Boolean))"
 
+section "Typing for terms"
 
 types tys  =        "ty + ty list"
 translations
@@ -295,23 +292,22 @@
 
 translations
 	"E\<turnstile>t\<Colon> T" == "E,empty_dt\<Turnstile>t\<Colon> T"
-	"E\<turnstile>s\<Colon>\<surd>" == "E,empty_dt\<Turnstile>s\<Colon>\<surd>"
-	"E\<turnstile>e\<Colon>-T" == "E,empty_dt\<Turnstile>e\<Colon>-T"
-	"E\<turnstile>e\<Colon>=T" == "E,empty_dt\<Turnstile>e\<Colon>=T"
-	"E\<turnstile>e\<Colon>\<doteq>T" == "E,empty_dt\<Turnstile>e\<Colon>\<doteq>T"
-
+        "E\<turnstile>s\<Colon>\<surd>"  == "E\<turnstile>In1r s\<Colon>Inl (PrimT Void)"
+	"E\<turnstile>e\<Colon>-T" == "E\<turnstile>In1l e\<Colon>Inl T"
+	"E\<turnstile>e\<Colon>=T" == "E\<turnstile>In2  e\<Colon>Inl T"
+	"E\<turnstile>e\<Colon>\<doteq>T" == "E\<turnstile>In3  e\<Colon>Inr T"
 
 
 inductive wt intros 
 
 
-(* well-typed statements *)
+--{* well-typed statements *}
 
   Skip:					"E,dt\<Turnstile>Skip\<Colon>\<surd>"
 
   Expr:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>Expr e\<Colon>\<surd>"
-  (* cf. 14.6 *)
+  --{* cf. 14.6 *}
   Lab:  "E,dt\<Turnstile>c\<Colon>\<surd> \<Longrightarrow>                   
                                          E,dt\<Turnstile>l\<bullet> c\<Colon>\<surd>" 
 
@@ -319,61 +315,62 @@
 	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>c1;; c2\<Colon>\<surd>"
 
-  (* cf. 14.8 *)
+  --{* cf. 14.8 *}
   If:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
 	  E,dt\<Turnstile>c1\<Colon>\<surd>;
 	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>If(e) c1 Else c2\<Colon>\<surd>"
 
-  (* cf. 14.10 *)
+  --{* cf. 14.10 *}
   Loop:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
 	  E,dt\<Turnstile>c\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>l\<bullet> While(e) c\<Colon>\<surd>"
-  (* cf. 14.13, 14.15, 14.16 *)
-  Do:                                   "E,dt\<Turnstile>Do jump\<Colon>\<surd>"
+  --{* cf. 14.13, 14.15, 14.16 *}
+  Jmp:                                   "E,dt\<Turnstile>Jmp jump\<Colon>\<surd>"
 
-  (* cf. 14.16 *)
+  --{* cf. 14.16 *}
   Throw: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class tn;
 	  prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>Throw e\<Colon>\<surd>"
-  (* cf. 14.18 *)
+  --{* cf. 14.18 *}
   Try:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable;
 	  lcl E (VName vn)=None; E \<lparr>lcl := lcl E(VName vn\<mapsto>Class tn)\<rparr>,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk>
           \<Longrightarrow>
 					 E,dt\<Turnstile>Try c1 Catch(tn vn) c2\<Colon>\<surd>"
 
-  (* cf. 14.18 *)
+  --{* cf. 14.18 *}
   Fin:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>c1 Finally c2\<Colon>\<surd>"
 
   Init:	"\<lbrakk>is_class (prg E) C\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>Init C\<Colon>\<surd>"
-  (* Init is created on the fly during evaluation (see Eval.thy). The class
-   * isn't necessarily accessible from the points Init is called. Therefor
-   * we only demand is_class and not is_acc_class here 
-   *)
+  --{* @{term Init} is created on the fly during evaluation (see Eval.thy). 
+     The class isn't necessarily accessible from the points @{term Init} 
+     is called. Therefor we only demand @{term is_class} and not 
+     @{term is_acc_class} here. 
+   *}
 
-(* well-typed expressions *)
+--{* well-typed expressions *}
 
-  (* cf. 15.8 *)
+  --{* cf. 15.8 *}
   NewC:	"\<lbrakk>is_acc_class (prg E) (pkg E) C\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>NewC C\<Colon>-Class C"
-  (* cf. 15.9 *)
+  --{* cf. 15.9 *}
   NewA:	"\<lbrakk>is_acc_type (prg E) (pkg E) T;
 	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>New T[i]\<Colon>-T.[]"
 
-  (* cf. 15.15 *)
+  --{* cf. 15.15 *}
   Cast:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T; is_acc_type (prg E) (pkg E) T';
 	  prg E\<turnstile>T\<preceq>? T'\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>Cast T' e\<Colon>-T'"
 
-  (* cf. 15.19.2 *)
+  --{* cf. 15.19.2 *}
   Inst:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT T; is_acc_type (prg E) (pkg E) (RefT T');
 	  prg E\<turnstile>RefT T\<preceq>? RefT T'\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>e InstOf T'\<Colon>-PrimT Boolean"
 
-  (* cf. 15.7.1 *)
+  --{* cf. 15.7.1 *}
   Lit:	"\<lbrakk>typeof dt x = Some T\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>Lit x\<Colon>-T"
 
@@ -386,28 +383,28 @@
            \<Longrightarrow>
 	   E,dt\<Turnstile>BinOp binop e1 e2\<Colon>-T"
   
-  (* cf. 15.10.2, 15.11.1 *)
+  --{* cf. 15.10.2, 15.11.1 *}
   Super: "\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object;
 	  class (prg E) C = Some c\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>Super\<Colon>-Class (super c)"
 
-  (* cf. 15.13.1, 15.10.1, 15.12 *)
+  --{* cf. 15.13.1, 15.10.1, 15.12 *}
   Acc:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>Acc va\<Colon>-T"
 
-  (* cf. 15.25, 15.25.1 *)
+  --{* cf. 15.25, 15.25.1 *}
   Ass:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T; va \<noteq> LVar This;
 	  E,dt\<Turnstile>v \<Colon>-T';
 	  prg E\<turnstile>T'\<preceq>T\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>va:=v\<Colon>-T'"
 
-  (* cf. 15.24 *)
+  --{* cf. 15.24 *}
   Cond:	"\<lbrakk>E,dt\<Turnstile>e0\<Colon>-PrimT Boolean;
 	  E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2;
 	  prg E\<turnstile>T1\<preceq>T2 \<and> T = T2  \<or>  prg E\<turnstile>T2\<preceq>T1 \<and> T = T1\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>e0 ? e1 : e2\<Colon>-T"
 
-  (* cf. 15.11.1, 15.11.2, 15.11.3 *)
+  --{* cf. 15.11.1, 15.11.2, 15.11.3 *}
   Call:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT;
 	  E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;
 	  max_spec (prg E) (cls E) statT \<lparr>name=mn,parTs=pTs\<rparr> 
@@ -419,48 +416,51 @@
 	  methd (prg E) C sig = Some m;
 	  E,dt\<Turnstile>Body (declclass m) (stmt (mbody (mthd m)))\<Colon>-T\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>Methd C sig\<Colon>-T"
- (* The class C is the dynamic class of the method call (cf. Eval.thy). 
-  * It hasn't got to be directly accessible from the current package (pkg E). 
-  * Only the static class must be accessible (enshured indirectly by Call). 
-  * Note that l is just a dummy value. It is only used in the smallstep 
-  * semantics. To proof typesafety directly for the smallstep semantics 
-  * we would have to assume conformance of l here!
-  *)
+ --{* The class @{term C} is the dynamic class of the method call 
+    (cf. Eval.thy). 
+    It hasn't got to be directly accessible from the current package 
+    @{term "(pkg E)"}. 
+    Only the static class must be accessible (enshured indirectly by 
+    @{term Call}). 
+    Note that l is just a dummy value. It is only used in the smallstep 
+    semantics. To proof typesafety directly for the smallstep semantics 
+    we would have to assume conformance of l here!
+  *}
 
   Body:	"\<lbrakk>is_class (prg E) D;
 	  E,dt\<Turnstile>blk\<Colon>\<surd>;
 	  (lcl E) Result = Some T;
           is_type (prg E) T\<rbrakk> \<Longrightarrow>
    					 E,dt\<Turnstile>Body D blk\<Colon>-T"
-  (* the class D implementing the method must not directly be accessible 
-   * from the current package (pkg E), but can also be indirectly accessible 
-   * due to inheritance (enshured in Call)
-   * The result type hasn't got to be accessible in Java! (If it is not 
-   * accessible you can only assign it to Object).
-   * For dummy value l see rule Methd. 
-   *)
+--{* The class @{term D} implementing the method must not directly be 
+     accessible  from the current package @{term "(pkg E)"}, but can also 
+     be indirectly accessible due to inheritance (enshured in @{term Call})
+    The result type hasn't got to be accessible in Java! (If it is not 
+    accessible you can only assign it to Object).
+    For dummy value l see rule @{term Methd}. 
+   *}
 
-(* well-typed variables *)
+--{* well-typed variables *}
 
-  (* cf. 15.13.1 *)
+  --{* cf. 15.13.1 *}
   LVar:	"\<lbrakk>lcl E vn = Some T; is_acc_type (prg E) (pkg E) T\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>LVar vn\<Colon>=T"
-  (* cf. 15.10.1 *)
+  --{* cf. 15.10.1 *}
   FVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; 
 	  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 *)
+  --{* cf. 15.12 *}
   AVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T.[]; 
 	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>e.[i]\<Colon>=T"
 
 
-(* well-typed expression lists *)
+--{* well-typed expression lists *}
 
-  (* cf. 15.11.??? *)
+  --{* cf. 15.11.??? *}
   Nil:					"E,dt\<Turnstile>[]\<Colon>\<doteq>[]"
 
-  (* cf. 15.11.??? *)
+  --{* cf. 15.11.??? *}
   Cons:	"\<lbrakk>E,dt\<Turnstile>e \<Colon>-T;
 	  E,dt\<Turnstile>es\<Colon>\<doteq>Ts\<rbrakk> \<Longrightarrow>
 					 E,dt\<Turnstile>e#es\<Colon>\<doteq>T#Ts"
@@ -472,8 +472,8 @@
 ML_setup {*
 simpset_ref() := simpset() delloop "split_all_tac"
 *}
-inductive_cases wt_stmt_cases: "E,dt\<Turnstile>c\<Colon>\<surd>"
-inductive_cases wt_elim_cases:
+
+inductive_cases wt_elim_cases [cases set]:
 	"E,dt\<Turnstile>In2  (LVar vn)               \<Colon>T"
 	"E,dt\<Turnstile>In2  ({accC,statDeclC,s}e..fn)\<Colon>T"
 	"E,dt\<Turnstile>In2  (e.[i])                 \<Colon>T"
@@ -499,7 +499,7 @@
         "E,dt\<Turnstile>In1r (l\<bullet> c)                  \<Colon>x" 
 	"E,dt\<Turnstile>In1r (If(e) c1 Else c2)      \<Colon>x"
 	"E,dt\<Turnstile>In1r (l\<bullet> While(e) c)         \<Colon>x"
-        "E,dt\<Turnstile>In1r (Do jump)               \<Colon>x"
+        "E,dt\<Turnstile>In1r (Jmp jump)              \<Colon>x"
 	"E,dt\<Turnstile>In1r (Throw e)               \<Colon>x"
 	"E,dt\<Turnstile>In1r (Try c1 Catch(tn vn) c2)\<Colon>x"
 	"E,dt\<Turnstile>In1r (c1 Finally c2)         \<Colon>x"
@@ -536,9 +536,10 @@
 apply (erule is_methdI, assumption)
 done
 
-(* Special versions of some typing rules, better suited to pattern match the
- * conclusion (no selectors in the conclusion\<dots>)
- *)
+
+text {* Special versions of some typing rules, better suited to pattern 
+        match the conclusion (no selectors in the conclusion)
+*}
 
 lemma wt_Call: 
 "\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT; E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;  
@@ -602,6 +603,14 @@
 apply auto
 done
 
+--{* In the special syntax to distinguish the typing judgements for expressions, 
+     statements, variables and expression lists the kind of term corresponds
+     to the kind of type in the end e.g. An statement (injection @{term In3} 
+    into terms, always has type void (injection @{term Inl} into the generalised
+    types. The following simplification procedures establish these kinds of
+    correlation. 
+ *}
+
 ML {*
 fun wt_fun name inj rhs =
 let
@@ -627,6 +636,19 @@
 Addsimprocs [wt_expr_proc,wt_var_proc,wt_exprs_proc,wt_stmt_proc]
 *}
 
+lemma wt_elim_BinOp:
+  "\<lbrakk>E,dt\<Turnstile>In1l (BinOp binop e1 e2)\<Colon>T;
+    \<And>T1 T2 T3.
+       \<lbrakk>E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2; wt_binop (prg E) binop T1 T2;
+          E,dt\<Turnstile>(if b then In1l e2 else In1r Skip)\<Colon>T3;
+          T = Inl (PrimT (binop_type binop))\<rbrakk>
+       \<Longrightarrow> P\<rbrakk>
+\<Longrightarrow> P"
+apply (erule wt_elim_cases)
+apply (cases b)
+apply auto
+done
+
 lemma Inj_eq_lemma [simp]: 
   "(\<forall>T. (\<exists>T'. T = Inj T' \<and> P T') \<longrightarrow> Q T) = (\<forall>T'. P T' \<longrightarrow> Q (Inj T'))"
 by auto
--- a/src/HOL/Bali/document/root.tex	Wed Oct 30 12:44:18 2002 +0100
+++ b/src/HOL/Bali/document/root.tex	Thu Oct 31 18:27:10 2002 +0100
@@ -4,6 +4,7 @@
 \usepackage{latexsym}
 \usepackage{graphicx}
 \usepackage{pdfsetup}
+\usepackage[english,french]{babel}
 
 \urlstyle{rm}
 \isabellestyle{it}
@@ -69,14 +70,16 @@
       \end{itemize}
 \item Packages
 \item Access Modifiers (\texttt{private}, \texttt{protected}, \texttt{public})
+\item A ``definite assignment'' check
 \end{itemize}
 
 The following features are missing in Bali wrt.{} JavaCard:
 \begin{itemize}
 \item Some primitive types (\texttt{byte, short})
-\item Most numeric operations, syntactic variants of statements
+\item Syntactic variants of statements
   (\texttt{do}-loop, \texttt{for}-loop)
-\item A ``definite assignment'' check
+\item Interface fields
+\item Inner Classes
 \end{itemize}
 
 In addition, features are missing that are not part of the JavaCard
@@ -123,6 +126,9 @@
 \item[WellType]
 Typesystem on the JavaCard term level.
 
+\item[DefiniteAssignment]
+The definite assignment analysis on the JavaCard term level.
+
 \item[WellForm]
 Typesystem on the JavaCard class, interface and program level.
 
@@ -142,6 +148,10 @@
 Conformance predicate for states. When does an execution state conform to the
 static types of the program given by the typesystem.
 
+\item[DefiniteAssignmentCorrect]
+Correctness of the definite assignment analysis. If the analysis regards a variable as definitely assigned at a
+certain program point, the variable will actually be assigned there during execution.
+
 \item[TypeSafe]
 Typesafety proof of the execution of JavaCard. ''Welltyped programs don't go
 wrong'' or more technical: The execution of a welltyped JavaCard program 
@@ -174,31 +184,43 @@
 
 \chapter{Basis}
 \input{Basis}
+
 \chapter{Table}
 \input{Table}    
 
 \chapter{Name}
 \input{Name}
+
 \chapter{Value}     
 \input{Value}
+
 \chapter{Type}
 \input{Type}      
+
 \chapter{Term}
 \input{Term}     
+
 \chapter{Decl}
 \input{Decl}          
+
 \chapter{TypeRel}
 \input{TypeRel}   
+
 \chapter{DeclConcepts}
 \input{DeclConcepts}  
 
 \chapter{WellType}
 \input{WellType}
+
+\chapter{DefiniteAssignment}
+\input{DefiniteAssignment}
+
 \chapter{WellForm}
 \input{WellForm}
 
 \chapter{State}
 \input{State}    
+
 \chapter{Eval}
 \input{Eval}          
 
@@ -207,19 +229,28 @@
 
 \chapter{Conform}
 \input{Conform}       
+
+\chapter{DefiniteAssignmentCorrect}
+\input{DefiniteAssignmentCorrect}
+
 \chapter{TypeSafe}
 \input{TypeSafe}
 
 \chapter{Evaln}
 \input{Evaln}         
+
 \chapter{Trans}
 \input{Trans}         
+
 \chapter{AxSem}
 \input{AxSem}      
+
 \chapter{AxSound}
 \input{AxSound}    
+
 \chapter{AxCompl}
 \input{AxCompl}    
+
 \chapter{AxExample}
 \input{AxExample}