Adapted to new inductive definition package.
authorberghofe
Mon, 11 Dec 2006 16:06:14 +0100
changeset 21765 89275a3ed7be
parent 21764 720b0add5206
child 21766 3eb48154388e
Adapted to new inductive definition package.
src/HOL/Bali/AxSem.thy
src/HOL/Bali/AxSound.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/Trans.thy
src/HOL/Bali/TypeRel.thy
src/HOL/Bali/TypeSafe.thy
src/HOL/Bali/WellForm.thy
src/HOL/Bali/WellType.thy
--- a/src/HOL/Bali/AxSem.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/AxSem.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -435,7 +435,6 @@
                                                 (   "_\<Turnstile>_:_" [61,0, 58] 57)
     ax_valids :: "prog \<Rightarrow> 'b triples \<Rightarrow> 'a triples \<Rightarrow> bool"
                                                 ("_,_|\<Turnstile>_"   [61,58,58] 57)
-    ax_derivs :: "prog \<Rightarrow> ('a triples \<times> 'a triples) set"
 
 syntax
 
@@ -443,10 +442,6 @@
                                                 (  "_||=_:_" [61,0, 58] 57)
      ax_valid :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
                                                 ( "_,_|=_"   [61,58,58] 57)
-     ax_Derivs:: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triples \<Rightarrow> bool"
-                                                ("_,_||-_"   [61,58,58] 57)
-     ax_Deriv :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
-                                                ( "_,_|-_"   [61,58,58] 57)
 
 syntax (xsymbols)
 
@@ -454,10 +449,6 @@
                                                 (  "_|\<Turnstile>_:_" [61,0, 58] 57)
      ax_valid :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
                                                 ( "_,_\<Turnstile>_"   [61,58,58] 57)
-     ax_Derivs:: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triples \<Rightarrow> bool"
-                                                ("_,_|\<turnstile>_"   [61,58,58] 57)
-     ax_Deriv :: "prog \<Rightarrow>  'b triples \<Rightarrow> 'a triple  \<Rightarrow> bool"
-                                                ( "_,_\<turnstile>_"   [61,58,58] 57)
 
 defs  triple_valid_def:  "G\<Turnstile>n:t  \<equiv> case t of {P} t\<succ> {Q} \<Rightarrow>
                           \<forall>Y s Z. P Y s Z \<longrightarrow> type_ok G t s \<longrightarrow>
@@ -465,8 +456,6 @@
 translations         "G|\<Turnstile>n:ts" == "Ball ts (triple_valid G n)"
 defs   ax_valids_def:"G,A|\<Turnstile>ts  \<equiv>  \<forall>n. G|\<Turnstile>n:A \<longrightarrow> G|\<Turnstile>n:ts"
 translations         "G,A \<Turnstile>t"  == "G,A|\<Turnstile>{t}"
-                     "G,A|\<turnstile>ts" == "(A,ts) \<in> ax_derivs G"
-                     "G,A \<turnstile>t"  == "G,A|\<turnstile>{t}"
 
 lemma triple_valid_def2: "G\<Turnstile>n:{P} t\<succ> {Q} =  
  (\<forall>Y s Z. P Y s Z 
@@ -487,63 +476,69 @@
 change_claset (fn cs => cs delSWrapper "split_all_tac");
 *}
 
-inductive "ax_derivs G" intros
+inductive2
+  ax_derivs :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool" ("_,_|\<turnstile>_" [61,58,58] 57)
+  and ax_deriv :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triple  \<Rightarrow> bool" ("_,_\<turnstile>_" [61,58,58] 57)
+  for G :: prog
+where
 
-  empty: " G,A|\<turnstile>{}"
-  insert:"\<lbrakk>G,A\<turnstile>t; G,A|\<turnstile>ts\<rbrakk> \<Longrightarrow>
+  "G,A \<turnstile>t \<equiv> G,A|\<turnstile>{t}"
+
+| empty: " G,A|\<turnstile>{}"
+| insert:"\<lbrakk>G,A\<turnstile>t; G,A|\<turnstile>ts\<rbrakk> \<Longrightarrow>
           G,A|\<turnstile>insert t ts"
 
-  asm:   "ts\<subseteq>A \<Longrightarrow> G,A|\<turnstile>ts"
+| asm:   "ts\<subseteq>A \<Longrightarrow> G,A|\<turnstile>ts"
 
 (* could be added for convenience and efficiency, but is not necessary
   cut:   "\<lbrakk>G,A'|\<turnstile>ts; G,A|\<turnstile>A'\<rbrakk> \<Longrightarrow>
            G,A |\<turnstile>ts"
 *)
-  weaken:"\<lbrakk>G,A|\<turnstile>ts'; ts \<subseteq> ts'\<rbrakk> \<Longrightarrow> G,A|\<turnstile>ts"
+| weaken:"\<lbrakk>G,A|\<turnstile>ts'; ts \<subseteq> ts'\<rbrakk> \<Longrightarrow> G,A|\<turnstile>ts"
 
-  conseq:"\<forall>Y s Z . P  Y s Z  \<longrightarrow> (\<exists>P' Q'. G,A\<turnstile>{P'} t\<succ> {Q'} \<and> (\<forall>Y' s'. 
+| conseq:"\<forall>Y s Z . P  Y s Z  \<longrightarrow> (\<exists>P' Q'. G,A\<turnstile>{P'} t\<succ> {Q'} \<and> (\<forall>Y' s'. 
          (\<forall>Y   Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>
                                  Q  Y' s' Z ))
                                          \<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q }"
 
-  hazard:"G,A\<turnstile>{P \<and>. Not \<circ> type_ok G t} t\<succ> {Q}"
+| hazard:"G,A\<turnstile>{P \<and>. Not \<circ> type_ok G t} t\<succ> {Q}"
 
-  Abrupt:  "G,A\<turnstile>{P\<leftarrow>(arbitrary3 t) \<and>. Not \<circ> normal} t\<succ> {P}"
+| Abrupt:  "G,A\<turnstile>{P\<leftarrow>(arbitrary3 t) \<and>. Not \<circ> normal} t\<succ> {P}"
 
   --{* variables *}
-  LVar:  " G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Var (lvar vn s))} LVar vn=\<succ> {P}"
+| 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};
+| FVar: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Q};
           G,A\<turnstile>{Q} e-\<succ> {\<lambda>Val:a:. fvar C stat fn a ..; R}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} {accC,C,stat}e..fn=\<succ> {R}"
 
-  AVar:  "\<lbrakk>G,A\<turnstile>{Normal P} e1-\<succ> {Q};
+| AVar:  "\<lbrakk>G,A\<turnstile>{Normal P} e1-\<succ> {Q};
           \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} e2-\<succ> {\<lambda>Val:i:. avar G i a ..; R}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} e1.[e2]=\<succ> {R}"
   --{* expressions *}
 
-  NewC: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Alloc G (CInst C) Q}\<rbrakk> \<Longrightarrow>
+| NewC: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Alloc G (CInst C) Q}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}"
 
-  NewA: "\<lbrakk>G,A\<turnstile>{Normal P} .init_comp_ty T. {Q};  G,A\<turnstile>{Q} e-\<succ>
+| NewA: "\<lbrakk>G,A\<turnstile>{Normal P} .init_comp_ty T. {Q};  G,A\<turnstile>{Q} e-\<succ>
 	  {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} New T[e]-\<succ> {R}"
 
-  Cast: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
+| Cast: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
           abupd (raise_if (\<not>G,s\<turnstile>v fits T) ClassCast) .; Q\<leftarrow>Val v}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} Cast T e-\<succ> {Q}"
 
-  Inst: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
+| Inst: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
                   Q\<leftarrow>Val (Bool (v\<noteq>Null \<and> G,s\<turnstile>v fits RefT T))}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} e InstOf T-\<succ> {Q}"
 
-  Lit:                          "G,A\<turnstile>{Normal (P\<leftarrow>Val v)} Lit v-\<succ> {P}"
+| Lit:                          "G,A\<turnstile>{Normal (P\<leftarrow>Val v)} Lit v-\<succ> {P}"
 
-  UnOp: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. Q\<leftarrow>Val (eval_unop unop v)}\<rbrakk>
+| UnOp: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. Q\<leftarrow>Val (eval_unop unop v)}\<rbrakk>
           \<Longrightarrow>
           G,A\<turnstile>{Normal P} UnOp unop e-\<succ> {Q}"
 
-  BinOp:
+| BinOp:
    "\<lbrakk>G,A\<turnstile>{Normal P} e1-\<succ> {Q};
      \<forall>v1. G,A\<turnstile>{Q\<leftarrow>Val v1} 
                (if need_second_arg binop v1 then (In1l e2) else (In1r Skip))\<succ>
@@ -551,20 +546,20 @@
     \<Longrightarrow>
     G,A\<turnstile>{Normal P} BinOp binop e1 e2-\<succ> {R}" 
 
-  Super:" G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Val (val_this s))} Super-\<succ> {P}"
+| Super:" G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Val (val_this s))} Super-\<succ> {P}"
 
-  Acc:  "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {\<lambda>Var:(v,f):. Q\<leftarrow>Val v}\<rbrakk> \<Longrightarrow>
+| Acc:  "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {\<lambda>Var:(v,f):. Q\<leftarrow>Val v}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} Acc va-\<succ> {Q}"
 
-  Ass:  "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {Q};
+| Ass:  "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {Q};
      \<forall>vf. G,A\<turnstile>{Q\<leftarrow>Var vf} e-\<succ> {\<lambda>Val:v:. assign (snd vf) v .; R}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} va:=e-\<succ> {R}"
 
-  Cond: "\<lbrakk>G,A \<turnstile>{Normal P} e0-\<succ> {P'};
+| Cond: "\<lbrakk>G,A \<turnstile>{Normal P} e0-\<succ> {P'};
           \<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} (if b then e1 else e2)-\<succ> {Q}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} e0 ? e1 : e2-\<succ> {Q}"
 
-  Call: 
+| Call: 
 "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q}; \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {R a};
   \<forall>a vs invC declC l. G,A\<turnstile>{(R a\<leftarrow>Vals vs \<and>.
  (\<lambda>s. declC=invocation_declclass G mode (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> \<and>
@@ -575,37 +570,37 @@
  Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}\<rbrakk> \<Longrightarrow>
          G,A\<turnstile>{Normal P} {accC,statT,mode}e\<cdot>mn({pTs}args)-\<succ> {S}"
 
-  Methd:"\<lbrakk>G,A\<union> {{P} Methd-\<succ> {Q} | ms} |\<turnstile> {{P} body G-\<succ> {Q} | ms}\<rbrakk> \<Longrightarrow>
+| Methd:"\<lbrakk>G,A\<union> {{P} Methd-\<succ> {Q} | ms} |\<turnstile> {{P} body G-\<succ> {Q} | ms}\<rbrakk> \<Longrightarrow>
                                  G,A|\<turnstile>{{P} Methd-\<succ>  {Q} | ms}"
 
-  Body: "\<lbrakk>G,A\<turnstile>{Normal P} .Init D. {Q}; 
+| Body: "\<lbrakk>G,A\<turnstile>{Normal P} .Init D. {Q}; 
   G,A\<turnstile>{Q} .c. {\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>(In1 (the (locals s Result)))}\<rbrakk> 
     \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} Body D c-\<succ> {R}"
   
   --{* expression lists *}
 
-  Nil:                          "G,A\<turnstile>{Normal (P\<leftarrow>Vals [])} []\<doteq>\<succ> {P}"
+| Nil:                          "G,A\<turnstile>{Normal (P\<leftarrow>Vals [])} []\<doteq>\<succ> {P}"
 
-  Cons: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q};
+| Cons: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q};
           \<forall>v. G,A\<turnstile>{Q\<leftarrow>Val v} es\<doteq>\<succ> {\<lambda>Vals:vs:. R\<leftarrow>Vals (v#vs)}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} e#es\<doteq>\<succ> {R}"
 
   --{* statements *}
 
-  Skip:                         "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit>)} .Skip. {P}"
+| Skip:                         "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit>)} .Skip. {P}"
 
-  Expr: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>
+| Expr: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} .Expr e. {Q}"
 
-  Lab: "\<lbrakk>G,A\<turnstile>{Normal P} .c. {abupd (absorb l) .; Q}\<rbrakk> \<Longrightarrow>
+| Lab: "\<lbrakk>G,A\<turnstile>{Normal P} .c. {abupd (absorb l) .; Q}\<rbrakk> \<Longrightarrow>
                            G,A\<turnstile>{Normal P} .l\<bullet> c. {Q}"
 
-  Comp: "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q};
+| Comp: "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q};
           G,A\<turnstile>{Q} .c2. {R}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} .c1;;c2. {R}"
 
-  If:   "\<lbrakk>G,A \<turnstile>{Normal P} e-\<succ> {P'};
+| If:   "\<lbrakk>G,A \<turnstile>{Normal P} e-\<succ> {P'};
           \<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c1 else c2). {Q}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} .If(e) c1 Else c2. {Q}"
 (* unfolding variant of Loop, not needed here
@@ -613,28 +608,28 @@
           \<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c;;While(e) c else Skip).{Q}\<rbrakk>
          \<Longrightarrow>              G,A\<turnstile>{Normal P} .While(e) c. {Q}"
 *)
-  Loop: "\<lbrakk>G,A\<turnstile>{P} e-\<succ> {P'}; 
+| 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>}"
   
-  Jmp: "G,A\<turnstile>{Normal (abupd (\<lambda>a. (Some (Jump j))) .; P\<leftarrow>\<diamondsuit>)} .Jmp 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>
+| Throw:"\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} .Throw e. {Q}"
 
-  Try:  "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {SXAlloc G Q};
+| Try:  "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {SXAlloc G Q};
           G,A\<turnstile>{Q \<and>. (\<lambda>s.  G,s\<turnstile>catch C) ;. new_xcpt_var vn} .c2. {R};
               (Q \<and>. (\<lambda>s. \<not>G,s\<turnstile>catch C)) \<Rightarrow> R\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} .Try c1 Catch(C vn) c2. {R}"
 
-  Fin:  "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q};
+| Fin:  "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q};
       \<forall>x. G,A\<turnstile>{Q \<and>. (\<lambda>s. x = fst s) ;. abupd (\<lambda>x. None)}
               .c2. {abupd (abrupt_if (x\<noteq>None) x) .; R}\<rbrakk> \<Longrightarrow>
                                  G,A\<turnstile>{Normal P} .c1 Finally c2. {R}"
 
-  Done:                       "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit> \<and>. initd C)} .Init C. {P}"
+| Done:                       "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit> \<and>. initd C)} .Init C. {P}"
 
-  Init: "\<lbrakk>the (class G C) = c;
+| Init: "\<lbrakk>the (class G C) = c;
           G,A\<turnstile>{Normal ((P \<and>. Not \<circ> initd C) ;. supd (init_class_obj G C))}
               .(if C = Object then Skip else Init (super c)). {Q};
       \<forall>l. G,A\<turnstile>{Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars empty}
@@ -645,10 +640,10 @@
 @{text InsInitE}, @{text InsInitV}, @{text FinA} only used by the smallstep 
 semantics.
 *}
-  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}"
+| 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 
 *)
@@ -1032,7 +1027,7 @@
 *}
 declare ax_Abrupts [intro!]
 
-lemmas ax_Normal_cases = ax_cases [of _ _ normal]
+lemmas ax_Normal_cases = ax_cases [of _ _ _ normal]
 
 lemma ax_Skip [intro!]: "G,(A::'a triple set)\<turnstile>{P\<leftarrow>\<diamondsuit>} .Skip. {P::'a assn}"
 apply (rule ax_Normal_cases)
--- a/src/HOL/Bali/AxSound.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/AxSound.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -375,18 +375,18 @@
   thus ?case
     by (unfold ax_valids2_def) blast
 next
-  case (asm A ts)
+  case (asm ts A)
   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')
+  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)
+  case (conseq P A t Q)
   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>
@@ -431,7 +431,7 @@
     qed
   qed
 next
-  case (hazard A P Q t)
+  case (hazard A P t Q)
   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
@@ -474,7 +474,7 @@
     qed
   qed
 next
-  case (FVar A statDeclC P Q R accC e fn stat)
+  case (FVar A P statDeclC Q e stat fn R accC)
   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} }"
@@ -569,7 +569,7 @@
     qed
   qed
 next
-  case (AVar A P Q R e1 e2)
+  case (AVar A P e1 Q e2 R)
   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
@@ -633,7 +633,7 @@
     qed
   qed
 next
-  case (NewC A C P Q)
+  case (NewC A P C 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)
@@ -673,7 +673,7 @@
     qed
   qed
 next
-  case (NewA A P Q R T e)
+  case (NewA A P T Q e R)
   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}}" .
@@ -746,7 +746,7 @@
     qed
   qed
 next
-  case (Cast A P Q T e)
+  case (Cast A P e T Q)
   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} }" .
@@ -786,7 +786,7 @@
     qed
   qed
 next
-  case (Inst A P Q T e)
+  case (Inst A P e Q T)
   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} }"
@@ -841,7 +841,7 @@
     qed
   qed
 next
-  case (UnOp A P Q e unop)
+  case (UnOp A P e Q 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)
@@ -878,7 +878,7 @@
     qed
   qed
 next
-  case (BinOp A P Q R binop e1 e2)
+  case (BinOp A P e1 Q binop e2 R)
   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>
@@ -977,7 +977,7 @@
     qed
   qed
 next
-  case (Acc A P Q var)
+  case (Acc A P var Q)
   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)
@@ -1013,7 +1013,7 @@
     qed
   qed
 next
-  case (Ass A P Q R e var)
+  case (Ass A P var Q e R)
   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} }"
@@ -1125,7 +1125,7 @@
     qed
   qed
 next
-  case (Cond A P P' Q e0 e1 e2)
+  case (Cond A P e0 P' e1 e2 Q)
   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
@@ -1215,7 +1215,7 @@
     qed
   qed
 next
-  case (Call A P Q R S accC' args e mn mode pTs' statT)
+  case (Call A P e Q args R mode statT mn pTs' S accC')
   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
@@ -1604,7 +1604,7 @@
   show "G,A|\<Turnstile>\<Colon>{{P} Methd-\<succ> {Q} | ms}"
     by (rule Methd_sound)
 next
-  case (Body A D P Q R c)
+  case (Body A P D Q c R)
   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))} }".
@@ -1697,7 +1697,7 @@
     qed
   qed
 next
-  case (Cons A P Q R e es)
+  case (Cons A P e Q es R)
   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
@@ -1779,7 +1779,7 @@
     qed
   qed
 next
-  case (Expr A P Q e)
+  case (Expr A P e Q)
   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)
@@ -1809,7 +1809,7 @@
     qed
   qed
 next
-  case (Lab A P Q c l)
+  case (Lab A P c l Q)
   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)
@@ -1846,7 +1846,7 @@
     qed
   qed
 next
-  case (Comp A P Q R c1 c2)
+  case (Comp A P c1 Q c2 R)
   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} }"
@@ -1905,7 +1905,7 @@
     qed
   qed
 next
-  case (If A P P' Q c1 c2 e)
+  case (If A P e P' c1 c2 Q)
   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
@@ -1982,7 +1982,7 @@
     qed
   qed
 next
-  case (Loop A P P' c e l)
+  case (Loop A P e P' c l)
   have valid_e: "G,A|\<Turnstile>\<Colon>{ {P} e-\<succ> {P'} }".
   have valid_c: "G,A|\<Turnstile>\<Colon>{ {Normal (P'\<leftarrow>=True)} 
                          .c. 
@@ -2020,7 +2020,7 @@
                 \<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)
+	  case (Loop s0' e' n' b s1' c' s2' l' 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". 
@@ -2173,7 +2173,7 @@
 	    qed
 	  qed
 	next
-	  case (Abrupt n' s t' abr Y' T E)
+	  case (Abrupt abr s t' n' 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".
@@ -2212,7 +2212,7 @@
     qed
   qed
 next
-  case (Jmp A P j)
+  case (Jmp A j P)
   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
@@ -2239,7 +2239,7 @@
     qed
   qed
 next
-  case (Throw A P Q e)
+  case (Throw A P e Q)
   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)
@@ -2277,7 +2277,7 @@
     qed
   qed
 next
-  case (Try A C P Q R c1 c2 vn)
+  case (Try A P c1 Q C vn c2 R)
   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. 
@@ -2409,7 +2409,7 @@
     qed
   qed
 next
-  case (Fin A P Q R c1 c2)
+  case (Fin A P c1 Q  c2 R)
   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.
@@ -2483,7 +2483,7 @@
     qed
   qed
 next
-  case (Done A C P)
+  case (Done A P C)
   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
@@ -2506,7 +2506,7 @@
     qed
   qed
 next
-  case (Init A C P Q R c)
+  case (Init C c A P Q R)
   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))}
@@ -2618,7 +2618,7 @@
     qed
   qed
 next
-  case (InsInitV A P Q c v)
+  case (InsInitV A P c v Q)
   show "G,A|\<Turnstile>\<Colon>{ {Normal P} InsInitV c v=\<succ> {Q} }"
   proof (rule valid_var_NormalI)
     fix s0 vf n s1 L Z
@@ -2630,7 +2630,7 @@
     thus "Q \<lfloor>vf\<rfloor>\<^sub>v s1 Z \<and> s1\<Colon>\<preceq>(G, L)"..
   qed
 next
-  case (InsInitE A P Q c e)
+  case (InsInitE A P c e Q)
   show "G,A|\<Turnstile>\<Colon>{ {Normal P} InsInitE c e-\<succ> {Q} }"
   proof (rule valid_expr_NormalI)
     fix s0 v n s1 L Z
@@ -2642,7 +2642,7 @@
     thus "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"..
   qed
 next
-  case (Callee A P Q e l)
+  case (Callee A P l e Q)
   show "G,A|\<Turnstile>\<Colon>{ {Normal P} Callee l e-\<succ> {Q} }"
   proof (rule valid_expr_NormalI)
     fix s0 v n s1 L Z
@@ -2654,7 +2654,7 @@
     thus "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"..
   qed
 next
-  case (FinA A P Q a c)
+  case (FinA A P a c Q)
   show "G,A|\<Turnstile>\<Colon>{ {Normal P} .FinA a c. {Q} }"
   proof (rule valid_stmt_NormalI)
     fix s0 v n s1 L Z
--- a/src/HOL/Bali/DeclConcepts.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/DeclConcepts.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -536,26 +536,21 @@
 
 subsubsection "members"
 
-consts
-members:: "prog \<Rightarrow> (qtname \<times> (qtname \<times> memberdecl)) set"
 (* Can't just take a function: prog \<Rightarrow> qtname \<Rightarrow> memberdecl set because
    the class qtname changes to the superclass in the inductive definition
    below
 *)
 
-syntax
-member_of:: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
-                           ("_ \<turnstile> _ member'_of _" [61,61,61] 60)
-
-translations
- "G\<turnstile>m member_of C" \<rightleftharpoons> "(C,m) \<in> members G"
+inductive2
+  members :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
+    ("_ \<turnstile> _ member'_of _" [61,61,61] 60)
+  for G :: prog
+where
 
-inductive "members G"  intros
-
-Immediate: "\<lbrakk>G\<turnstile>mbr m declared_in C;declclass m = C\<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
-Inherited: "\<lbrakk>G\<turnstile>m inheritable_in (pid C); G\<turnstile>memberid m undeclared_in C; 
-             G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S 
-            \<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
+  Immediate: "\<lbrakk>G\<turnstile>mbr m declared_in C;declclass m = C\<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
+| Inherited: "\<lbrakk>G\<turnstile>m inheritable_in (pid C); G\<turnstile>memberid m undeclared_in C; 
+               G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S 
+              \<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
 text {* Note that in the case of an inherited member only the members of the
 direct superclass are concerned. If a member of a superclass of the direct
 superclass isn't inherited in the direct superclass (not member of the
@@ -619,9 +614,6 @@
 translations
  "G\<turnstile>Methd s m member_in C" \<rightleftharpoons> "G\<turnstile>(method s m) member_in C" 
 
-consts stat_overridesR::
-  "prog  \<Rightarrow> ((qtname \<times> mdecl) \<times> (qtname \<times> mdecl)) set"
-
 lemma member_inD: "G\<turnstile>m member_in C 
  \<Longrightarrow> \<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC"
 by (auto simp add: member_in_def)
@@ -641,48 +633,43 @@
 *}
 
 text {* Static overriding (used during the typecheck of the compiler) *}
-syntax
-stat_overrides:: "prog  \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
-                                  ("_ \<turnstile> _ overrides\<^sub>S _" [61,61,61] 60)
-translations
- "G\<turnstile>new overrides\<^sub>S  old" == "(new,old) \<in> stat_overridesR G "
 
-inductive "stat_overridesR G" intros
+inductive2
+  stat_overridesR :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
+    ("_ \<turnstile> _ overrides\<^sub>S _" [61,61,61] 60)
+  for G :: prog
+where
 
-Direct: "\<lbrakk>\<not> is_static new; msig new = msig old; 
-         G\<turnstile>Method new declared_in (declclass new);  
-         G\<turnstile>Method old declared_in (declclass old); 
-         G\<turnstile>Method old inheritable_in pid (declclass new);
-         G\<turnstile>(declclass new) \<prec>\<^sub>C\<^sub>1 superNew;
-         G \<turnstile>Method old member_of superNew
-         \<rbrakk> \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
+  Direct: "\<lbrakk>\<not> is_static new; msig new = msig old; 
+           G\<turnstile>Method new declared_in (declclass new);  
+           G\<turnstile>Method old declared_in (declclass old); 
+           G\<turnstile>Method old inheritable_in pid (declclass new);
+           G\<turnstile>(declclass new) \<prec>\<^sub>C\<^sub>1 superNew;
+           G \<turnstile>Method old member_of superNew
+           \<rbrakk> \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
 
-Indirect: "\<lbrakk>G\<turnstile>new overrides\<^sub>S inter; G\<turnstile>inter overrides\<^sub>S old\<rbrakk>
-           \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
+| Indirect: "\<lbrakk>G\<turnstile>new overrides\<^sub>S inter; G\<turnstile>inter overrides\<^sub>S old\<rbrakk>
+             \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
 
 text {* Dynamic overriding (used during the typecheck of the compiler) *}
-consts overridesR::
-  "prog  \<Rightarrow> ((qtname \<times> mdecl) \<times> (qtname \<times> mdecl)) set"
 
-
-overrides:: "prog  \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
-                                  ("_ \<turnstile> _ overrides _" [61,61,61] 60)
-translations
- "G\<turnstile>new overrides old" == "(new,old) \<in> overridesR G "
-
-inductive "overridesR G" intros
+inductive2
+  overridesR :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
+    ("_ \<turnstile> _ overrides _" [61,61,61] 60)
+  for G :: prog
+where
 
-Direct: "\<lbrakk>\<not> is_static new; \<not> is_static old; accmodi new \<noteq> Private;
-         msig new = msig old; 
-         G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
-         G\<turnstile>Method new declared_in (declclass new);  
-         G\<turnstile>Method old declared_in (declclass old);    
-         G\<turnstile>Method old inheritable_in pid (declclass new);
-         G\<turnstile>resTy new \<preceq> resTy old
-         \<rbrakk> \<Longrightarrow> G\<turnstile>new overrides old"
+  Direct: "\<lbrakk>\<not> is_static new; \<not> is_static old; accmodi new \<noteq> Private;
+           msig new = msig old; 
+           G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
+           G\<turnstile>Method new declared_in (declclass new);  
+           G\<turnstile>Method old declared_in (declclass old);    
+           G\<turnstile>Method old inheritable_in pid (declclass new);
+           G\<turnstile>resTy new \<preceq> resTy old
+           \<rbrakk> \<Longrightarrow> G\<turnstile>new overrides old"
 
-Indirect: "\<lbrakk>G\<turnstile>new overrides inter; G\<turnstile>inter overrides old\<rbrakk>
-           \<Longrightarrow> G\<turnstile>new overrides old"
+| Indirect: "\<lbrakk>G\<turnstile>new overrides inter; G\<turnstile>inter overrides old\<rbrakk>
+            \<Longrightarrow> G\<turnstile>new overrides old"
 
 syntax
 sig_stat_overrides:: 
@@ -797,28 +784,31 @@
 \end{itemize} 
 *} 
 
-
-consts
-accessible_fromR:: 
- "prog \<Rightarrow> qtname \<Rightarrow> ((qtname \<times> memberdecl) \<times>  qtname) set"
+inductive2
+  accessible_fromR :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
+  and accessible_from :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+    ("_ \<turnstile> _ of _ accessible'_from _" [61,61,61,61] 60)
+  and method_accessible_from :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+    ("_ \<turnstile>Method _ of _ accessible'_from _" [61,61,61,61] 60)
+  for G :: prog and accclass :: qtname
+where
+  "G\<turnstile>membr of cls accessible_from accclass \<equiv> accessible_fromR G accclass membr cls"
 
-syntax 
-accessible_from:: 
- "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
-                   ("_ \<turnstile> _ of _ accessible'_from _" [61,61,61,61] 60)
+| "G\<turnstile>Method m of cls accessible_from accclass \<equiv>
+   G\<turnstile>methdMembr m of cls accessible_from accclass"  
 
-translations
-"G\<turnstile>membr of cls accessible_from accclass"  
- \<rightleftharpoons> "(membr,cls) \<in> accessible_fromR G accclass"
+| Immediate:  "\<lbrakk>G\<turnstile>membr member_of class;
+                G\<turnstile>(Class class) accessible_in (pid accclass);
+                G\<turnstile>membr in class permits_acc_from accclass 
+               \<rbrakk> \<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
 
-syntax 
-method_accessible_from:: 
- "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
-                   ("_ \<turnstile>Method _ of _ accessible'_from _" [61,61,61,61] 60)
-
-translations
-"G\<turnstile>Method m of cls accessible_from accclass"  
- \<rightleftharpoons> "G\<turnstile>methdMembr m of cls accessible_from accclass"  
+| Overriding: "\<lbrakk>G\<turnstile>membr member_of class;
+                G\<turnstile>(Class class) accessible_in (pid accclass);
+                membr=(C,mdecl new);
+                G\<turnstile>(C,new) overrides\<^sub>S old; 
+                G\<turnstile>class \<prec>\<^sub>C sup;
+                G\<turnstile>Method old of sup accessible_from accclass
+               \<rbrakk>\<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
 
 syntax 
 methd_accessible_from:: 
@@ -838,41 +828,29 @@
 "G\<turnstile>Field fn f of C accessible_from accclass"  
  \<rightleftharpoons> "G\<turnstile>(fieldm fn f) of C accessible_from accclass" 
 
-inductive "accessible_fromR G accclass" intros
-Immediate:  "\<lbrakk>G\<turnstile>membr member_of class;
-              G\<turnstile>(Class class) accessible_in (pid accclass);
-              G\<turnstile>membr in class permits_acc_from accclass 
-             \<rbrakk> \<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
-
-Overriding: "\<lbrakk>G\<turnstile>membr member_of class;
-              G\<turnstile>(Class class) accessible_in (pid accclass);
-              membr=(C,mdecl new);
-              G\<turnstile>(C,new) overrides\<^sub>S old; 
-              G\<turnstile>class \<prec>\<^sub>C sup;
-              G\<turnstile>Method old of sup accessible_from accclass
-             \<rbrakk>\<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
-
-consts
-dyn_accessible_fromR:: 
- "prog \<Rightarrow> qtname \<Rightarrow> ((qtname \<times> memberdecl) \<times>  qtname) set"
+inductive2
+  dyn_accessible_fromR :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
+  and dyn_accessible_from' ::  "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+    ("_ \<turnstile> _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
+  and method_dyn_accessible_from :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+    ("_ \<turnstile>Method _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
+  for G :: prog and accclass :: qtname
+where
+  "G\<turnstile>membr in C dyn_accessible_from accC \<equiv> dyn_accessible_fromR G accC membr C"
 
-syntax 
-dyn_accessible_from:: 
- "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
-                   ("_ \<turnstile> _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
+| "G\<turnstile>Method m in C dyn_accessible_from accC \<equiv>
+   G\<turnstile>methdMembr m in C dyn_accessible_from accC"
 
-translations
-"G\<turnstile>membr in C dyn_accessible_from accC"  
- \<rightleftharpoons> "(membr,C) \<in> dyn_accessible_fromR G accC"
+| Immediate:  "\<lbrakk>G\<turnstile>membr member_in class;
+                G\<turnstile>membr in class permits_acc_from accclass 
+               \<rbrakk> \<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
 
-syntax 
-method_dyn_accessible_from:: 
- "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
-                 ("_ \<turnstile>Method _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
-
-translations
-"G\<turnstile>Method m in C dyn_accessible_from accC"  
- \<rightleftharpoons> "G\<turnstile>methdMembr m in C dyn_accessible_from accC"  
+| Overriding: "\<lbrakk>G\<turnstile>membr member_in class;
+                membr=(C,mdecl new);
+                G\<turnstile>(C,new) overrides old; 
+                G\<turnstile>class \<prec>\<^sub>C sup;
+                G\<turnstile>Method old in sup dyn_accessible_from accclass
+               \<rbrakk>\<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
 
 syntax 
 methd_dyn_accessible_from:: 
@@ -891,18 +869,6 @@
 translations
 "G\<turnstile>Field fn f in dynC dyn_accessible_from accC"  
  \<rightleftharpoons> "G\<turnstile>(fieldm fn f) in dynC dyn_accessible_from accC"
-  
-inductive "dyn_accessible_fromR G accclass" intros
-Immediate:  "\<lbrakk>G\<turnstile>membr member_in class;
-              G\<turnstile>membr in class permits_acc_from accclass 
-             \<rbrakk> \<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
-
-Overriding: "\<lbrakk>G\<turnstile>membr member_in class;
-              membr=(C,mdecl new);
-              G\<turnstile>(C,new) overrides old; 
-              G\<turnstile>class \<prec>\<^sub>C sup;
-              G\<turnstile>Method old in sup dyn_accessible_from accclass
-             \<rbrakk>\<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
 
 
 lemma accessible_from_commonD: "G\<turnstile>m of C accessible_from S
@@ -967,13 +933,13 @@
   from n m eqid  
   show "n=m"
   proof (induct)
-    case (Immediate C n)
+    case (Immediate n C)
     assume member_n: "G\<turnstile> mbr n declared_in C" "declclass n = C"
     assume eqid: "memberid n = memberid m"
     assume "G \<turnstile> m member_of C"
     then show "n=m"
     proof (cases)
-      case (Immediate _ m')
+      case (Immediate m' _)
       with eqid 
       have "m=m'"
            "memberid n = memberid m" 
@@ -987,7 +953,7 @@
                            cdeclaredmethd_def cdeclaredfield_def
                     split: memberdecl.splits)
     next
-      case (Inherited _ _ m')
+      case (Inherited m' _ _)
       then have "G\<turnstile> memberid m undeclared_in C"
 	by simp
       with eqid member_n
@@ -995,7 +961,7 @@
 	by (cases n) (auto dest: declared_not_undeclared)
     qed
   next
-    case (Inherited C S n)
+    case (Inherited n C S)
     assume undecl: "G\<turnstile> memberid n undeclared_in C"
     assume  super: "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S"
     assume    hyp: "\<lbrakk>G \<turnstile> m member_of S; memberid n = memberid m\<rbrakk> \<Longrightarrow> n = m"
@@ -1021,13 +987,13 @@
 
 lemma member_of_is_classD: "G\<turnstile>m member_of C \<Longrightarrow> is_class G C"
 proof (induct set: members)
-  case (Immediate C m)
+  case (Immediate m C)
   assume "G\<turnstile> mbr m declared_in C"
   then show "is_class G C"
     by (cases "mbr m")
        (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)
 next
-  case (Inherited C S m)  
+  case (Inherited m C S)  
   assume "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S" and "is_class G S"
   then show "is_class G C"
     by - (rule subcls_is_class2,auto)
@@ -1046,10 +1012,10 @@
 lemma member_of_class_relation:
   "G\<turnstile>m member_of C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
 proof (induct set: members)
-  case (Immediate C m)
+  case (Immediate m C)
   then show "G\<turnstile>C \<preceq>\<^sub>C declclass m" by simp
 next
-  case (Inherited C S m)
+  case (Inherited m C S)
   then show "G\<turnstile>C \<preceq>\<^sub>C declclass m" 
     by (auto dest: r_into_rtrancl intro: rtrancl_trans)
 qed
@@ -1152,12 +1118,12 @@
   from acc_C static
   show "G\<turnstile>m in (declclass m) dyn_accessible_from accC"
   proof (induct)
-    case (Immediate C m)
+    case (Immediate m C)
     then show ?case 
       by (auto intro!: dyn_accessible_fromR.Immediate
                  dest: member_in_declC permits_acc_static_declC) 
   next 
-    case (Overriding declCNew C m new old sup)
+    case (Overriding m C declCNew new old sup)
     then have "\<not> is_static m"
       by (auto dest: overrides_commonD)
     moreover
@@ -1259,7 +1225,7 @@
   from m subclseq_D_C subclseq_C_m
   show ?thesis
   proof (induct)
-    case (Immediate D m)
+    case (Immediate m D)
     assume "declclass m = D" and
            "G\<turnstile>D\<preceq>\<^sub>C C" and "G\<turnstile>C\<preceq>\<^sub>C declclass m"
     with ws have "D=C" 
@@ -1268,7 +1234,7 @@
     show "G\<turnstile>m member_of C"
       by (auto intro: members.Immediate)
   next
-    case (Inherited D S m)
+    case (Inherited m D S)
     assume member_of_D_props: 
             "G \<turnstile> m inheritable_in pid D" 
             "G\<turnstile> memberid m undeclared_in D"  
@@ -1714,7 +1680,7 @@
     from member_of
     show "?Methd C"
     proof (cases)
-      case (Immediate Ca membr)
+      case (Immediate membr Ca)
       then have "Ca=C" "membr = method sig m" and 
                 "G\<turnstile>Methd sig m declared_in C" "declclass m = C"
 	by (cases m,auto)
@@ -1727,7 +1693,7 @@
       show ?thesis
 	by (simp add: methd_rec)
     next
-      case (Inherited Ca S membr)
+      case (Inherited membr Ca S)
       with clsC
       have eq_Ca_C: "Ca=C" and
             undecl: "G\<turnstile>mid sig undeclared_in C" and
--- a/src/HOL/Bali/DefiniteAssignment.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/DefiniteAssignment.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -513,25 +513,6 @@
          brk :: "breakass" --{* Definetly assigned variables for 
                                 abrupt 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"
  
@@ -543,35 +524,45 @@
 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>"
+text {*
+In @{text "E\<turnstile> B \<guillemotright>t\<guillemotright> A"},
+@{text B} denotes the ''assigned'' variables before evaluating term @{text t},
+whereas @{text A} denotes the ''assigned'' variables after evaluating term @{text t}.
+The environment @{term E} is only needed for the conditional @{text "_ ? _ : _"}.
+The definite assignment rules refer to the typing rules here to
+distinguish boolean and other expressions.
+*}
 
- 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" 
+inductive2
+  da :: "env \<Rightarrow> lname set \<Rightarrow> term \<Rightarrow> assigned \<Rightarrow> bool" ("_\<turnstile> _ \<guillemotright>_\<guillemotright> _" [65,65,65,65] 71)
+where
+  Skip: "Env\<turnstile> B \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>"
 
- 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"
+| 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" 
 
- 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"
+| 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
+     @{term "Env\<turnstile>B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E"} is just used to ensure 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 
@@ -585,12 +576,12 @@
      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"
+| 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
@@ -602,14 +593,14 @@
      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"
+| 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
@@ -618,21 +609,21 @@
      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"
+| 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"
+| 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" 
+| 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}.
@@ -671,7 +662,7 @@
      and @{text NewA}
 *}
 
- Init: "Env\<turnstile> B \<guillemotright>\<langle>Init C\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>"
+| 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 
@@ -679,91 +670,91 @@
      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>" 
+| 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"
 
- 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"
 
- 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"
 
- 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>"
 
- 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"
 
- 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"
 
- 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));
+| 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 CondAnd e1 e2\<rangle>\<guillemotright> A"
+            \<Longrightarrow>
+            Env\<turnstile> B \<guillemotright>\<langle>BinOp CondOr 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"
 
- 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>"
 
- 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"
+| 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"
+| 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"
+| 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"
+| 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" 
+| 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" 
+| 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"
+| 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
@@ -786,16 +777,16 @@
       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>
+| 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>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"
+         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
@@ -809,21 +800,21 @@
       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>" 
+| 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>
+| FVar: "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> A 
          \<Longrightarrow> 
-         Env\<turnstile> B \<guillemotright>\<langle>e1.[e2]\<rangle>\<guillemotright> A" 
+         Env\<turnstile> B \<guillemotright>\<langle>{accC,statDeclC,stat}e..fn\<rangle>\<guillemotright> A" 
 
- Nil: "Env\<turnstile> B \<guillemotright>\<langle>[]::expr list\<rangle>\<guillemotright> \<lparr>nrm=B, brk=\<lambda> l. UNIV\<rparr>" 
+| 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" 
 
- 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" 
+| 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]
@@ -832,7 +823,7 @@
 ML_setup {*
 change_simpset (fn ss => ss delloop "split_all_tac");
 *}
-inductive_cases da_elim_cases [cases set]:
+inductive_cases2 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"
@@ -1076,7 +1067,7 @@
     from Expr.prems Expr.hyps 
     show ?case by cases simp
   next
-    case (Lab A B C Env c l B' A')
+    case (Lab Env B c C A 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
@@ -1102,7 +1093,7 @@
     ultimately show ?case
       by simp
   next
-    case (Comp A B C1 C2 Env c1 c2 B' A')
+    case (Comp Env B c1 C1 c2 C2 A 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'
@@ -1123,7 +1114,7 @@
     show ?case
       by auto
   next
-    case (If A B C1 C2 E Env c1 c2 e B' A')
+    case (If Env B e E c1 C1 c2 C2 A 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'
@@ -1144,7 +1135,7 @@
     show ?case
       by auto
   next
-    case (Loop A B C E Env c e l B' A')
+    case (Loop Env B e E c C A 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'" .
@@ -1180,12 +1171,12 @@
     ultimately show ?case
       by auto
   next
-    case (Jmp A B Env jump B' A')
+    case (Jmp jump B A Env 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')
+    case (Try Env B c1 C1 vn C c2 C2 A 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'" .
@@ -1210,7 +1201,7 @@
     show ?case
       by auto
   next
-    case (Fin A B C1 C2 Env c1 c2 B' A')
+    case (Fin Env B c1 C1 c2 C2 A 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'" .
@@ -1247,7 +1238,7 @@
    next
      case UnOp thus ?case by -  (erule da_elim_cases, auto)
    next
-     case (CondAnd A B E1 E2 Env e1 e2 B' A')
+     case (CondAnd Env B e1 E1 e2 E2 A B' A')
      have A: "nrm A = B \<union>
                        assigns_if True (BinOp CondAnd e1 e2) \<inter>
                        assigns_if False (BinOp CondAnd e1 e2)"
@@ -1276,7 +1267,7 @@
    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')
+     case (CondBool Env c e1 e2 B C E1 E2 A B' A')
      have A: "nrm A = B \<union> 
                         assigns_if True (c ? e1 : e2) \<inter> 
                         assigns_if False (c ? e1 : e2)"
@@ -1295,7 +1286,7 @@
      with A A' show ?case 
        by auto 
    next
-     case (Cond A B C E1 E2 Env c e1 e2 B' A')  
+     case (Cond Env c e1 e2 B C E1 E2 A 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)" .
@@ -1354,7 +1345,7 @@
   next
     case Expr thus ?case by (iprover intro: da.Expr)
   next
-    case (Lab A B C Env c l B')  
+    case (Lab Env B c C A l B')  
     have "PROP ?Hyp Env B \<langle>c\<rangle>" .
     moreover
     have B': "B \<subseteq> B'" .
@@ -1364,7 +1355,7 @@
       by (iprover intro: da.Lab)
     thus ?case ..
   next
-    case (Comp A B C1 C2 Env c1 c2 B')
+    case (Comp Env B c1 C1 c2 C2 A B')
     have da_c1: "Env\<turnstile> B \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1" .
     have "PROP ?Hyp Env B \<langle>c1\<rangle>" .
     moreover
@@ -1383,7 +1374,7 @@
       by (iprover intro: da.Comp)
     thus ?case ..
   next
-    case (If A B C1 C2 E Env c1 c2 e B')
+    case (If Env B e E c1 C1 c2 C2 A B')
     have B': "B \<subseteq> B'" .
     obtain  E' where "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
     proof -
@@ -1417,7 +1408,7 @@
       by (iprover intro: da.If)
     thus ?case ..
   next  
-    case (Loop A B C E Env c e l B')
+    case (Loop Env B e E c C A l B')
     have B': "B \<subseteq> B'" .
     obtain  E' where "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
     proof -
@@ -1441,7 +1432,7 @@
       by (iprover intro: da.Loop )
     thus ?case ..
   next
-    case (Jmp A B Env jump B') 
+    case (Jmp jump B A Env B') 
     have B': "B \<subseteq> B'" .
     with Jmp.hyps have "jump = Ret \<longrightarrow> Result \<in> B' "
       by auto
@@ -1460,7 +1451,7 @@
   next
     case Throw thus ?case by (iprover intro: da.Throw )
   next
-    case (Try A B C C1 C2 Env c1 c2 vn B')
+    case (Try Env B c1 C1 vn C c2 C2 A B')
     have B': "B \<subseteq> B'" .
     obtain C1' where "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'"
     proof -
@@ -1485,7 +1476,7 @@
       by (iprover intro: da.Try )
     thus ?case ..
   next
-    case (Fin A B C1 C2 Env c1 c2 B')
+    case (Fin Env B c1 C1 c2 C2 A B')
     have B': "B \<subseteq> B'" .
     obtain C1' where C1': "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'"
     proof -
@@ -1519,7 +1510,7 @@
   next
     case UnOp thus ?case by (iprover intro: da.UnOp)
   next
-    case (CondAnd A B E1 E2 Env e1 e2 B')
+    case (CondAnd Env B e1 E1 e2 E2 A B')
     have B': "B \<subseteq> B'" .
     obtain E1' where "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
     proof -
@@ -1541,7 +1532,7 @@
       by (iprover intro: da.CondAnd)
     thus ?case ..
   next
-    case (CondOr A B E1 E2 Env e1 e2 B')
+    case (CondOr Env B e1 E1 e2 E2 A B')
     have B': "B \<subseteq> B'" .
     obtain E1' where "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
     proof -
@@ -1563,7 +1554,7 @@
       by (iprover intro: da.CondOr)
     thus ?case ..
   next
-    case (BinOp A B E1 Env binop e1 e2 B')
+    case (BinOp Env B e1 E1 e2 A binop B')
     have B': "B \<subseteq> B'" .
     obtain E1' where E1': "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
     proof -
@@ -1593,7 +1584,7 @@
       by auto
     thus ?case by (iprover intro: da.Super)
   next
-    case (AccLVar A B Env vn B')
+    case (AccLVar vn B A Env B')
     have "vn \<in> B" .
     moreover
     have "B \<subseteq> B'" .
@@ -1602,7 +1593,7 @@
   next
     case Acc thus ?case by (iprover intro: da.Acc)
   next 
-    case (AssLVar A B E Env e vn B')
+    case (AssLVar Env B e E A 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]) iprover
@@ -1611,7 +1602,7 @@
       by (iprover intro: da.AssLVar)
     thus ?case ..
   next
-    case (Ass A B Env V e v B') 
+    case (Ass v Env B V e A B') 
     have B': "B \<subseteq> B'" .
     have "\<forall>vn. v \<noteq> LVar vn".
     moreover
@@ -1637,7 +1628,7 @@
       by (iprover intro: da.Ass)
     thus ?case ..
   next
-    case (CondBool A B C E1 E2 Env c e1 e2 B')
+    case (CondBool Env c e1 e2 B C E1 E2 A 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'"
@@ -1673,7 +1664,7 @@
       by (iprover intro: da.CondBool)
     thus ?case ..
   next
-    case (Cond A B C E1 E2 Env c e1 e2 B')
+    case (Cond Env c e1 e2 B C E1 E2 A 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'"
@@ -1709,7 +1700,7 @@
       by (iprover intro: da.Cond)
     thus ?case ..
   next
-    case (Call A B E Env accC args e mn mode pTs statT B')
+    case (Call Env B e E args A accC statT mode mn pTs B')
     have B': "B \<subseteq> B'" .
     obtain E' where E': "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
     proof -
@@ -1735,7 +1726,7 @@
   next
     case Methd thus ?case by (iprover intro: da.Methd)
   next
-    case (Body A B C D Env c B')  
+    case (Body Env B c C A D 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 -
@@ -1763,7 +1754,7 @@
   next
     case FVar thus ?case by (iprover intro: da.FVar)
   next
-    case (AVar A B E1 Env e1 e2 B')
+    case (AVar Env B e1 E1 e2 A B')
     have B': "B \<subseteq> B'" .
     obtain E1' where E1': "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
     proof -
@@ -1789,7 +1780,7 @@
   next
     case Nil thus ?case by (iprover intro: da.Nil)
   next
-    case (Cons A B E Env e es B')
+    case (Cons Env B e E es A B')
     have B': "B \<subseteq> B'" .
     obtain E' where E': "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
     proof -
--- a/src/HOL/Bali/DefiniteAssignmentCorrect.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/DefiniteAssignmentCorrect.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -309,7 +309,7 @@
   next
     case Expr thus ?case by (elim wt_elim_cases) simp
   next
-    case (Lab c jmp s0 s1 jmps T Env) 
+    case (Lab s0 c s1 jmp jmps T Env) 
     have jmpOK: "jumpNestingOk jmps (In1r (jmp\<bullet> c))" .
     have G: "prg Env = G" .
     have wt_c: "Env\<turnstile>c\<Colon>\<surd>" 
@@ -338,7 +338,7 @@
     }
     thus ?case by simp
   next
-    case (Comp c1 c2 s0 s1 s2 jmps T Env)
+    case (Comp s0 c1 s1 c2 s2 jmps T Env)
     have jmpOk: "jumpNestingOk jmps (In1r (c1;; c2))" .
     have G: "prg Env = G" .
     from Comp.prems obtain 
@@ -363,7 +363,7 @@
       qed
     } thus ?case by simp
   next
-    case (If b c1 c2 e s0 s1 s2 jmps T Env)
+    case (If s0 e b s1 c1 c2 s2 jmps T Env)
     have jmpOk: "jumpNestingOk jmps (In1r (If(e) c1 Else c2))" .
     have G: "prg Env = G" .
     from If.prems obtain 
@@ -389,7 +389,7 @@
     }
     thus ?case by simp
   next
-    case (Loop b c e l s0 s1 s2 s3 jmps T Env)
+    case (Loop s0 e b s1 c s2 l 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" .
@@ -467,9 +467,9 @@
     }
     thus ?case by simp
   next
-    case (Jmp j s jmps T Env) thus ?case by simp
+    case (Jmp s j jmps T Env) thus ?case by simp
   next
-    case (Throw a e s0 s1 jmps T Env)
+    case (Throw s0 e a s1 jmps T Env)
     have jmpOk: "jumpNestingOk jmps (In1r (Throw e))" .
     have G: "prg Env = G" .
     from Throw.prems obtain Te where 
@@ -491,7 +491,7 @@
     }
     thus ?case by simp
   next
-    case (Try C c1 c2 s0 s1 s2 s3 vn jmps T Env)
+    case (Try s0 c1 s1 s2 C vn c2 s3 jmps T Env)
     have jmpOk: "jumpNestingOk jmps (In1r (Try c1 Catch(C vn) c2))" .
     have G: "prg Env = G" .
     from Try.prems obtain 
@@ -543,7 +543,7 @@
     }
     thus ?case by simp
   next
-    case (Fin c1 c2 s0 s1 s2 s3 x1 jmps T Env)
+    case (Fin s0 c1 x1 s1 c2 s2 s3 jmps T Env)
     have jmpOk: " jumpNestingOk jmps (In1r (c1 Finally c2))" .
     have G: "prg Env = G" .
     from Fin.prems obtain 
@@ -571,7 +571,7 @@
     }
     thus ?case by simp
   next
-    case (Init C c s0 s1 s2 s3 jmps T Env)
+    case (Init C c s0 s3 s1 s2 jmps T Env)
     have "jumpNestingOk jmps (In1r (Init C))".
     have G: "prg Env = G" .
     have "the (class G C) = c" .
@@ -636,7 +636,7 @@
     }
     thus ?case by simp
   next
-    case (NewC C a s0 s1 s2 jmps T Env)
+    case (NewC s0 C s1 a s2 jmps T Env)
     {
       fix j
       assume jmp: "abrupt s2 = Some (Jump j)"
@@ -660,7 +660,7 @@
     }
     thus ?case by simp
   next
-    case (NewA elT a e i s0 s1 s2 s3 jmps T Env)
+    case (NewA s0 elT s1 e i s2 a s3 jmps T Env)
     {
       fix j
       assume jmp: "abrupt s3 = Some (Jump j)"
@@ -692,7 +692,7 @@
     }
     thus ?case by simp
   next
-    case (Cast cT e s0 s1 s2 v jmps T Env)
+    case (Cast s0 e v s1 s2 cT jmps T Env)
     {
       fix j
       assume jmp: "abrupt s2 = Some (Jump j)"
@@ -715,7 +715,7 @@
     }
     thus ?case by simp
   next
-    case (Inst eT b e s0 s1 v jmps T Env)
+    case (Inst s0 e v s1 b eT jmps T Env)
     {
       fix j
       assume jmp: "abrupt s1 = Some (Jump j)"
@@ -734,7 +734,7 @@
   next
     case Lit thus ?case by simp
   next
-    case (UnOp e s0 s1 unop v jmps T Env)
+    case (UnOp s0 e v s1 unop jmps T Env)
     {
       fix j
       assume jmp: "abrupt s1 = Some (Jump j)"
@@ -751,7 +751,7 @@
     }
     thus ?case by simp
   next
-    case (BinOp binop e1 e2 s0 s1 s2 v1 v2 jmps T Env)
+    case (BinOp s0 e1 v1 s1 binop e2 v2 s2 jmps T Env)
     {
       fix j
       assume jmp: "abrupt s2 = Some (Jump j)"
@@ -784,7 +784,7 @@
   next
     case Super thus ?case by simp
   next
-    case (Acc f s0 s1 v va jmps T Env)
+    case (Acc s0 va v f s1 jmps T Env)
     {
       fix j
       assume jmp: "abrupt s1 = Some (Jump j)"
@@ -801,7 +801,7 @@
     }
     thus ?case by simp
   next
-    case (Ass e f s0 s1 s2 v va w jmps T Env)
+    case (Ass s0 va w f s1 e v s2 jmps T Env)
     have G: "prg Env = G" .
     from Ass.prems
     obtain vT eT where
@@ -841,7 +841,7 @@
     }
     thus ?case by simp
   next
-    case (Cond b e0 e1 e2 s0 s1 s2 v jmps T Env)
+    case (Cond s0 e0 b s1 e1 e2 v s2 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)) 
@@ -876,7 +876,7 @@
     }
     thus ?case by simp
   next
-    case (Call D a accC args e mn mode pTs s0 s1 s2 s3 s3' s4 statT v vs 
+    case (Call s0 e a s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4
                jmps T Env)
     have G: "prg Env = G" .
     from Call.prems
@@ -919,14 +919,14 @@
     }
     thus ?case by simp
   next
-    case (Methd D s0 s1 sig v jmps T Env)
+    case (Methd s0 D sig v s1 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)
+    case (Body s0 D s1 c 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)
@@ -937,7 +937,7 @@
     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)
+    case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC jmps T Env)
     have G: "prg Env = G" .
     from wf FVar.prems 
     obtain  statC f where
@@ -996,7 +996,7 @@
     }
     ultimately show ?case using v by simp
   next
-    case (AVar a e1 e2 i s0 s1 s2 s2' v jmps T Env)
+    case (AVar s0 e1 a s1 e2 i s2 v s2' jmps T Env)
     have G: "prg Env = G" .
     from AVar.prems 
     obtain  e1T e2T where
@@ -1042,7 +1042,7 @@
   next
     case Nil thus ?case by simp
   next
-    case (Cons e es s0 s1 s2 v vs jmps T Env)
+    case (Cons s0 e v s1 es vs s2 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"
@@ -1290,7 +1290,7 @@
   next
     case Lab thus ?case by simp
   next
-    case (Comp c1 c2 s0 s1 s2) 
+    case (Comp s0 c1 s1 c2 s2) 
     from Comp.hyps 
     have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
       by simp
@@ -1300,7 +1300,7 @@
       by simp
     finally show ?case by simp
   next
-    case (If b c1 c2 e s0 s1 s2)
+    case (If s0 e b s1 c1 c2 s2)
     from If.hyps 
     have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
       by simp
@@ -1310,7 +1310,7 @@
       by simp
     finally show ?case by simp
   next
-    case (Loop b c e l s0 s1 s2 s3) 
+    case (Loop s0 e b s1 c s2 l s3) 
     show ?case
     proof (cases "the_Bool b")
       case True
@@ -1334,7 +1334,7 @@
   next
     case Throw thus ?case by simp
   next
-    case (Try C c1 c2 s0 s1 s2 s3 vn)
+    case (Try s0 c1 s1 s2 C vn c2 s3)
     then
     have s0_s1: "dom (locals (store ((Norm s0)::state))) 
                   \<subseteq> dom (locals (store s1))" by simp
@@ -1361,7 +1361,7 @@
 	using False Try.hyps by simp
     qed
   next
-    case (Fin c1 c2 s0 s1 s2 s3 x1) 
+    case (Fin s0 c1 x1 s1 c2 s2 s3) 
     show ?case
     proof (cases "\<exists>err. x1 = Some (Error err)")
       case True
@@ -1384,7 +1384,7 @@
 	using Fin.hyps by simp
     qed
   next
-    case (Init C c s0 s1 s2 s3)
+    case (Init C c s0 s3 s1 s2)
     show ?case
     proof (cases "inited C (globs s0)")
       case True
@@ -1405,7 +1405,7 @@
       thus ?thesis by simp
     qed
   next
-    case (NewC C a s0 s1 s2)
+    case (NewC s0 C s1 a 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))" 
@@ -1415,7 +1415,7 @@
     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)
+    case (NewA s0 T s1 e i s2 a 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))" 
@@ -1437,7 +1437,7 @@
   next
     case UnOp thus ?case by simp
   next
-    case (BinOp binop e1 e2 s0 s1 s2 v1 v2) 
+    case (BinOp s0 e1 v1 s1 binop e2 v2 s2) 
     from BinOp.hyps
     have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))" 
       by simp
@@ -1450,7 +1450,7 @@
   next
     case Acc thus ?case by simp
   next
-    case (Ass e f s0 s1 s2 v va w)
+    case (Ass s0 va w f s1 e v s2)
     from Ass.hyps
     have s0_s1: 
       "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))" 
@@ -1486,7 +1486,7 @@
 	by simp
     qed
   next
-    case (Cond b e0 e1 e2 s0 s1 s2 v)
+    case (Cond s0 e0 b s1 e1 e2 v s2)
     from Cond.hyps 
     have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
       by simp
@@ -1496,7 +1496,7 @@
       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)
+    case (Call s0 e a' s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4)
     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))"
@@ -1512,7 +1512,7 @@
   next
     case Methd thus ?case by simp
   next
-    case (Body D c s0 s1 s2 s3)
+    case (Body s0 D s1 c s2 s3)
     from Body.hyps 
     have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
       by simp
@@ -1540,7 +1540,7 @@
       using dom_locals_lvar_mono
       by simp
   next
-    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v)
+    case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC)
     from FVar.hyps
     obtain s2': "s2' = snd (fvar statDeclC stat fn a s2)" and
              v: "v = fst (fvar statDeclC stat fn a s2)"
@@ -1572,7 +1572,7 @@
       using v_ok
       by simp
   next
-    case (AVar a e1 e2 i s0 s1 s2 s2' v)
+    case (AVar s0 e1 a s1 e2 i s2 v s2')
     from AVar.hyps
     obtain s2': "s2' = snd (avar G i a s2)" and
              v: "v   = fst (avar G i a s2)"
@@ -1599,7 +1599,7 @@
   next
     case Nil thus ?case by simp
   next
-    case (Cons e es s0 s1 s2 v vs)
+    case (Cons s0 e v s1 es vs s2)
     from Cons.hyps
     have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
       by simp
@@ -1689,7 +1689,7 @@
   next
     case NewC show ?case by simp
   next
-    case (NewA T a e i s0 s1 s2 s3)
+    case (NewA s0 T s1 e i s2 a 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 -
@@ -1706,7 +1706,7 @@
       by (rule dom_locals_halloc_mono [elim_format]) simp
     finally show ?case by simp 
   next
-    case (Cast T e s0 s1 s2 v)
+    case (Cast s0 e v s1 s2 T)
     hence "normal s1" by (cases s1,simp)
     with Cast.hyps
     have "assigns (In1l e) \<subseteq> dom (locals (store s1))"
@@ -1725,7 +1725,7 @@
   next
     case UnOp thus ?case by simp
   next
-    case (BinOp binop e1 e2 s0 s1 s2 v1 v2)
+    case (BinOp s0 e1 v1 s1 binop e2 v2 s2)
     hence "normal s1" by - (erule eval_no_abrupt_lemma [rule_format]) 
     with BinOp.hyps
     have "assigns (In1l e1) \<subseteq> dom (locals (store s1))"
@@ -1756,7 +1756,7 @@
   next
     case Acc thus ?case by simp
   next 
-    case (Ass e f s0 s1 s2 v va w)
+    case (Ass s0 va w f s1 e v s2)
     have  nrm_ass_s2: "normal (assign f v s2)" .
     hence nrm_s2: "normal s2"
       by (cases s2, simp add: assign_def Let_def)
@@ -1808,7 +1808,7 @@
       show ?thesis by (simp add: Un_assoc)
     qed
   next
-    case (Cond b e0 e1 e2 s0 s1 s2 v) 
+    case (Cond s0 e0 b s1 e1 e2 v s2) 
     hence "normal s1"
       by - (erule eval_no_abrupt_lemma [rule_format])
     with Cond.hyps
@@ -1845,7 +1845,7 @@
       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)
+    case (Call s0 e a' s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4)
     have nrm_s2: "normal s2"
     proof -
       have "normal ((set_lvars (locals (snd s2))) s4)" .
@@ -1887,7 +1887,7 @@
   next
     case LVar thus ?case by simp
   next
-    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v)
+    case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC)
     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"
@@ -1916,7 +1916,7 @@
     finally show ?case
       by simp
   next
-    case (AVar a e1 e2 i s0 s1 s2 s2' v)
+    case (AVar s0 e1 a s1 e2 i s2 v s2')
     have avar: "(v, s2') = avar G i a s2" .
     have nrm_s2: "normal s2"
     proof -
@@ -1954,7 +1954,7 @@
   next
     case Nil show ?case by simp
   next
-    case (Cons e es s0 s1 s2 v vs)
+    case (Cons s0 e v s1 es vs s2)
     have "assigns (In1l e) \<subseteq> dom (locals (store s1))"
     proof -
       from Cons
@@ -2254,19 +2254,19 @@
    proof (induct)
      case Abrupt thus ?case by simp
    next
-     case (NewC C a s0 s1 s2)
+     case (NewC s0 C s1 a 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)
+     case (NewA s0 T s1 e i s2 a 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)
+     case (Cast s0 e b s1 s2 T)
      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 -
@@ -2285,7 +2285,7 @@
        by simp
      finally show ?case by simp
    next
-     case (Inst T b e s0 s1 v)
+     case (Inst s0 e v s1 b T)
      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)
@@ -2301,7 +2301,7 @@
      thus ?case
        by simp
    next
-     case (UnOp e s0 s1 unop v)
+     case (UnOp s0 e v s1 unop)
      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)
@@ -2341,7 +2341,7 @@
        thus ?thesis by simp
      qed
    next
-     case (BinOp binop e1 e2 s0 s1 s2 v1 v2)
+     case (BinOp s0 e1 v1 s1 binop e2 v2 s2)
      have bool: "Env\<turnstile>BinOp binop e1 e2\<Colon>-PrimT Boolean" .
      show ?case
      proof (cases "constVal (BinOp binop e1 e2)") 
@@ -2507,13 +2507,13 @@
        by cases simp
      thus ?case ..
    next
-     case (Acc f s0 s1 v va)
+     case (Acc s0 va v f s1)
      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)
+     case (Ass s0 va w f s1 e v s2)
      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)" .
@@ -2522,7 +2522,7 @@
        by (rule assignsE_good_approx)
      thus ?case by simp
    next
-     case (Cond b e0 e1 e2 s0 s1 s2 v)
+     case (Cond s0 e0 b s1 e1 e2 v s2)
      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"
@@ -2599,7 +2599,7 @@
        qed
      qed
    next
-     case (Call D a accC args e mn mode pTs s0 s1 s2 s3 s3' s4 statT v vs)
+     case (Call s0 e a s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4)
      hence
      "prg Env\<turnstile>Norm s0 \<midarrow>({accC,statT,mode}e\<cdot>mn( {pTs}args))-\<succ>v\<rightarrow> 
                        (set_lvars (locals (store s2)) s4)"
@@ -2662,7 +2662,7 @@
   from eval and wt da G
   show ?thesis
   proof (induct arbitrary: Env T A)
-    case (Abrupt s t xc Env T A)
+    case (Abrupt xc s t 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
@@ -2687,14 +2687,14 @@
       by simp
     ultimately show ?case by (intro conjI)
   next
-    case (Expr e s0 s1 v Env T A)
+    case (Expr s0 e v s1 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)
+    case (Lab s0 c s1 j Env T A)
     have G: "prg Env = G" .
     from Lab.prems
     obtain C l where
@@ -2753,7 +2753,7 @@
       by (cases s1) (simp add: absorb_def)
     ultimately show ?case by (intro conjI)
   next
-    case (Comp c1 c2 s0 s1 s2 Env T A)
+    case (Comp s0 c1 s1 c2 s2 Env T A)
     have G: "prg Env = G" .
     from Comp.prems
     obtain C1 C2
@@ -2826,7 +2826,7 @@
       ultimately show ?thesis by (intro conjI)
     qed
   next
-    case (If b c1 c2 e s0 s1 s2 Env T A)
+    case (If s0 e b s1 c1 c2 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
@@ -2930,7 +2930,7 @@
       ultimately show ?thesis by simp
     qed
   next
-    case (Loop b c e l s0 s1 s2 s3 Env T A)
+    case (Loop s0 e b s1 c s2 l 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))
@@ -3118,7 +3118,7 @@
 	by simp 
     qed
   next 
-    case (Jmp j s Env T A)
+    case (Jmp s j Env T A)
     have "?NormalAssigned (Some (Jump j),s) A" by simp
     moreover
     from Jmp.prems
@@ -3136,7 +3136,7 @@
       by simp
     ultimately show ?case by (intro conjI)
   next
-    case (Throw a e s0 s1 Env T A)
+    case (Throw s0 e a 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"
@@ -3166,7 +3166,7 @@
       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)
+    case (Try s0 c1 s1 s2 C vn c2 s3 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
@@ -3313,7 +3313,7 @@
       qed
     qed
   next
-    case (Fin c1 c2 s0 s1 s2 s3 x1 Env T A)
+    case (Fin s0 c1 x1 s1 c2 s2 s3 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
@@ -3472,7 +3472,7 @@
       by simp
     ultimately show ?case by (intro conjI)
   next 
-    case (Init C c s0 s1 s2 s3 Env T A)
+    case (Init C c s0 s3 s1 s2 Env T A)
     have G: "prg Env = G" .
     from Init.hyps
     have eval: "prg Env\<turnstile> Norm s0 \<midarrow>Init C\<rightarrow> s3"
@@ -3528,7 +3528,7 @@
       ultimately show ?thesis by (intro conjI)
     qed
   next 
-    case (NewC C a s0 s1 s2 Env T A)
+    case (NewC s0 C s1 a s2 Env T A)
     have G: "prg Env = G" .
     from NewC.prems
     obtain A: "nrm A = dom (locals (store ((Norm s0)::state)))"
@@ -3568,7 +3568,7 @@
       by simp_all      
     ultimately show ?case by (intro conjI)
   next
-    case (NewA elT a e i s0 s1 s2 s3 Env T A) 
+    case (NewA s0 elT s1 e i s2 a 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"
@@ -3633,7 +3633,7 @@
       by simp_all
     ultimately show ?case by (intro conjI)
   next
-    case (Cast cT e s0 s1 s2 v Env T A)
+    case (Cast s0 e v s1 s2 cT 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"
@@ -3678,7 +3678,7 @@
       by simp_all
     ultimately show ?case by (intro conjI)
   next
-    case (Inst iT b e s0 s1 v Env T A)
+    case (Inst s0 e v s1 b iT 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"
@@ -3700,7 +3700,7 @@
       by (elim da_elim_cases) simp
     thus ?case by simp
   next
-    case (UnOp e s0 s1 unop v Env T A)
+    case (UnOp s0 e v s1 unop 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"
@@ -3716,7 +3716,7 @@
       by simp
     thus ?case by (intro conjI)
   next
-    case (BinOp binop e1 e2 s0 s1 s2 v1 v2 Env T A)
+    case (BinOp s0 e1 v1 s1 binop e2 v2 s2 Env T A)
     have G: "prg Env = G". 
     from BinOp.hyps 
     have 
@@ -3875,7 +3875,7 @@
       by (elim da_elim_cases) simp
     thus ?case by simp
   next
-    case (Acc upd s0 s1 w v Env T A)
+    case (Acc s0 v w upd s1 Env T A)
     show ?case
     proof (cases "\<exists> vn. v = LVar vn")
       case True
@@ -3905,7 +3905,7 @@
       thus ?thesis by (intro conjI)
     qed
   next
-    case (Ass e upd s0 s1 s2 v var w Env T A)
+    case (Ass s0 var w upd s1 e v s2 Env T A)
     have G: "prg Env = G" .
     from Ass.prems obtain varT eT where
       wt_var: "Env\<turnstile>var\<Colon>=varT" and
@@ -4019,7 +4019,7 @@
       by simp_all
     ultimately show ?case by (intro conjI)
   next
-    case (Cond b e0 e1 e2 s0 s1 s2 v Env T A)
+    case (Cond s0 e0 b s1 e1 e2 v s2 Env T A)
     have G: "prg Env = G" .
     have "?NormalAssigned s2 A"
     proof 
@@ -4140,7 +4140,7 @@
       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 
+    case (Call s0 e a s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4
            Env T A)
     have G: "prg Env = G" .
     have "?NormalAssigned (restore_lvars s2 s4) A"
@@ -4213,7 +4213,7 @@
       by simp_all
     ultimately show ?case by (intro conjI)
   next
-    case (Methd D s0 s1 sig v Env T A)
+    case (Methd s0 D sig v s1 Env T A)
     have G: "prg Env = G". 
     from Methd.prems obtain m where
        m:      "methd (prg Env) D sig = Some m" and
@@ -4235,7 +4235,7 @@
     ultimately show ?case
       using G by simp
   next
-    case (Body D c s0 s1 s2 s3 Env T A)
+    case (Body s0 D s1 c s2 s3 Env T A)
     have G: "prg Env = G" .
     from Body.prems 
     have nrm_A: "nrm A = dom (locals (store ((Norm s0)::state)))"
@@ -4261,7 +4261,7 @@
       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)
+    case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC Env T A)
     have G: "prg Env = G" .
     have "?NormalAssigned s3 A"
     proof 
@@ -4333,7 +4333,7 @@
       by simp_all
     ultimately show ?case by (intro conjI)
   next
-    case (AVar a e1 e2 i s0 s1 s2 s2' v Env T A)
+    case (AVar s0 e1 a s1 e2 i s2 v s2' Env T A)
     have G: "prg Env = G" .
     have "?NormalAssigned s2' A"
     proof 
@@ -4405,7 +4405,7 @@
       by (elim da_elim_cases) simp
     thus ?case by simp
   next 
-    case (Cons e es s0 s1 s2 v vs Env T A)
+    case (Cons s0 e v s1 es vs s2 Env T A)
     have G: "prg Env = G" .
     have "?NormalAssigned s2 A"
     proof 
--- a/src/HOL/Bali/Eval.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/Eval.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -474,104 +474,75 @@
        
 section "evaluation judgments"
 
-consts
-  eval   :: "prog \<Rightarrow> (state \<times> term    \<times> vals \<times> state) set"
-  halloc::  "prog \<Rightarrow> (state \<times> obj_tag \<times> loc  \<times> state) set"
-  sxalloc:: "prog \<Rightarrow> (state                  \<times> state) set"
-
-
-syntax
-eval ::"[prog,state,term,vals*state]=>bool"("_|-_ -_>-> _"  [61,61,80,   61]60)
-exec ::"[prog,state,stmt      ,state]=>bool"("_|-_ -_-> _"   [61,61,65,   61]60)
-evar ::"[prog,state,var  ,vvar,state]=>bool"("_|-_ -_=>_-> _"[61,61,90,61,61]60)
-eval_::"[prog,state,expr ,val, state]=>bool"("_|-_ -_->_-> _"[61,61,80,61,61]60)
-evals::"[prog,state,expr list ,
-		    val  list ,state]=>bool"("_|-_ -_#>_-> _"[61,61,61,61,61]60)
-hallo::"[prog,state,obj_tag,
-	             loc,state]=>bool"("_|-_ -halloc _>_-> _"[61,61,61,61,61]60)
-sallo::"[prog,state        ,state]=>bool"("_|-_ -sxalloc-> _"[61,61,      61]60)
-
-syntax (xsymbols)
-eval ::"[prog,state,term,vals\<times>state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>_\<succ>\<rightarrow> _"  [61,61,80,   61]60)
-exec ::"[prog,state,stmt      ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<rightarrow> _"   [61,61,65,   61]60)
-evar ::"[prog,state,var  ,vvar,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_=\<succ>_\<rightarrow> _"[61,61,90,61,61]60)
-eval_::"[prog,state,expr ,val ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_-\<succ>_\<rightarrow> _"[61,61,80,61,61]60)
-evals::"[prog,state,expr list ,
-		    val  list ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
-hallo::"[prog,state,obj_tag,
-	             loc,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>halloc _\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
-sallo::"[prog,state,        state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>sxalloc\<rightarrow> _"[61,61,      61]60)
-
-translations
-  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow>  w___s' " == "(s,t,w___s') \<in> eval G"
-  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow> (w,  s')" <= "(s,t,w,  s') \<in> eval G"
-  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow> (w,x,s')" <= "(s,t,w,x,s') \<in> eval G"
-  "G\<turnstile>s \<midarrow>c    \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In1r c\<succ>\<rightarrow> (\<diamondsuit>,x,s')"
-  "G\<turnstile>s \<midarrow>c    \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In1r c\<succ>\<rightarrow> (\<diamondsuit>  ,  s')"
-  "G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In1l e\<succ>\<rightarrow> (In1 v ,x,s')"
-  "G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In1l e\<succ>\<rightarrow> (In1 v ,  s')"
-  "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In2  e\<succ>\<rightarrow> (In2 vf,x,s')"
-  "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In2  e\<succ>\<rightarrow> (In2 vf,  s')"
-  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In3  e\<succ>\<rightarrow> (In3 v ,x,s')"
-  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In3  e\<succ>\<rightarrow> (In3 v ,  s')"
-  "G\<turnstile>s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,s')" <= "(s,oi,a,x,s') \<in> halloc G"
-  "G\<turnstile>s \<midarrow>halloc oi\<succ>a\<rightarrow>    s' " == "(s,oi,a,  s') \<in> halloc G"
-  "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>     (x,s')" <= "(s     ,x,s') \<in> sxalloc G"
-  "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>        s' " == "(s     ,  s') \<in> sxalloc G"
-
-inductive "halloc G" intros --{* allocating objects on the heap, cf. 12.5 *}
+inductive2
+  halloc :: "[prog,state,obj_tag,loc,state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>halloc _\<succ>_\<rightarrow> _"[61,61,61,61,61]60) for G::prog
+where --{* allocating objects on the heap, cf. 12.5 *}
 
   Abrupt: 
   "G\<turnstile>(Some x,s) \<midarrow>halloc oi\<succ>arbitrary\<rightarrow> (Some x,s)"
 
-  New:	  "\<lbrakk>new_Addr (heap s) = Some a; 
+| New:	  "\<lbrakk>new_Addr (heap s) = Some a; 
 	    (x,oi') = (if atleast_free (heap s) (Suc (Suc 0)) then (None,oi)
 		       else (Some (Xcpt (Loc a)),CInst (SXcpt OutOfMemory)))\<rbrakk>
             \<Longrightarrow>
 	    G\<turnstile>Norm s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,init_obj G oi' (Heap a) s)"
 
-inductive "sxalloc G" intros --{* allocating exception objects for
-	 	 	      standard exceptions (other than OutOfMemory) *}
+inductive2 sxalloc :: "[prog,state,state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>sxalloc\<rightarrow> _"[61,61,61]60) for G::prog
+where --{* 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)"
+| 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)"
+| 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)"
+| 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>
+| SXcpt: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>halloc (CInst (SXcpt xn))\<succ>a\<rightarrow> (x,s1)\<rbrakk> \<Longrightarrow>
 	  G\<turnstile>(Some (Xcpt (Std xn)),s0) \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s1)"
 
 
-inductive "eval G" intros
+inductive2
+  eval :: "[prog,state,term,vals,state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>_\<succ>\<rightarrow> '(_, _')"  [61,61,80,0,0]60)
+  and exec ::"[prog,state,stmt      ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<rightarrow> _"   [61,61,65,   61]60)
+  and evar ::"[prog,state,var  ,vvar,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_=\<succ>_\<rightarrow> _"[61,61,90,61,61]60)
+  and eval'::"[prog,state,expr ,val ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_-\<succ>_\<rightarrow> _"[61,61,80,61,61]60)
+  and evals::"[prog,state,expr list ,
+		    val  list ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
+  for G::prog
+where
+
+  "G\<turnstile>s \<midarrow>c    \<rightarrow>     s' \<equiv> G\<turnstile>s \<midarrow>In1r c\<succ>\<rightarrow> (\<diamondsuit>,  s')"
+| "G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow>     s' \<equiv> G\<turnstile>s \<midarrow>In1l e\<succ>\<rightarrow> (In1 v,  s')"
+| "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow>     s' \<equiv> G\<turnstile>s \<midarrow>In2  e\<succ>\<rightarrow> (In2 vf, s')"
+| "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow>     s' \<equiv> G\<turnstile>s \<midarrow>In3  e\<succ>\<rightarrow> (In3 v,  s')"
 
 --{* propagation of abrupt completion *}
 
   --{* cf. 14.1, 15.5 *}
-  Abrupt: 
-   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (arbitrary3 t,(Some xc,s))"
+| Abrupt: 
+   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (arbitrary3 t, (Some xc, s))"
 
 
 --{* execution of statements *}
 
   --{* cf. 14.5 *}
-  Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<rightarrow> Norm s"
+| Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<rightarrow> Norm s"
 
   --{* cf. 14.7 *}
-  Expr:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+| 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>
+| 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 *}
-  Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<rightarrow> s1;
+| Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<rightarrow> s1;
 	  G\<turnstile>     s1 \<midarrow>c2 \<rightarrow> s2\<rbrakk> \<Longrightarrow>
 				 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<rightarrow> s2"
 
   --{* cf. 14.8.2 *}
-  If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
+| 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"
 
@@ -583,26 +554,26 @@
      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;
+| Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
 	  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"
 
-  Jmp: "G\<turnstile>Norm s \<midarrow>Jmp j\<rightarrow> (Some (Jump j), s)"
+| Jmp: "G\<turnstile>Norm s \<midarrow>Jmp j\<rightarrow> (Some (Jump j), s)"
    
   --{* cf. 14.16 *}
-  Throw: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+| Throw: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 				 G\<turnstile>Norm s0 \<midarrow>Throw e\<rightarrow> abupd (throw a') s1"
 
   --{* cf. 14.18.1 *}
-  Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; 
+| Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; 
 	  if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2\<rbrakk> \<Longrightarrow>
 		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(C vn) c2\<rightarrow> s3"
 
   --{* cf. 14.18.2 *}
-  Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1,s1);
+| 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)) 
               then (x1,s1) 
@@ -610,7 +581,7 @@
           \<Longrightarrow>
           G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<rightarrow> s3"
   --{* cf. 12.4.2, 8.5 *}
-  Init:	"\<lbrakk>the (class G C) = c;
+| Init:	"\<lbrakk>the (class G C) = c;
 	  if inited C (globs s0) then s3 = Norm s0
 	  else (G\<turnstile>Norm (init_class_obj G C s0) 
 		  \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1 \<and>
@@ -642,51 +613,51 @@
 --{* evaluation of expressions *}
 
   --{* cf. 15.8.1, 12.4.1 *}
-  NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1;
+| NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1;
 	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
 
   --{* cf. 15.9.1, 12.4.1 *}
-  NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<rightarrow> s2; 
+| NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<rightarrow> s2; 
 	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
 	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<rightarrow> s3"
 
   --{* cf. 15.15 *}
-  Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
+| Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
 	  s2 = abupd (raise_if (\<not>G,store s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
 			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<rightarrow> s2"
 
   --{* cf. 15.19.2 *}
-  Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
+| Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
 	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
 			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<rightarrow> s1"
 
   --{* cf. 15.7.1 *}
-  Lit:	"G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<rightarrow> Norm s"
+| 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> 
+| UnOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> 
          \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>(eval_unop unop v)\<rightarrow> s1"
  
-  BinOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1; 
+| BinOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1; 
           G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then (In1l e2) else (In1r Skip))
-                \<succ>\<rightarrow> (In1 v2,s2)
+                \<succ>\<rightarrow> (In1 v2, s2)
           \<rbrakk> 
          \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<rightarrow> s2"
    
   --{* cf. 15.10.2 *}
-  Super: "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<rightarrow> Norm s"
+| Super: "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<rightarrow> Norm s"
 
   --{* cf. 15.2 *}
-  Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+| Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<rightarrow> s1"
 
   --{* cf. 15.25.1 *}
-  Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<rightarrow> s1;
+| Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<rightarrow> s1;
           G\<turnstile>     s1 \<midarrow>e-\<succ>v  \<rightarrow> s2\<rbrakk> \<Longrightarrow>
 				   G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v\<rightarrow> assign f v s2"
 
   --{* cf. 15.24 *}
-  Cond:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1;
+| 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"
 
@@ -710,7 +681,7 @@
       \end{itemize}
    *}
   --{* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5 *}
-  Call:	
+| 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>;
     s3=init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs s2;
@@ -723,10 +694,10 @@
    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>
+| 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"
   
-  Body:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<rightarrow> s2; 
+| 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 
@@ -742,10 +713,10 @@
 --{* 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"
+| LVar:	"G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<rightarrow> Norm s"
 
   --{* cf. 15.10.1, 12.4.1 *}
-  FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2;
+| FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2;
 	  (v,s2') = fvar statDeclC stat fn a s2;
           s3 = check_field_access G accC statDeclC fn stat a s2' \<rbrakk> \<Longrightarrow>
 	  G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v\<rightarrow> s3"
@@ -755,7 +726,7 @@
   *}
 
   --{* 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;
+| 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'"
 
@@ -763,11 +734,11 @@
 --{* evaluation of expression lists *}
 
   --{* cf. 15.11.4.2 *}
-  Nil:
+| Nil:
 				    "G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<rightarrow> Norm s0"
 
   --{* cf. 15.6.4 *}
-  Cons:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<rightarrow> s1;
+| 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"
 
@@ -794,17 +765,17 @@
 
 declare split_if     [split del] split_if_asm     [split del] 
         option.split [split del] option.split_asm [split del]
-inductive_cases halloc_elim_cases: 
+inductive_cases2 halloc_elim_cases: 
   "G\<turnstile>(Some xc,s) \<midarrow>halloc oi\<succ>a\<rightarrow> s'"
   "G\<turnstile>(Norm    s) \<midarrow>halloc oi\<succ>a\<rightarrow> s'"
 
-inductive_cases sxalloc_elim_cases:
+inductive_cases2 sxalloc_elim_cases:
  	"G\<turnstile> Norm                 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'"
+inductive_cases2 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;  
@@ -822,40 +793,40 @@
 ML_setup {*
 change_simpset (fn ss => ss delloop "split_all_tac");
 *}
-inductive_cases eval_cases: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> vs'"
+inductive_cases2 eval_cases: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v, s')"
 
-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 (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'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Lit w)                        \<succ>\<rightarrow> vs'"
-        "G\<turnstile>Norm s \<midarrow>In1l (UnOp unop e)                  \<succ>\<rightarrow> vs'"
-        "G\<turnstile>Norm s \<midarrow>In1l (BinOp binop e1 e2)            \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In2  (LVar vn)                      \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Cast T e)                     \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (e InstOf T)                   \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Super)                        \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Acc va)                       \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (Expr e)                       \<succ>\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (c1;; c2)                      \<succ>\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Methd C sig)                  \<succ>\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Body D c)                     \<succ>\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2)                 \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2)             \<succ>\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> While(e) c)                \<succ>\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2)                \<succ>\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (Throw e)                      \<succ>\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (NewC C)                       \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (New T[e])                     \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Ass va e)                     \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2)       \<succ>\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In2  ({accC,statDeclC,stat}e..fn)   \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                      \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l ({accC,statT,mode}e\<cdot>mn({pT}p)) \<succ>\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (Init C)                       \<succ>\<rightarrow> xs'"
+inductive_cases2 eval_elim_cases [cases set]:
+        "G\<turnstile>(Some xc,s) \<midarrow>t                              \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r Skip                           \<succ>\<rightarrow> (x, s')"
+        "G\<turnstile>Norm s \<midarrow>In1r (Jmp j)                        \<succ>\<rightarrow> (x, s')"
+        "G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> c)                         \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In3  ([])                           \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In3  (e#es)                         \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Lit w)                        \<succ>\<rightarrow> (v, s')"
+        "G\<turnstile>Norm s \<midarrow>In1l (UnOp unop e)                  \<succ>\<rightarrow> (v, s')"
+        "G\<turnstile>Norm s \<midarrow>In1l (BinOp binop e1 e2)            \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In2  (LVar vn)                      \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Cast T e)                     \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (e InstOf T)                   \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Super)                        \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Acc va)                       \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (Expr e)                       \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (c1;; c2)                      \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Methd C sig)                  \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Body D c)                     \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2)                 \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2)             \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> While(e) c)                \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2)                \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (Throw e)                      \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (NewC C)                       \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (New T[e])                     \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Ass va e)                     \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2)       \<succ>\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In2  ({accC,statDeclC,stat}e..fn)   \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                      \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l ({accC,statT,mode}e\<cdot>mn({pT}p)) \<succ>\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (Init C)                       \<succ>\<rightarrow> (x, s')"
 declare not_None_eq [simp]  (* IntDef.Zero_def [simp] *)
 declare split_paired_All [simp] split_paired_Ex [simp]
 ML_setup {*
@@ -885,25 +856,33 @@
  (injection @{term In1l} into terms) always evaluates to ordinary values 
  (injection @{term In1} into generalised values @{term vals}). 
 *}
+
+lemma eval_expr_eq: "G\<turnstile>s \<midarrow>In1l t\<succ>\<rightarrow> (w, s') = (\<exists>v. w=In1 v \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<rightarrow> s')"
+  by (auto, frule eval_Inj_elim, auto)
+
+lemma eval_var_eq: "G\<turnstile>s \<midarrow>In2 t\<succ>\<rightarrow> (w, s') = (\<exists>vf. w=In2 vf \<and> G\<turnstile>s \<midarrow>t=\<succ>vf\<rightarrow> s')"
+  by (auto, frule eval_Inj_elim, auto)
+
+lemma eval_exprs_eq: "G\<turnstile>s \<midarrow>In3 t\<succ>\<rightarrow> (w, s') = (\<exists>vs. w=In3 vs \<and> G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs\<rightarrow> s')"
+  by (auto, frule eval_Inj_elim, auto)
+
+lemma eval_stmt_eq: "G\<turnstile>s \<midarrow>In1r t\<succ>\<rightarrow> (w, s') = (w=\<diamondsuit> \<and> G\<turnstile>s \<midarrow>t \<rightarrow> s')"
+  by (auto, frule eval_Inj_elim, auto, frule eval_Inj_elim, auto)
+
 ML_setup {*
-fun eval_fun nam inj rhs =
+fun eval_fun name lhs =
 let
-  val name = "eval_" ^ nam ^ "_eq"
-  val lhs = "G\<turnstile>s \<midarrow>" ^ inj ^ " t\<succ>\<rightarrow> (w, s')"
-  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
-	(K [Auto_tac, ALLGOALS (ftac (thm "eval_Inj_elim")) THEN Auto_tac])
   fun is_Inj (Const (inj,_) $ _) = true
     | is_Inj _                   = false
-  fun pred (_ $ (Const ("Pair",_) $ _ $ 
-      (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ x $ _ ))) $ _ ) = is_Inj x
+  fun pred (_ $ _ $ _ $ _ $ x $ _) = is_Inj x
 in
   cond_simproc name lhs pred (thm name)
 end
 
-val eval_expr_proc =eval_fun "expr" "In1l" "\<exists>v.  w=In1 v   \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<rightarrow> s'"
-val eval_var_proc  =eval_fun "var"  "In2"  "\<exists>vf. w=In2 vf  \<and> G\<turnstile>s \<midarrow>t=\<succ>vf\<rightarrow> s'"
-val eval_exprs_proc=eval_fun "exprs""In3"  "\<exists>vs. w=In3 vs  \<and> G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs\<rightarrow> s'"
-val eval_stmt_proc =eval_fun "stmt" "In1r" "     w=\<diamondsuit> \<and> G\<turnstile>s \<midarrow>t    \<rightarrow> s'";
+val eval_expr_proc  = eval_fun "eval_expr_eq"  "G\<turnstile>s \<midarrow>In1l t\<succ>\<rightarrow> (w, s')";
+val eval_var_proc   = eval_fun "eval_var_eq"   "G\<turnstile>s \<midarrow>In2 t\<succ>\<rightarrow> (w, s')";
+val eval_exprs_proc = eval_fun "eval_exprs_eq" "G\<turnstile>s \<midarrow>In3 t\<succ>\<rightarrow> (w, s')";
+val eval_stmt_proc  = eval_fun "eval_stmt_eq"  "G\<turnstile>s \<midarrow>In1r t\<succ>\<rightarrow> (w, s')";
 Addsimprocs [eval_expr_proc,eval_var_proc,eval_exprs_proc,eval_stmt_proc];
 bind_thms ("AbruptIs", sum3_instantiate (thm "eval.Abrupt"))
 *}
@@ -922,9 +901,7 @@
     assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s')" and
          normal: "normal s" and
          callee: "t=In1l (Callee l e)"
-    then have "False"
-    proof (induct)
-    qed (auto)
+    then have "False" by induct auto
   }  
   then show ?thesis
     by (cases s') fastsimp
@@ -937,9 +914,7 @@
     assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s')" and
          normal: "normal s" and
          callee: "t=In1l (InsInitE c e)"
-    then have "False"
-    proof (induct)
-    qed (auto)
+    then have "False" by induct auto
   }
   then show ?thesis
     by (cases s') fastsimp
@@ -951,9 +926,7 @@
     assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s')" and
          normal: "normal s" and
          callee: "t=In2 (InsInitV c w)"
-    then have "False"
-    proof (induct)
-    qed (auto)
+    then have "False" by induct auto
   }  
   then show ?thesis
     by (cases s') fastsimp
@@ -965,9 +938,7 @@
     assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s')" and
          normal: "normal s" and
          callee: "t=In1r (FinA a c)"
-    then have "False"
-    proof (induct)
-    qed (auto)
+    then have "False" by induct auto
   }  
   then show ?thesis
     by (cases s') fastsimp 
@@ -988,8 +959,7 @@
 local
   fun is_None (Const ("Datatype.option.None",_)) = true
     | is_None _ = false
-  fun pred (t as (_ $ (Const ("Pair",_) $
-     (Const ("Pair", _) $ x $ _) $ _ ) $ _)) = is_None x
+  fun pred (_ $ _ $ (Const ("Pair", _) $ x $ _) $ _  $ _ $ _) = is_None x
 in
   val eval_no_abrupt_proc = 
   cond_simproc "eval_no_abrupt" "G\<turnstile>(x,s) \<midarrow>e\<succ>\<rightarrow> (w,Norm s')" pred 
@@ -1015,9 +985,7 @@
 local
   fun is_Some (Const ("Pair",_) $ (Const ("Datatype.option.Some",_) $ _)$ _) =true
     | is_Some _ = false
-  fun pred (_ $ (Const ("Pair",_) $
-     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
-       x))) $ _ ) = is_Some x
+  fun pred (_ $ _ $ _ $ _ $ _ $ x) = is_Some x
 in
   val eval_abrupt_proc = 
   cond_simproc "eval_abrupt" 
@@ -1174,8 +1142,7 @@
 section "single valued"
 
 lemma unique_halloc [rule_format (no_asm)]: 
-  "\<And>s as as'. (s,oi,as)\<in>halloc G \<Longrightarrow> (s,oi,as')\<in>halloc G \<longrightarrow> as'=as"
-apply (simp (no_asm_simp) only: split_tupled_all)
+  "G\<turnstile>s \<midarrow>halloc oi\<succ>a \<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>halloc oi\<succ>a' \<rightarrow> s'' \<longrightarrow> a' = a \<and> s'' = s'"
 apply (erule halloc.induct)
 apply  (auto elim!: halloc_elim_cases split del: split_if split_if_asm)
 apply (drule trans [THEN sym], erule sym) 
@@ -1192,8 +1159,7 @@
 
 
 lemma unique_sxalloc [rule_format (no_asm)]: 
-  "\<And>s s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'' \<longrightarrow> s'' = s'"
-apply (simp (no_asm_simp) only: split_tupled_all)
+  "G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'' \<longrightarrow> s'' = s'"
 apply (erule sxalloc.induct)
 apply   (auto dest: unique_halloc elim!: sxalloc_elim_cases 
               split del: split_if split_if_asm)
@@ -1210,9 +1176,7 @@
 
 
 lemma unique_eval [rule_format (no_asm)]: 
-  "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws \<Longrightarrow> (\<forall>ws'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws' \<longrightarrow> ws' = ws)"
-apply (case_tac "ws")
-apply hypsubst
+  "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w, s') \<Longrightarrow> (\<forall>w' s''. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w', s'') \<longrightarrow> w' = w \<and> s'' = s')"
 apply (erule eval_induct)
 apply (tactic {* ALLGOALS (EVERY'
       [strip_tac, rotate_tac ~1, eresolve_tac (thms "eval_elim_cases")]) *})
@@ -1221,19 +1185,16 @@
 apply (simp (no_asm_use) only: split add: split_if_asm)
 (* 34 subgoals *)
 prefer 30 (* Init *)
-apply (case_tac "inited C (globs s0)", (simp only: if_True if_False)+)
+apply (case_tac "inited C (globs s0)", (simp only: if_True if_False simp_thms)+)
 prefer 26 (* While *)
 apply (simp (no_asm_use) only: split add: split_if_asm, blast)
-apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp)
-apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp)
-apply blast
-(* 33 subgoals *)
+(* 36 subgoals *)
 apply (blast dest: unique_sxalloc unique_halloc split_pairD)+
 done
 
 (* unused *)
 lemma single_valued_eval: 
- "single_valued {((s,t),vs'). G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> vs'}"
+ "single_valued {((s, t), (v, s')). G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v, s')}"
 apply (unfold single_valued_def)
 by (clarify, drule (1) unique_eval, auto)
 
--- a/src/HOL/Bali/Evaln.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/Evaln.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -28,68 +28,40 @@
 for welltyped, and definitely assigned terms.
 *}
 
-consts
-
-  evaln	:: "prog \<Rightarrow> (state \<times> term \<times> nat \<times> vals \<times> state) set"
-
-syntax
-
-  evaln	:: "[prog, state, term,        nat, vals * state] => bool"
-				("_|-_ -_>-_-> _"   [61,61,80,   61,61] 60)
-  evarn	:: "[prog, state, var  , vvar        , nat, state] => bool"
-				("_|-_ -_=>_-_-> _" [61,61,90,61,61,61] 60)
-  eval_n:: "[prog, state, expr , val         , nat, state] => bool"
-				("_|-_ -_->_-_-> _" [61,61,80,61,61,61] 60)
-  evalsn:: "[prog, state, expr list, val list, nat, state] => bool"
-				("_|-_ -_#>_-_-> _" [61,61,61,61,61,61] 60)
-  execn	:: "[prog, state, stmt ,               nat, state] => bool"
-				("_|-_ -_-_-> _"    [61,61,65,   61,61] 60)
-
-syntax (xsymbols)
+inductive2
+  evaln	:: "[prog, state, term, nat, vals, state] \<Rightarrow> bool"
+    ("_\<turnstile>_ \<midarrow>_\<succ>\<midarrow>_\<rightarrow> '(_, _')" [61,61,80,61,0,0] 60)
+  and evarn :: "[prog, state, var, vvar, nat, state] \<Rightarrow> bool"
+    ("_\<turnstile>_ \<midarrow>_=\<succ>_\<midarrow>_\<rightarrow> _" [61,61,90,61,61,61] 60)
+  and eval_n:: "[prog, state, expr, val, nat, state] \<Rightarrow> bool"
+    ("_\<turnstile>_ \<midarrow>_-\<succ>_\<midarrow>_\<rightarrow> _" [61,61,80,61,61,61] 60)
+  and evalsn :: "[prog, state, expr list, val  list, nat, state] \<Rightarrow> bool"
+    ("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<midarrow>_\<rightarrow> _" [61,61,61,61,61,61] 60)
+  and execn	:: "[prog, state, stmt, nat, state] \<Rightarrow> bool"
+    ("_\<turnstile>_ \<midarrow>_\<midarrow>_\<rightarrow> _"     [61,61,65,   61,61] 60)
+  for G :: prog
+where
 
-  evaln	:: "[prog, state, term,         nat, vals \<times> state] \<Rightarrow> bool"
-				("_\<turnstile>_ \<midarrow>_\<succ>\<midarrow>_\<rightarrow> _"   [61,61,80,   61,61] 60)
-  evarn	:: "[prog, state, var  , vvar         , nat, state] \<Rightarrow> bool"
-				("_\<turnstile>_ \<midarrow>_=\<succ>_\<midarrow>_\<rightarrow> _" [61,61,90,61,61,61] 60)
-  eval_n:: "[prog, state, expr , val ,          nat, state] \<Rightarrow> bool"
-				("_\<turnstile>_ \<midarrow>_-\<succ>_\<midarrow>_\<rightarrow> _" [61,61,80,61,61,61] 60)
-  evalsn:: "[prog, state, expr list, val  list, nat, state] \<Rightarrow> bool"
-				("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<midarrow>_\<rightarrow> _" [61,61,61,61,61,61] 60)
-  execn	:: "[prog, state, stmt ,                nat, state] \<Rightarrow> bool"
-				("_\<turnstile>_ \<midarrow>_\<midarrow>_\<rightarrow> _"     [61,61,65,   61,61] 60)
-
-translations
-
-  "G\<turnstile>s \<midarrow>t    \<succ>\<midarrow>n\<rightarrow>  w___s' " == "(s,t,n,w___s') \<in> evaln G"
-  "G\<turnstile>s \<midarrow>t    \<succ>\<midarrow>n\<rightarrow> (w,  s')" <= "(s,t,n,w,  s') \<in> evaln G"
-  "G\<turnstile>s \<midarrow>t    \<succ>\<midarrow>n\<rightarrow> (w,x,s')" <= "(s,t,n,w,x,s') \<in> evaln G"
-  "G\<turnstile>s \<midarrow>c     \<midarrow>n\<rightarrow> (x,s')" <= "G\<turnstile>s \<midarrow>In1r  c\<succ>\<midarrow>n\<rightarrow> (\<diamondsuit>    ,x,s')"
-  "G\<turnstile>s \<midarrow>c     \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In1r  c\<succ>\<midarrow>n\<rightarrow> (\<diamondsuit>    ,  s')"
-  "G\<turnstile>s \<midarrow>e-\<succ>v  \<midarrow>n\<rightarrow> (x,s')" <= "G\<turnstile>s \<midarrow>In1l e\<succ>\<midarrow>n\<rightarrow> (In1 v ,x,s')"
-  "G\<turnstile>s \<midarrow>e-\<succ>v  \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In1l e\<succ>\<midarrow>n\<rightarrow> (In1 v ,  s')"
-  "G\<turnstile>s \<midarrow>e=\<succ>vf \<midarrow>n\<rightarrow> (x,s')" <= "G\<turnstile>s \<midarrow>In2  e\<succ>\<midarrow>n\<rightarrow> (In2 vf,x,s')"
-  "G\<turnstile>s \<midarrow>e=\<succ>vf \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In2  e\<succ>\<midarrow>n\<rightarrow> (In2 vf,  s')"
-  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v  \<midarrow>n\<rightarrow> (x,s')" <= "G\<turnstile>s \<midarrow>In3  e\<succ>\<midarrow>n\<rightarrow> (In3 v ,x,s')"
-  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v  \<midarrow>n\<rightarrow>    s' " == "G\<turnstile>s \<midarrow>In3  e\<succ>\<midarrow>n\<rightarrow> (In3 v ,  s')"
-
-
-inductive "evaln G" intros
+  "G\<turnstile>s \<midarrow>c     \<midarrow>n\<rightarrow>    s' \<equiv> G\<turnstile>s \<midarrow>In1r  c\<succ>\<midarrow>n\<rightarrow> (\<diamondsuit>    ,  s')"
+| "G\<turnstile>s \<midarrow>e-\<succ>v  \<midarrow>n\<rightarrow>    s' \<equiv> G\<turnstile>s \<midarrow>In1l e\<succ>\<midarrow>n\<rightarrow> (In1 v ,  s')"
+| "G\<turnstile>s \<midarrow>e=\<succ>vf \<midarrow>n\<rightarrow>    s' \<equiv> G\<turnstile>s \<midarrow>In2  e\<succ>\<midarrow>n\<rightarrow> (In2 vf,  s')"
+| "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v  \<midarrow>n\<rightarrow>    s' \<equiv> G\<turnstile>s \<midarrow>In3  e\<succ>\<midarrow>n\<rightarrow> (In3 v ,  s')"
 
 --{* propagation of abrupt completion *}
 
-  Abrupt:   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (arbitrary3 t,(Some xc,s))"
+| Abrupt:   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (arbitrary3 t,(Some xc,s))"
 
 
 --{* evaluation of variables *}
 
-  LVar:	"G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<midarrow>n\<rightarrow> Norm s"
+| 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;
+| 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; 
+| AVar:	"\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1 ; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2; 
 	  (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
 	              G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<midarrow>n\<rightarrow> s2'"
 
@@ -98,46 +70,46 @@
 
 --{* evaluation of expressions *}
 
-  NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1;
+| NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1;
 	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<midarrow>n\<rightarrow> s2"
 
-  NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<midarrow>n\<rightarrow> s2; 
+| NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<midarrow>n\<rightarrow> s2; 
 	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
 	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<midarrow>n\<rightarrow> s3"
 
-  Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
+| Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
 	  s2 = abupd (raise_if (\<not>G,snd s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
 			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<midarrow>n\<rightarrow> s2"
 
-  Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
+| Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
 	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
 			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<midarrow>n\<rightarrow> s1"
 
-  Lit:			   "G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<midarrow>n\<rightarrow> Norm s"
+| Lit:			   "G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<midarrow>n\<rightarrow> Norm s"
 
-  UnOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> 
+| UnOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> 
          \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>(eval_unop unop v)\<midarrow>n\<rightarrow> s1"
 
-  BinOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<midarrow>n\<rightarrow> s1; 
+| BinOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<midarrow>n\<rightarrow> s1; 
            G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then (In1l e2) else (In1r Skip))
             \<succ>\<midarrow>n\<rightarrow> (In1 v2,s2)\<rbrakk> 
          \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<midarrow>n\<rightarrow> s2"
 
-  Super:		   "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<midarrow>n\<rightarrow> Norm s"
+| Super:		   "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<midarrow>n\<rightarrow> Norm s"
 
-  Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+| Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<midarrow>n\<rightarrow> s1"
 
-  Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<midarrow>n\<rightarrow> s1;
+| Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<midarrow>n\<rightarrow> s1;
           G\<turnstile>     s1 \<midarrow>e-\<succ>v     \<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 				   G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v\<midarrow>n\<rightarrow> assign f v s2"
 
-  Cond:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1;
+| Cond:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1;
           G\<turnstile>     s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 			    G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<midarrow>n\<rightarrow> s2"
 
-  Call:	
+| 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>; 
     s3=init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs s2;
@@ -147,10 +119,10 @@
    \<Longrightarrow> 
     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>
+| 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;
+| 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 
@@ -160,57 +132,57 @@
 
 --{* evaluation of expression lists *}
 
-  Nil:
+| Nil:
 				"G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<midarrow>n\<rightarrow> Norm s0"
 
-  Cons:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<midarrow>n\<rightarrow> s1;
+| Cons:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<midarrow>n\<rightarrow> s1;
           G\<turnstile>     s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 			     G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<midarrow>n\<rightarrow> s2"
 
 
 --{* execution of statements *}
 
-  Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<midarrow>n\<rightarrow> Norm s"
+| Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<midarrow>n\<rightarrow> Norm s"
 
-  Expr:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+| Expr:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 				  G\<turnstile>Norm s0 \<midarrow>Expr e\<midarrow>n\<rightarrow> s1"
 
-  Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+| Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
                              G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<midarrow>n\<rightarrow> abupd (absorb l) s1"
 
-  Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1;
+| Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1;
 	  G\<turnstile>     s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 				 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<midarrow>n\<rightarrow> s2"
 
-  If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
+| If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
 	  G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
 		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<midarrow>n\<rightarrow> s2"
 
-  Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
+| Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
 	  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"
   
-  Jmp: "G\<turnstile>Norm s \<midarrow>Jmp 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>
+| Throw:"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
 				 G\<turnstile>Norm s0 \<midarrow>Throw e\<midarrow>n\<rightarrow> abupd (throw a') s1"
 
-  Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2;
+| Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2;
 	  if G,s2\<turnstile>catch tn then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3 else s3 = s2\<rbrakk>
           \<Longrightarrow>
 		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(tn vn) c2\<midarrow>n\<rightarrow> s3"
 
-  Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> (x1,s1);
+| Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> (x1,s1);
 	  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;
+| Init:	"\<lbrakk>the (class G C) = c;
 	  if inited C (globs s0) then s3 = Norm s0
 	  else (G\<turnstile>Norm (init_class_obj G C s0)
 	          \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1 \<and>
@@ -229,41 +201,41 @@
 ML_setup {*
 change_simpset (fn ss => ss delloop "split_all_tac");
 *}
-inductive_cases evaln_cases: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> vs'"
+inductive_cases2 evaln_cases: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
 
-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 (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'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Lit w)                   \<succ>\<midarrow>n\<rightarrow> vs'"
-        "G\<turnstile>Norm s \<midarrow>In1l (UnOp unop e)             \<succ>\<midarrow>n\<rightarrow> vs'"
-        "G\<turnstile>Norm s \<midarrow>In1l (BinOp binop e1 e2)       \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In2  (LVar vn)                 \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Cast T e)                \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (e InstOf T)              \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Super)                   \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Acc va)                  \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (Expr e)                  \<succ>\<midarrow>n\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (c1;; c2)                 \<succ>\<midarrow>n\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Methd C sig)             \<succ>\<midarrow>n\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Body D c)                \<succ>\<midarrow>n\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2)            \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2)        \<succ>\<midarrow>n\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> While(e) c)           \<succ>\<midarrow>n\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2)           \<succ>\<midarrow>n\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (Throw e)                 \<succ>\<midarrow>n\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (NewC C)                  \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (New T[e])                \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l (Ass va e)                \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2)  \<succ>\<midarrow>n\<rightarrow> xs'"
-	"G\<turnstile>Norm s \<midarrow>In2  ({accC,statDeclC,stat}e..fn) \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                 \<succ>\<midarrow>n\<rightarrow> vs'"
-	"G\<turnstile>Norm s \<midarrow>In1l ({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'"
+inductive_cases2 evaln_elim_cases:
+	"G\<turnstile>(Some xc, s) \<midarrow>t                        \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r Skip                      \<succ>\<midarrow>n\<rightarrow> (x, s')"
+        "G\<turnstile>Norm s \<midarrow>In1r (Jmp j)                   \<succ>\<midarrow>n\<rightarrow> (x, s')"
+        "G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> c)                    \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In3  ([])                      \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In3  (e#es)                    \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Lit w)                   \<succ>\<midarrow>n\<rightarrow> (v, s')"
+        "G\<turnstile>Norm s \<midarrow>In1l (UnOp unop e)             \<succ>\<midarrow>n\<rightarrow> (v, s')"
+        "G\<turnstile>Norm s \<midarrow>In1l (BinOp binop e1 e2)       \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In2  (LVar vn)                 \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Cast T e)                \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (e InstOf T)              \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Super)                   \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Acc va)                  \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (Expr e)                  \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (c1;; c2)                 \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Methd C sig)             \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Body D c)                \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2)            \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2)        \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> While(e) c)           \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2)           \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (Throw e)                 \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (NewC C)                  \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (New T[e])                \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l (Ass va e)                \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2)  \<succ>\<midarrow>n\<rightarrow> (x, s')"
+	"G\<turnstile>Norm s \<midarrow>In2  ({accC,statDeclC,stat}e..fn) \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                 \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1l ({accC,statT,mode}e\<cdot>mn({pT}p)) \<succ>\<midarrow>n\<rightarrow> (v, s')"
+	"G\<turnstile>Norm s \<midarrow>In1r (Init C)                  \<succ>\<midarrow>n\<rightarrow> (x, s')"
+        "G\<turnstile>Norm s \<midarrow>In1r (Init C)                  \<succ>\<midarrow>n\<rightarrow> (x, s')"
 
 declare split_if     [split] split_if_asm     [split] 
         option.split [split] option.split_asm [split]
@@ -288,25 +260,32 @@
  (injection @{term In1} into generalised values @{term vals}). 
 *}
 
+lemma evaln_expr_eq: "G\<turnstile>s \<midarrow>In1l t\<succ>\<midarrow>n\<rightarrow> (w, s') = (\<exists>v. w=In1 v \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<midarrow>n\<rightarrow> s')"
+  by (auto, frule evaln_Inj_elim, auto)
+
+lemma evaln_var_eq: "G\<turnstile>s \<midarrow>In2 t\<succ>\<midarrow>n\<rightarrow> (w, s') = (\<exists>vf. w=In2 vf \<and> G\<turnstile>s \<midarrow>t=\<succ>vf\<midarrow>n\<rightarrow> s')"
+  by (auto, frule evaln_Inj_elim, auto)
+
+lemma evaln_exprs_eq: "G\<turnstile>s \<midarrow>In3 t\<succ>\<midarrow>n\<rightarrow> (w, s') = (\<exists>vs. w=In3 vs \<and> G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s')"
+  by (auto, frule evaln_Inj_elim, auto)
+
+lemma evaln_stmt_eq: "G\<turnstile>s \<midarrow>In1r t\<succ>\<midarrow>n\<rightarrow> (w, s') = (w=\<diamondsuit> \<and> G\<turnstile>s \<midarrow>t \<midarrow>n\<rightarrow> s')"
+  by (auto, frule evaln_Inj_elim, auto, frule evaln_Inj_elim, auto)
+
 ML_setup {*
-fun enf nam inj rhs =
+fun enf name lhs =
 let
-  val name = "evaln_" ^ nam ^ "_eq"
-  val lhs = "G\<turnstile>s \<midarrow>" ^ inj ^ " t\<succ>\<midarrow>n\<rightarrow> (w, s')"
-  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
-	(K [Auto_tac, ALLGOALS (ftac (thm "evaln_Inj_elim")) THEN Auto_tac])
   fun is_Inj (Const (inj,_) $ _) = true
     | is_Inj _                   = false
-  fun pred (_ $ (Const ("Pair",_) $ _ $ (Const ("Pair", _) $ _ $ 
-    (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ x $ _ )))) $ _ ) = is_Inj x
+  fun pred (_ $ _ $ _ $ _ $ _ $ x $ _) = is_Inj x
 in
   cond_simproc name lhs pred (thm name)
 end;
 
-val evaln_expr_proc = enf "expr" "In1l" "\<exists>v.  w=In1 v  \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<midarrow>n\<rightarrow> s'";
-val evaln_var_proc  = enf "var"  "In2"  "\<exists>vf. w=In2 vf \<and> G\<turnstile>s \<midarrow>t=\<succ>vf\<midarrow>n\<rightarrow> s'";
-val evaln_exprs_proc= enf "exprs""In3"  "\<exists>vs. w=In3 vs \<and> G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s'";
-val evaln_stmt_proc = enf "stmt" "In1r" "     w=\<diamondsuit>      \<and> G\<turnstile>s \<midarrow>t     \<midarrow>n\<rightarrow> s'";
+val evaln_expr_proc  = enf "evaln_expr_eq"  "G\<turnstile>s \<midarrow>In1l t\<succ>\<midarrow>n\<rightarrow> (w, s')";
+val evaln_var_proc   = enf "evaln_var_eq"   "G\<turnstile>s \<midarrow>In2 t\<succ>\<midarrow>n\<rightarrow> (w, s')";
+val evaln_exprs_proc = enf "evaln_exprs_eq" "G\<turnstile>s \<midarrow>In3 t\<succ>\<midarrow>n\<rightarrow> (w, s')";
+val evaln_stmt_proc  = enf "evaln_stmt_eq"  "G\<turnstile>s \<midarrow>In1r t\<succ>\<midarrow>n\<rightarrow> (w, s')";
 Addsimprocs [evaln_expr_proc,evaln_var_proc,evaln_exprs_proc,evaln_stmt_proc];
 
 bind_thms ("evaln_AbruptIs", sum3_instantiate (thm "evaln.Abrupt"))
@@ -319,9 +298,7 @@
     assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s')" and
          normal: "normal s" and
          callee: "t=In1l (Callee l e)"
-    then have "False"
-    proof (induct)
-    qed (auto)
+    then have "False" by induct auto
   }
   then show ?thesis
     by (cases s') fastsimp 
@@ -333,9 +310,7 @@
     assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s')" and
          normal: "normal s" and
          callee: "t=In1l (InsInitE c e)"
-    then have "False"
-    proof (induct)
-    qed (auto)
+    then have "False" by induct auto
   }
   then show ?thesis
     by (cases s') fastsimp
@@ -347,9 +322,7 @@
     assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s')" and
          normal: "normal s" and
          callee: "t=In2 (InsInitV c w)"
-    then have "False"
-    proof (induct)
-    qed (auto)
+    then have "False" by induct auto
   }  
   then show ?thesis
     by (cases s') fastsimp
@@ -361,9 +334,7 @@
     assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s')" and
          normal: "normal s" and
          callee: "t=In1r (FinA a c)"
-    then have "False"
-    proof (induct)
-    qed (auto)
+    then have "False" by induct auto
   } 
   then show ?thesis
     by (cases s') fastsimp
@@ -385,9 +356,7 @@
 local
   fun is_Some (Const ("Pair",_) $ (Const ("Datatype.option.Some",_) $ _)$ _) =true
     | is_Some _ = false
-  fun pred (_ $ (Const ("Pair",_) $
-     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
-       (Const ("Pair", _) $ _ $ x)))) $ _ ) = is_Some x
+  fun pred (_ $ _ $ _ $ _ $ _ $ _ $ x) = is_Some x
 in
   val evaln_abrupt_proc = 
  cond_simproc "evaln_abrupt" "G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (w,s')" pred (thm "evaln_abrupt")
@@ -440,7 +409,7 @@
   shows "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
 using evaln 
 proof (induct)
-  case (Loop b c e l n s0 s1 s2 s3)
+  case (Loop s0 e n b s1 c s2 l s3)
   have "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1".
   moreover
   have "if the_Bool b
@@ -450,7 +419,7 @@
     using Loop.hyps by simp
   ultimately show ?case by (rule eval.Loop)
 next
-  case (Try c1 c2 n s0 s1 s2 s3 C vn)
+  case (Try s0 c1 n s1 s2 C vn c2 s3)
   have "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1".
   moreover
   have "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2".
@@ -459,7 +428,7 @@
     using Try.hyps by simp
   ultimately show ?case by (rule eval.Try)
 next
-  case (Init C c n s0 s1 s2 s3)
+  case (Init C c s0 s3 n s1 s2)
   have "the (class G C) = c".
   moreover
   have "if inited C (globs s0) 
@@ -478,8 +447,7 @@
 done
 
 lemma evaln_nonstrict [rule_format (no_asm), elim]: 
-  "\<And>ws. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> ws \<Longrightarrow> \<forall>m. n\<le>m \<longrightarrow> G\<turnstile>s \<midarrow>t\<succ>\<midarrow>m\<rightarrow> ws"
-apply (simp (no_asm_simp) only: split_tupled_all)
+  "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (w, s') \<Longrightarrow> \<forall>m. n\<le>m \<longrightarrow> G\<turnstile>s \<midarrow>t\<succ>\<midarrow>m\<rightarrow> (w, s')"
 apply (erule evaln.induct)
 apply (tactic {* ALLGOALS (EVERY'[strip_tac, TRY o etac (thm "Suc_le_D_lemma"),
   REPEAT o smp_tac 1, 
@@ -490,30 +458,30 @@
 
 lemmas evaln_nonstrict_Suc = evaln_nonstrict [OF _ le_refl [THEN le_SucI]]
 
-lemma evaln_max2: "\<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>n1\<rightarrow> ws1; G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>n2\<rightarrow> ws2\<rbrakk> \<Longrightarrow> 
-             G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max n1 n2\<rightarrow> ws1 \<and> G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max n1 n2\<rightarrow> ws2"
+lemma evaln_max2: "\<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>n1\<rightarrow> (w1, s1'); G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>n2\<rightarrow> (w2, s2')\<rbrakk> \<Longrightarrow> 
+             G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max n1 n2\<rightarrow> (w1, s1') \<and> G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max n1 n2\<rightarrow> (w2, s2')"
 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"
+  "\<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>n1\<rightarrow> (w1, s1'); G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>n2\<rightarrow> (w2, s2'); 
+    \<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max n1 n2\<rightarrow> (w1, s1');G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max n1 n2\<rightarrow> (w2, s2') \<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>
- G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws1 \<and>
- G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws2 \<and> 
- G\<turnstile>s3 \<midarrow>t3\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> ws3"
+"\<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>n1\<rightarrow> (w1, s1'); G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>n2\<rightarrow> (w2, s2'); G\<turnstile>s3 \<midarrow>t3\<succ>\<midarrow>n3\<rightarrow> (w3, s3')\<rbrakk> \<Longrightarrow>
+ G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> (w1, s1') \<and>
+ G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> (w2, s2') \<and> 
+ G\<turnstile>s3 \<midarrow>t3\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> (w3, s3')"
 apply (drule (1) evaln_max2, erule thin_rl)
 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
+"\<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>n1\<rightarrow> (w1, s1'); G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>n2\<rightarrow> (w2, s2'); G\<turnstile>s3 \<midarrow>t3\<succ>\<midarrow>n3\<rightarrow> (w3, s3');
+   \<lbrakk>G\<turnstile>s1 \<midarrow>t1\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> (w1, s1');
+    G\<turnstile>s2 \<midarrow>t2\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> (w2, s2'); 
+    G\<turnstile>s3 \<midarrow>t3\<succ>\<midarrow>max (max n1 n2) n3\<rightarrow> (w3, s3')
    \<rbrakk> \<Longrightarrow> P
   \<rbrakk> \<Longrightarrow> P"
 by (drule (2) evaln_max3) simp
@@ -551,16 +519,16 @@
   shows  "\<exists>n. G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)"
 using eval 
 proof (induct)
-  case (Abrupt s t xc)
+  case (Abrupt xc s t)
   obtain n where
-    "G\<turnstile>(Some xc, s) \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (arbitrary3 t, Some xc, s)"
+    "G\<turnstile>(Some xc, s) \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (arbitrary3 t, (Some xc, s))"
     by (iprover intro: evaln.Abrupt)
   then show ?case ..
 next
   case Skip
   show ?case by (blast intro: evaln.Skip)
 next
-  case (Expr e s0 s1 v)
+  case (Expr s0 e v s1)
   then obtain n where
     "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
     by (iprover)
@@ -568,7 +536,7 @@
     by (rule evaln.Expr) 
   then show ?case ..
 next
-  case (Lab c l s0 s1)
+  case (Lab s0 c s1 l)
   then obtain n where
     "G\<turnstile>Norm s0 \<midarrow>c\<midarrow>n\<rightarrow> s1"
     by (iprover)
@@ -576,7 +544,7 @@
     by (rule evaln.Lab)
   then show ?case ..
 next
-  case (Comp c1 c2 s0 s1 s2)
+  case (Comp s0 c1 s1 c2 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"
@@ -585,7 +553,7 @@
     by (blast intro: evaln.Comp dest: evaln_max2 )
   then show ?case ..
 next
-  case (If b c1 c2 e s0 s1 s2)
+  case (If s0 e b s1 c1 c2 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"
@@ -594,7 +562,7 @@
     by (blast intro: evaln.If dest: evaln_max2)
   then show ?case ..
 next
-  case (Loop b c e l s0 s1 s2 s3)
+  case (Loop s0 e b s1 c s2 l s3)
   from Loop.hyps obtain n1 where
     "G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n1\<rightarrow> s1"
     by (iprover)
@@ -614,12 +582,12 @@
     done
   then show ?case ..
 next
-  case (Jmp j s)
+  case (Jmp s j)
   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)
+  case (Throw s0 e a s1)
   then obtain n where
     "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1"
     by (iprover)
@@ -627,7 +595,7 @@
     by (rule evaln.Throw)
   then show ?case ..
 next 
-  case (Try catchC c1 c2 s0 s1 s2 s3 vn)
+  case (Try s0 c1 s1 s2 catchC vn c2 s3)
   from Try.hyps obtain n1 where
     "G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n1\<rightarrow> s1"
     by (iprover)
@@ -642,7 +610,7 @@
     by (auto intro!: evaln.Try le_maxI1 le_maxI2)
   then show ?case ..
 next
-  case (Fin c1 c2 s0 s1 s2 s3 x1)
+  case (Fin s0 c1 x1 s1 c2 s2 s3)
   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"
@@ -657,7 +625,7 @@
     by (blast intro: evaln.Fin dest: evaln_max2)
   then show ?case ..
 next
-  case (Init C c s0 s1 s2 s3)
+  case (Init C c s0 s3 s1 s2)
   have     cls: "the (class G C) = c" .
   moreover from Init.hyps obtain n where
       "if inited C (globs s0) then s3 = Norm s0
@@ -670,7 +638,7 @@
     by (rule evaln.Init)
   then show ?case ..
 next
-  case (NewC C a s0 s1 s2)
+  case (NewC s0 C s1 a s2)
   then obtain n where 
     "G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1"
     by (iprover)
@@ -679,7 +647,7 @@
     by (iprover intro: evaln.NewC)
   then show ?case ..
 next
-  case (NewA T a e i s0 s1 s2 s3)
+  case (NewA s0 T s1 e i s2 a 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"      
@@ -691,7 +659,7 @@
     by (blast intro: evaln.NewA dest: evaln_max2)
   then show ?case ..
 next
-  case (Cast castT e s0 s1 s2 v)
+  case (Cast s0 e v s1 s2 castT)
   then obtain n where
     "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
     by (iprover)
@@ -702,7 +670,7 @@
     by (rule evaln.Cast)
   then show ?case ..
 next
-  case (Inst T b e s0 s1 v)
+  case (Inst s0 e v s1 b T)
   then obtain n where
     "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
     by (iprover)
@@ -718,7 +686,7 @@
     by (rule evaln.Lit)
   then show ?case ..
 next
-  case (UnOp e s0 s1 unop v )
+  case (UnOp s0 e v s1 unop)
   then obtain n where
     "G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
     by (iprover)
@@ -726,7 +694,7 @@
     by (rule evaln.UnOp)
   then show ?case ..
 next
-  case (BinOp binop e1 e2 s0 s1 s2 v1 v2)
+  case (BinOp s0 e1 v1 s1 binop e2 v2 s2)
   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
@@ -742,7 +710,7 @@
     by (rule evaln.Super)
   then show ?case ..
 next
-  case (Acc f s0 s1 v va)
+  case (Acc s0 va v f s1)
   then obtain n where
     "G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v, f)\<midarrow>n\<rightarrow> s1"
     by (iprover)
@@ -751,7 +719,7 @@
     by (rule evaln.Acc)
   then show ?case ..
 next
-  case (Ass e f s0 s1 s2 v var w)
+  case (Ass s0 var w f s1 e v s2)
   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"      
@@ -761,7 +729,7 @@
     by (blast intro: evaln.Ass dest: evaln_max2)
   then show ?case ..
 next
-  case (Cond b e0 e1 e2 s0 s1 s2 v)
+  case (Cond s0 e0 b s1 e1 e2 v s2)
   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"
@@ -771,8 +739,7 @@
     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)
+  case (Call s0 e a' s1 args vs s2 invDeclC mode statT mn pTs' s3 s3' accC' v s4)
   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"
@@ -795,7 +762,7 @@
     by (auto intro!: evaln.Call le_maxI1 le_max3I1 le_max3I2)
   thus ?case ..
 next
-  case (Methd D s0 s1 sig v )
+  case (Methd s0 D sig v s1)
   then obtain n where
     "G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<midarrow>n\<rightarrow> s1"
     by iprover
@@ -803,7 +770,7 @@
     by (rule evaln.Methd)
   then show ?case ..
 next
-  case (Body D c s0 s1 s2 s3 )
+  case (Body s0 D s1 c 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"
@@ -826,7 +793,7 @@
     by (iprover intro: evaln.LVar)
   then show ?case ..
 next
-  case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v)
+  case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC)
   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"
@@ -839,7 +806,7 @@
     by (iprover intro: evaln.FVar dest: evaln_max2)
   then show ?case ..
 next
-  case (AVar a e1 e2 i s0 s1 s2 s2' v )
+  case (AVar s0 e1 a s1 e2 i s2 v s2')
   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"      
@@ -854,7 +821,7 @@
   case (Nil s0)
   show ?case by (iprover intro: evaln.Nil)
 next
-  case (Cons e es s0 s1 s2 v vs)
+  case (Cons s0 e v s1 es vs s2)
   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"      
--- a/src/HOL/Bali/Trans.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/Trans.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -60,29 +60,22 @@
 by (simp)
 declare the_var_AVar_def [simp del]
 
-consts
-  step	:: "prog \<Rightarrow> ((term \<times> state) \<times> (term \<times> state)) set"
-
-syntax (symbols)
-  step :: "[prog,term \<times> state,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>1 _"[61,82,82] 81)
-  stepn:: "[prog, term \<times> state,nat,term \<times> state] \<Rightarrow> bool" 
-                                                  ("_\<turnstile>_ \<mapsto>_ _"[61,82,82] 81)
-"step*":: "[prog,term \<times> state,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>* _"[61,82,82] 81)
+syntax (xsymbols)
   Ref  :: "loc \<Rightarrow> expr"
   SKIP :: "expr"
 
 translations
-  "G\<turnstile>p \<mapsto>1 p' " == "(p,p') \<in> step G"
-  "G\<turnstile>p \<mapsto>n p' " == "(p,p') \<in> (step G)^n"
-  "G\<turnstile>p \<mapsto>* p' " == "(p,p') \<in> (step G)\<^sup>*"
   "Ref a" == "Lit (Addr a)"
   "SKIP"  == "Lit Unit"
 
-inductive "step G" intros 
+inductive2
+  step :: "[prog,term \<times> state,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>1 _"[61,82,82] 81)
+  for G :: prog
+where
 
 (* evaluation of expression *)
   (* cf. 15.5 *)
-Abrupt:	         "\<lbrakk>\<forall>v. t \<noteq> \<langle>Lit v\<rangle>;
+  Abrupt:       "\<lbrakk>\<forall>v. t \<noteq> \<langle>Lit v\<rangle>;
                   \<forall> t. t \<noteq> \<langle>l\<bullet> Skip\<rangle>;
                   \<forall> C vn c.  t \<noteq> \<langle>Try Skip Catch(C vn) c\<rangle>;
                   \<forall> x c. t \<noteq> \<langle>Skip Finally c\<rangle> \<and> xc \<noteq> Xcpt x;
@@ -90,19 +83,19 @@
                 \<Longrightarrow> 
                   G\<turnstile>(t,Some xc,s) \<mapsto>1 (\<langle>Lit arbitrary\<rangle>,Some xc,s)"
 
-InsInitE: "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,Norm s) \<mapsto>1 (\<langle>c'\<rangle>, s')\<rbrakk>
-           \<Longrightarrow> 
-           G\<turnstile>(\<langle>InsInitE c e\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE c' e\<rangle>, s')"
+| InsInitE: "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,Norm s) \<mapsto>1 (\<langle>c'\<rangle>, s')\<rbrakk>
+             \<Longrightarrow> 
+             G\<turnstile>(\<langle>InsInitE c e\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE c' e\<rangle>, s')"
 
 (* SeqE: "G\<turnstile>(\<langle>Seq Skip e\<rangle>,Norm s) \<mapsto>1 (\<langle>e\<rangle>, Norm s)" *)
 (* Specialised rules to evaluate: 
    InsInitE Skip (NewC C), InisInitE Skip (NewA T[e]) *)
  
   (* cf. 15.8.1 *)
-NewC: "G\<turnstile>(\<langle>NewC C\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE (Init C) (NewC C)\<rangle>, Norm s)"
-NewCInited: "\<lbrakk>G\<turnstile> Norm s \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s'\<rbrakk> 
-             \<Longrightarrow> 
-             G\<turnstile>(\<langle>InsInitE Skip (NewC C)\<rangle>,Norm s) \<mapsto>1 (\<langle>Ref a\<rangle>, s')"
+| NewC: "G\<turnstile>(\<langle>NewC C\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE (Init C) (NewC C)\<rangle>, Norm s)"
+| NewCInited: "\<lbrakk>G\<turnstile> Norm s \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s'\<rbrakk> 
+               \<Longrightarrow> 
+               G\<turnstile>(\<langle>InsInitE Skip (NewC C)\<rangle>,Norm s) \<mapsto>1 (\<langle>Ref a\<rangle>, s')"
 
 
 
@@ -119,68 +112,68 @@
 
 *)
   (* cf. 15.9.1 *)
-NewA: 
+| NewA: 
    "G\<turnstile>(\<langle>New T[e]\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE (init_comp_ty T) (New T[e])\<rangle>,Norm s)"
-InsInitNewAIdx: 
+| InsInitNewAIdx: 
    "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>, s')\<rbrakk>
     \<Longrightarrow>  
     G\<turnstile>(\<langle>InsInitE Skip (New T[e])\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE Skip (New T[e'])\<rangle>,s')"
-InsInitNewA: 
+| InsInitNewA: 
    "\<lbrakk>G\<turnstile>abupd (check_neg i) (Norm s) \<midarrow>halloc (Arr T (the_Intg i))\<succ>a\<rightarrow> s' \<rbrakk>
     \<Longrightarrow>
     G\<turnstile>(\<langle>InsInitE Skip (New T[Lit i])\<rangle>,Norm s) \<mapsto>1 (\<langle>Ref a\<rangle>,s')"
  
   (* cf. 15.15 *)
-CastE:	
+| CastE:	
    "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
     \<Longrightarrow>
     G\<turnstile>(\<langle>Cast T e\<rangle>,None,s) \<mapsto>1 (\<langle>Cast T e'\<rangle>,s')" 
-Cast:	
+| Cast:	
    "\<lbrakk>s' = abupd (raise_if (\<not>G,s\<turnstile>v fits T)  ClassCast) (Norm s)\<rbrakk> 
     \<Longrightarrow> 
     G\<turnstile>(\<langle>Cast T (Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,s')"
   (* can be written without abupd, since we know Norm s *)
 
 
-InstE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk> 
+| InstE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk> 
         \<Longrightarrow> 
         G\<turnstile>(\<langle>e InstOf T\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')" 
-Inst:  "\<lbrakk>b = (v\<noteq>Null \<and> G,s\<turnstile>v fits RefT T)\<rbrakk> 
-        \<Longrightarrow> 
-        G\<turnstile>(\<langle>(Lit v) InstOf T\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (Bool b)\<rangle>,s')"
+| Inst:  "\<lbrakk>b = (v\<noteq>Null \<and> G,s\<turnstile>v fits RefT T)\<rbrakk> 
+          \<Longrightarrow> 
+          G\<turnstile>(\<langle>(Lit v) InstOf T\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (Bool b)\<rangle>,s')"
 
   (* cf. 15.7.1 *)
 (*Lit				"G\<turnstile>(Lit v,None,s) \<mapsto>1 (Lit v,None,s)"*)
 
-UnOpE:  "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s') \<rbrakk>
-         \<Longrightarrow> 
-         G\<turnstile>(\<langle>UnOp unop e\<rangle>,Norm s) \<mapsto>1 (\<langle>UnOp unop e'\<rangle>,s')"
-UnOp:   "G\<turnstile>(\<langle>UnOp unop (Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (eval_unop unop v)\<rangle>,Norm s)"
-
-BinOpE1:  "\<lbrakk>G\<turnstile>(\<langle>e1\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'\<rangle>,s') \<rbrakk>
+| UnOpE:  "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s') \<rbrakk>
            \<Longrightarrow> 
-           G\<turnstile>(\<langle>BinOp binop e1 e2\<rangle>,Norm s) \<mapsto>1 (\<langle>BinOp binop e1' e2\<rangle>,s')"
-BinOpE2:  "\<lbrakk>need_second_arg binop v1; G\<turnstile>(\<langle>e2\<rangle>,Norm s) \<mapsto>1 (\<langle>e2'\<rangle>,s') \<rbrakk>
-           \<Longrightarrow> 
-           G\<turnstile>(\<langle>BinOp binop (Lit v1) e2\<rangle>,Norm s) 
-            \<mapsto>1 (\<langle>BinOp binop (Lit v1) e2'\<rangle>,s')"
-BinOpTerm:  "\<lbrakk>\<not> need_second_arg binop v1\<rbrakk>
+           G\<turnstile>(\<langle>UnOp unop e\<rangle>,Norm s) \<mapsto>1 (\<langle>UnOp unop e'\<rangle>,s')"
+| UnOp:   "G\<turnstile>(\<langle>UnOp unop (Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (eval_unop unop v)\<rangle>,Norm s)"
+
+| BinOpE1:  "\<lbrakk>G\<turnstile>(\<langle>e1\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'\<rangle>,s') \<rbrakk>
+             \<Longrightarrow> 
+             G\<turnstile>(\<langle>BinOp binop e1 e2\<rangle>,Norm s) \<mapsto>1 (\<langle>BinOp binop e1' e2\<rangle>,s')"
+| BinOpE2:  "\<lbrakk>need_second_arg binop v1; G\<turnstile>(\<langle>e2\<rangle>,Norm s) \<mapsto>1 (\<langle>e2'\<rangle>,s') \<rbrakk>
              \<Longrightarrow> 
              G\<turnstile>(\<langle>BinOp binop (Lit v1) e2\<rangle>,Norm s) 
-              \<mapsto>1 (\<langle>Lit v1\<rangle>,Norm s)"
-BinOp:    "G\<turnstile>(\<langle>BinOp binop (Lit v1) (Lit v2)\<rangle>,Norm s) 
-            \<mapsto>1 (\<langle>Lit (eval_binop binop v1 v2)\<rangle>,Norm s)"
+              \<mapsto>1 (\<langle>BinOp binop (Lit v1) e2'\<rangle>,s')"
+| BinOpTerm:  "\<lbrakk>\<not> need_second_arg binop v1\<rbrakk>
+               \<Longrightarrow> 
+               G\<turnstile>(\<langle>BinOp binop (Lit v1) e2\<rangle>,Norm s) 
+                \<mapsto>1 (\<langle>Lit v1\<rangle>,Norm s)"
+| BinOp:    "G\<turnstile>(\<langle>BinOp binop (Lit v1) (Lit v2)\<rangle>,Norm s) 
+              \<mapsto>1 (\<langle>Lit (eval_binop binop v1 v2)\<rangle>,Norm s)"
 (* Maybe its more convenient to add: need_second_arg as precondition to BinOp 
    to make the choice between BinOpTerm and BinOp deterministic *)
    
-Super: "G\<turnstile>(\<langle>Super\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (val_this s)\<rangle>,Norm s)"
+| Super: "G\<turnstile>(\<langle>Super\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (val_this s)\<rangle>,Norm s)"
 
-AccVA: "\<lbrakk>G\<turnstile>(\<langle>va\<rangle>,Norm s) \<mapsto>1 (\<langle>va'\<rangle>,s') \<rbrakk>
-        \<Longrightarrow> 
-        G\<turnstile>(\<langle>Acc va\<rangle>,Norm s) \<mapsto>1 (\<langle>Acc va'\<rangle>,s')"
-Acc:  "\<lbrakk>groundVar va; ((v,vf),s') = the_var G (Norm s) va\<rbrakk>
-       \<Longrightarrow>  
-       G\<turnstile>(\<langle>Acc va\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,s')"
+| AccVA: "\<lbrakk>G\<turnstile>(\<langle>va\<rangle>,Norm s) \<mapsto>1 (\<langle>va'\<rangle>,s') \<rbrakk>
+          \<Longrightarrow> 
+          G\<turnstile>(\<langle>Acc va\<rangle>,Norm s) \<mapsto>1 (\<langle>Acc va'\<rangle>,s')"
+| Acc:  "\<lbrakk>groundVar va; ((v,vf),s') = the_var G (Norm s) va\<rbrakk>
+         \<Longrightarrow>  
+         G\<turnstile>(\<langle>Acc va\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,s')"
 
 (*
 AccLVar: "G\<turnstile>(\<langle>Acc (LVar vn)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (fst (lvar vn s))\<rangle>,Norm s)"
@@ -192,68 +185,68 @@
           \<Longrightarrow>  
           G\<turnstile>(\<langle>Acc ((Lit a).[Lit i])\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,s')"
 *) 
-AssVA:  "\<lbrakk>G\<turnstile>(\<langle>va\<rangle>,Norm s) \<mapsto>1 (\<langle>va'\<rangle>,s')\<rbrakk> 
-         \<Longrightarrow> 
-         G\<turnstile>(\<langle>va:=e\<rangle>,Norm s) \<mapsto>1 (\<langle>va':=e\<rangle>,s')"
-AssE:   "\<lbrakk>groundVar va; G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
-         \<Longrightarrow> 
-         G\<turnstile>(\<langle>va:=e\<rangle>,Norm s) \<mapsto>1 (\<langle>va:=e'\<rangle>,s')"
-Ass:    "\<lbrakk>groundVar va; ((w,f),s') = the_var G (Norm s) va\<rbrakk> 
-         \<Longrightarrow> 
-         G\<turnstile>(\<langle>va:=(Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,assign f v s')"
+| AssVA:  "\<lbrakk>G\<turnstile>(\<langle>va\<rangle>,Norm s) \<mapsto>1 (\<langle>va'\<rangle>,s')\<rbrakk> 
+           \<Longrightarrow> 
+           G\<turnstile>(\<langle>va:=e\<rangle>,Norm s) \<mapsto>1 (\<langle>va':=e\<rangle>,s')"
+| AssE:   "\<lbrakk>groundVar va; G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
+           \<Longrightarrow> 
+           G\<turnstile>(\<langle>va:=e\<rangle>,Norm s) \<mapsto>1 (\<langle>va:=e'\<rangle>,s')"
+| Ass:    "\<lbrakk>groundVar va; ((w,f),s') = the_var G (Norm s) va\<rbrakk> 
+           \<Longrightarrow> 
+           G\<turnstile>(\<langle>va:=(Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,assign f v s')"
 
-CondC: "\<lbrakk>G\<turnstile>(\<langle>e0\<rangle>,Norm s) \<mapsto>1 (\<langle>e0'\<rangle>,s')\<rbrakk> 
-        \<Longrightarrow> 
-        G\<turnstile>(\<langle>e0? e1:e2\<rangle>,Norm s) \<mapsto>1 (\<langle>e0'? e1:e2\<rangle>,s')"
-Cond:  "G\<turnstile>(\<langle>Lit b? e1:e2\<rangle>,Norm s) \<mapsto>1 (\<langle>if the_Bool b then e1 else e2\<rangle>,Norm s)"
+| CondC: "\<lbrakk>G\<turnstile>(\<langle>e0\<rangle>,Norm s) \<mapsto>1 (\<langle>e0'\<rangle>,s')\<rbrakk> 
+          \<Longrightarrow> 
+          G\<turnstile>(\<langle>e0? e1:e2\<rangle>,Norm s) \<mapsto>1 (\<langle>e0'? e1:e2\<rangle>,s')"
+| Cond:  "G\<turnstile>(\<langle>Lit b? e1:e2\<rangle>,Norm s) \<mapsto>1 (\<langle>if the_Bool b then e1 else e2\<rangle>,Norm s)"
 
 
-CallTarget: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
-             \<Longrightarrow>
-	     G\<turnstile>(\<langle>{accC,statT,mode}e\<cdot>mn({pTs}args)\<rangle>,Norm s) 
-              \<mapsto>1 (\<langle>{accC,statT,mode}e'\<cdot>mn({pTs}args)\<rangle>,s')"
-CallArgs:   "\<lbrakk>G\<turnstile>(\<langle>args\<rangle>,Norm s) \<mapsto>1 (\<langle>args'\<rangle>,s')\<rbrakk> 
-             \<Longrightarrow>
-	     G\<turnstile>(\<langle>{accC,statT,mode}Lit a\<cdot>mn({pTs}args)\<rangle>,Norm s) 
-              \<mapsto>1 (\<langle>{accC,statT,mode}Lit a\<cdot>mn({pTs}args')\<rangle>,s')"
-Call:       "\<lbrakk>groundExprs args; vs = map the_val args;
-              D = invocation_declclass G mode s a statT \<lparr>name=mn,parTs=pTs\<rparr>;
-              s'=init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs (Norm s)\<rbrakk> 
-             \<Longrightarrow> 
-             G\<turnstile>(\<langle>{accC,statT,mode}Lit a\<cdot>mn({pTs}args)\<rangle>,Norm s) 
-              \<mapsto>1 (\<langle>Callee (locals s) (Methd D \<lparr>name=mn,parTs=pTs\<rparr>)\<rangle>,s')"
+| CallTarget: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
+               \<Longrightarrow>
+               G\<turnstile>(\<langle>{accC,statT,mode}e\<cdot>mn({pTs}args)\<rangle>,Norm s) 
+                \<mapsto>1 (\<langle>{accC,statT,mode}e'\<cdot>mn({pTs}args)\<rangle>,s')"
+| CallArgs:   "\<lbrakk>G\<turnstile>(\<langle>args\<rangle>,Norm s) \<mapsto>1 (\<langle>args'\<rangle>,s')\<rbrakk> 
+               \<Longrightarrow>
+               G\<turnstile>(\<langle>{accC,statT,mode}Lit a\<cdot>mn({pTs}args)\<rangle>,Norm s) 
+                \<mapsto>1 (\<langle>{accC,statT,mode}Lit a\<cdot>mn({pTs}args')\<rangle>,s')"
+| Call:       "\<lbrakk>groundExprs args; vs = map the_val args;
+                D = invocation_declclass G mode s a statT \<lparr>name=mn,parTs=pTs\<rparr>;
+                s'=init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs (Norm s)\<rbrakk> 
+               \<Longrightarrow> 
+               G\<turnstile>(\<langle>{accC,statT,mode}Lit a\<cdot>mn({pTs}args)\<rangle>,Norm s) 
+                \<mapsto>1 (\<langle>Callee (locals s) (Methd D \<lparr>name=mn,parTs=pTs\<rparr>)\<rangle>,s')"
            
-Callee:     "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk>
-             \<Longrightarrow> 
-             G\<turnstile>(\<langle>Callee lcls_caller e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')"
+| Callee:     "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk>
+               \<Longrightarrow> 
+               G\<turnstile>(\<langle>Callee lcls_caller e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')"
 
-CalleeRet:   "G\<turnstile>(\<langle>Callee lcls_caller (Lit v)\<rangle>,Norm s) 
-               \<mapsto>1 (\<langle>Lit v\<rangle>,(set_lvars lcls_caller (Norm s)))"
+| CalleeRet:   "G\<turnstile>(\<langle>Callee lcls_caller (Lit v)\<rangle>,Norm s) 
+                 \<mapsto>1 (\<langle>Lit v\<rangle>,(set_lvars lcls_caller (Norm s)))"
 
-Methd: "G\<turnstile>(\<langle>Methd D sig\<rangle>,Norm s) \<mapsto>1 (\<langle>body G D sig\<rangle>,Norm s)"
+| Methd: "G\<turnstile>(\<langle>Methd D sig\<rangle>,Norm s) \<mapsto>1 (\<langle>body G D sig\<rangle>,Norm s)"
 
-Body: "G\<turnstile>(\<langle>Body D c\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE (Init D) (Body D c)\<rangle>,Norm s)"
+| Body: "G\<turnstile>(\<langle>Body D c\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE (Init D) (Body D c)\<rangle>,Norm s)"
 
-InsInitBody: 
+| InsInitBody: 
     "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,Norm s) \<mapsto>1 (\<langle>c'\<rangle>,s')\<rbrakk>
      \<Longrightarrow> 
      G\<turnstile>(\<langle>InsInitE Skip (Body D c)\<rangle>,Norm s) \<mapsto>1(\<langle>InsInitE Skip (Body D c')\<rangle>,s')"
-InsInitBodyRet: 
+| InsInitBodyRet: 
      "G\<turnstile>(\<langle>InsInitE Skip (Body D Skip)\<rangle>,Norm s)
        \<mapsto>1 (\<langle>Lit (the ((locals s) Result))\<rangle>,abupd (absorb Ret) (Norm s))"
 
 (*   LVar: "G\<turnstile>(LVar vn,Norm s)" is already evaluated *)
   
-FVar: "\<lbrakk>\<not> inited statDeclC (globs s)\<rbrakk>
-       \<Longrightarrow> 
-       G\<turnstile>(\<langle>{accC,statDeclC,stat}e..fn\<rangle>,Norm s) 
-        \<mapsto>1 (\<langle>InsInitV (Init statDeclC) ({accC,statDeclC,stat}e..fn)\<rangle>,Norm s)"
-InsInitFVarE:
+| FVar: "\<lbrakk>\<not> inited statDeclC (globs s)\<rbrakk>
+         \<Longrightarrow> 
+         G\<turnstile>(\<langle>{accC,statDeclC,stat}e..fn\<rangle>,Norm s) 
+          \<mapsto>1 (\<langle>InsInitV (Init statDeclC) ({accC,statDeclC,stat}e..fn)\<rangle>,Norm s)"
+| InsInitFVarE:
       "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk>
        \<Longrightarrow>
        G\<turnstile>(\<langle>InsInitV Skip ({accC,statDeclC,stat}e..fn)\<rangle>,Norm s) 
         \<mapsto>1 (\<langle>InsInitV Skip ({accC,statDeclC,stat}e'..fn)\<rangle>,s')"
-InsInitFVar:
+| InsInitFVar:
       "G\<turnstile>(\<langle>InsInitV Skip ({accC,statDeclC,stat}Lit a..fn)\<rangle>,Norm s) 
         \<mapsto>1 (\<langle>{accC,statDeclC,stat}Lit a..fn\<rangle>,Norm s)"
 --  {* Notice, that we do not have literal values for @{text vars}. 
@@ -266,13 +259,13 @@
 *}
 
 
-AVarE1: "\<lbrakk>G\<turnstile>(\<langle>e1\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'\<rangle>,s')\<rbrakk> 
-         \<Longrightarrow> 
-         G\<turnstile>(\<langle>e1.[e2]\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'.[e2]\<rangle>,s')"
+| AVarE1: "\<lbrakk>G\<turnstile>(\<langle>e1\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'\<rangle>,s')\<rbrakk> 
+           \<Longrightarrow> 
+           G\<turnstile>(\<langle>e1.[e2]\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'.[e2]\<rangle>,s')"
 
-AVarE2: "G\<turnstile>(\<langle>e2\<rangle>,Norm s) \<mapsto>1 (\<langle>e2'\<rangle>,s') 
-         \<Longrightarrow> 
-         G\<turnstile>(\<langle>Lit a.[e2]\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit a.[e2']\<rangle>,s')"
+| AVarE2: "G\<turnstile>(\<langle>e2\<rangle>,Norm s) \<mapsto>1 (\<langle>e2'\<rangle>,s') 
+           \<Longrightarrow> 
+           G\<turnstile>(\<langle>Lit a.[e2]\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit a.[e2']\<rangle>,s')"
 
 (* AVar: \<langle>(Lit a).[Lit i]\<rangle> is fully evaluated *)
 
@@ -280,89 +273,97 @@
 
   -- {* @{text Nil}  is fully evaluated *}
 
-ConsHd: "\<lbrakk>G\<turnstile>(\<langle>e::expr\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk> 
-         \<Longrightarrow>
-         G\<turnstile>(\<langle>e#es\<rangle>,Norm s) \<mapsto>1 (\<langle>e'#es\<rangle>,s')"
+| ConsHd: "\<lbrakk>G\<turnstile>(\<langle>e::expr\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk> 
+           \<Longrightarrow>
+           G\<turnstile>(\<langle>e#es\<rangle>,Norm s) \<mapsto>1 (\<langle>e'#es\<rangle>,s')"
   
-ConsTl: "\<lbrakk>G\<turnstile>(\<langle>es\<rangle>,Norm s) \<mapsto>1 (\<langle>es'\<rangle>,s')\<rbrakk> 
-         \<Longrightarrow>
-         G\<turnstile>(\<langle>(Lit v)#es\<rangle>,Norm s) \<mapsto>1 (\<langle>(Lit v)#es'\<rangle>,s')"
+| ConsTl: "\<lbrakk>G\<turnstile>(\<langle>es\<rangle>,Norm s) \<mapsto>1 (\<langle>es'\<rangle>,s')\<rbrakk> 
+           \<Longrightarrow>
+           G\<turnstile>(\<langle>(Lit v)#es\<rangle>,Norm s) \<mapsto>1 (\<langle>(Lit v)#es'\<rangle>,s')"
 
 (* execution of statements *)
 
   (* cf. 14.5 *)
-Skip: "G\<turnstile>(\<langle>Skip\<rangle>,Norm s) \<mapsto>1 (\<langle>SKIP\<rangle>,Norm s)"
+| Skip: "G\<turnstile>(\<langle>Skip\<rangle>,Norm s) \<mapsto>1 (\<langle>SKIP\<rangle>,Norm s)"
 
-ExprE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
-        \<Longrightarrow> 
-        G\<turnstile>(\<langle>Expr e\<rangle>,Norm s) \<mapsto>1 (\<langle>Expr e'\<rangle>,s')"
-Expr:  "G\<turnstile>(\<langle>Expr (Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,Norm s)"
+| ExprE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
+          \<Longrightarrow> 
+          G\<turnstile>(\<langle>Expr e\<rangle>,Norm s) \<mapsto>1 (\<langle>Expr e'\<rangle>,s')"
+| Expr:  "G\<turnstile>(\<langle>Expr (Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,Norm s)"
 
 
-LabC: "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,Norm s) \<mapsto>1 (\<langle>c'\<rangle>,s')\<rbrakk> 
-       \<Longrightarrow>  
-       G\<turnstile>(\<langle>l\<bullet> c\<rangle>,Norm s) \<mapsto>1 (\<langle>l\<bullet> c'\<rangle>,s')"
-Lab:  "G\<turnstile>(\<langle>l\<bullet> Skip\<rangle>,s) \<mapsto>1 (\<langle>Skip\<rangle>, abupd (absorb l) s)"
+| LabC: "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,Norm s) \<mapsto>1 (\<langle>c'\<rangle>,s')\<rbrakk> 
+         \<Longrightarrow>  
+         G\<turnstile>(\<langle>l\<bullet> c\<rangle>,Norm s) \<mapsto>1 (\<langle>l\<bullet> c'\<rangle>,s')"
+| Lab:  "G\<turnstile>(\<langle>l\<bullet> Skip\<rangle>,s) \<mapsto>1 (\<langle>Skip\<rangle>, abupd (absorb l) s)"
 
   (* cf. 14.2 *)
-CompC1:	"\<lbrakk>G\<turnstile>(\<langle>c1\<rangle>,Norm s) \<mapsto>1 (\<langle>c1'\<rangle>,s')\<rbrakk> 
-         \<Longrightarrow> 
-         G\<turnstile>(\<langle>c1;; c2\<rangle>,Norm s) \<mapsto>1 (\<langle>c1';; c2\<rangle>,s')"
+| CompC1: "\<lbrakk>G\<turnstile>(\<langle>c1\<rangle>,Norm s) \<mapsto>1 (\<langle>c1'\<rangle>,s')\<rbrakk> 
+           \<Longrightarrow> 
+           G\<turnstile>(\<langle>c1;; c2\<rangle>,Norm s) \<mapsto>1 (\<langle>c1';; c2\<rangle>,s')"
 
-Comp:   "G\<turnstile>(\<langle>Skip;; c2\<rangle>,Norm s) \<mapsto>1 (\<langle>c2\<rangle>,Norm s)"
+| Comp:   "G\<turnstile>(\<langle>Skip;; c2\<rangle>,Norm s) \<mapsto>1 (\<langle>c2\<rangle>,Norm s)"
 
   (* cf. 14.8.2 *)
-IfE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle> ,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
-      \<Longrightarrow>
-      G\<turnstile>(\<langle>If(e) s1 Else s2\<rangle>,Norm s) \<mapsto>1 (\<langle>If(e') s1 Else s2\<rangle>,s')"
-If:  "G\<turnstile>(\<langle>If(Lit v) s1 Else s2\<rangle>,Norm s) 
-       \<mapsto>1 (\<langle>if the_Bool v then s1 else s2\<rangle>,Norm s)"
+| IfE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle> ,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
+        \<Longrightarrow>
+        G\<turnstile>(\<langle>If(e) s1 Else s2\<rangle>,Norm s) \<mapsto>1 (\<langle>If(e') s1 Else s2\<rangle>,s')"
+| If:  "G\<turnstile>(\<langle>If(Lit v) s1 Else s2\<rangle>,Norm s) 
+         \<mapsto>1 (\<langle>if the_Bool v then s1 else s2\<rangle>,Norm s)"
 
   (* cf. 14.10, 14.10.1 *)
-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)"
+| 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)"
 
-Jmp: "G\<turnstile>(\<langle>Jmp 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>
-         G\<turnstile>(\<langle>Throw e\<rangle>,Norm s) \<mapsto>1 (\<langle>Throw e'\<rangle>,s')"
-Throw:  "G\<turnstile>(\<langle>Throw (Lit a)\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,abupd (throw a) (Norm s))"
+| ThrowE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> 
+           \<Longrightarrow>
+           G\<turnstile>(\<langle>Throw e\<rangle>,Norm s) \<mapsto>1 (\<langle>Throw e'\<rangle>,s')"
+| Throw:  "G\<turnstile>(\<langle>Throw (Lit a)\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,abupd (throw a) (Norm s))"
 
-TryC1: "\<lbrakk>G\<turnstile>(\<langle>c1\<rangle>,Norm s) \<mapsto>1 (\<langle>c1'\<rangle>,s')\<rbrakk> 
-        \<Longrightarrow>
-        G\<turnstile>(\<langle>Try c1 Catch(C vn) c2\<rangle>, Norm s) \<mapsto>1 (\<langle>Try c1' Catch(C vn) c2\<rangle>,s')"
-Try:   "\<lbrakk>G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'\<rbrakk>
-        \<Longrightarrow>
-        G\<turnstile>(\<langle>Try Skip Catch(C vn) c2\<rangle>, s) 
-         \<mapsto>1 (if G,s'\<turnstile>catch C then (\<langle>c2\<rangle>,new_xcpt_var vn s')
-                              else (\<langle>Skip\<rangle>,s'))"
+| TryC1: "\<lbrakk>G\<turnstile>(\<langle>c1\<rangle>,Norm s) \<mapsto>1 (\<langle>c1'\<rangle>,s')\<rbrakk> 
+          \<Longrightarrow>
+          G\<turnstile>(\<langle>Try c1 Catch(C vn) c2\<rangle>, Norm s) \<mapsto>1 (\<langle>Try c1' Catch(C vn) c2\<rangle>,s')"
+| Try:   "\<lbrakk>G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'\<rbrakk>
+          \<Longrightarrow>
+          G\<turnstile>(\<langle>Try Skip Catch(C vn) c2\<rangle>, s) 
+           \<mapsto>1 (if G,s'\<turnstile>catch C then (\<langle>c2\<rangle>,new_xcpt_var vn s')
+                                else (\<langle>Skip\<rangle>,s'))"
 
-FinC1: "\<lbrakk>G\<turnstile>(\<langle>c1\<rangle>,Norm s) \<mapsto>1 (\<langle>c1'\<rangle>,s')\<rbrakk> 
-        \<Longrightarrow>
-        G\<turnstile>(\<langle>c1 Finally c2\<rangle>,Norm s) \<mapsto>1 (\<langle>c1' Finally c2\<rangle>,s')"
+| FinC1: "\<lbrakk>G\<turnstile>(\<langle>c1\<rangle>,Norm s) \<mapsto>1 (\<langle>c1'\<rangle>,s')\<rbrakk> 
+          \<Longrightarrow>
+          G\<turnstile>(\<langle>c1 Finally c2\<rangle>,Norm s) \<mapsto>1 (\<langle>c1' Finally c2\<rangle>,s')"
 
-Fin:    "G\<turnstile>(\<langle>Skip Finally c2\<rangle>,(a,s)) \<mapsto>1 (\<langle>FinA a c2\<rangle>,Norm s)"
+| Fin:    "G\<turnstile>(\<langle>Skip Finally c2\<rangle>,(a,s)) \<mapsto>1 (\<langle>FinA a c2\<rangle>,Norm s)"
 
-FinAC: "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,s) \<mapsto>1 (\<langle>c'\<rangle>,s')\<rbrakk>
-        \<Longrightarrow>
-        G\<turnstile>(\<langle>FinA a c\<rangle>,s) \<mapsto>1 (\<langle>FinA a c'\<rangle>,s')"
-FinA: "G\<turnstile>(\<langle>FinA a Skip\<rangle>,s) \<mapsto>1 (\<langle>Skip\<rangle>,abupd (abrupt_if (a\<noteq>None) a) s)"
+| FinAC: "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,s) \<mapsto>1 (\<langle>c'\<rangle>,s')\<rbrakk>
+          \<Longrightarrow>
+          G\<turnstile>(\<langle>FinA a c\<rangle>,s) \<mapsto>1 (\<langle>FinA a c'\<rangle>,s')"
+| FinA: "G\<turnstile>(\<langle>FinA a Skip\<rangle>,s) \<mapsto>1 (\<langle>Skip\<rangle>,abupd (abrupt_if (a\<noteq>None) a) s)"
 
 
-Init1: "\<lbrakk>inited C (globs s)\<rbrakk> 
-        \<Longrightarrow> 
-        G\<turnstile>(\<langle>Init C\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,Norm s)"
-Init: "\<lbrakk>the (class G C)=c; \<not> inited C (globs s)\<rbrakk>  
-       \<Longrightarrow> 
-       G\<turnstile>(\<langle>Init C\<rangle>,Norm s) 
-        \<mapsto>1 (\<langle>(if C = Object then Skip else (Init (super c)));;
-              Expr (Callee (locals s) (InsInitE (init c) SKIP))\<rangle>
-             ,Norm (init_class_obj G C s))"
+| Init1: "\<lbrakk>inited C (globs s)\<rbrakk> 
+          \<Longrightarrow> 
+          G\<turnstile>(\<langle>Init C\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,Norm s)"
+| Init: "\<lbrakk>the (class G C)=c; \<not> inited C (globs s)\<rbrakk>  
+         \<Longrightarrow> 
+         G\<turnstile>(\<langle>Init C\<rangle>,Norm s) 
+          \<mapsto>1 (\<langle>(if C = Object then Skip else (Init (super c)));;
+                Expr (Callee (locals s) (InsInitE (init c) SKIP))\<rangle>
+               ,Norm (init_class_obj G C s))"
 -- {* @{text InsInitE} is just used as trick to embed the statement 
 @{text "init c"} into an expression*} 
-InsInitESKIP:
-  "G\<turnstile>(\<langle>InsInitE Skip SKIP\<rangle>,Norm s) \<mapsto>1 (\<langle>SKIP\<rangle>,Norm s)"
+| InsInitESKIP:
+    "G\<turnstile>(\<langle>InsInitE Skip SKIP\<rangle>,Norm s) \<mapsto>1 (\<langle>SKIP\<rangle>,Norm s)"
+
+abbreviation
+  stepn:: "[prog, term \<times> state,nat,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>_ _"[61,82,82] 81)
+  where "G\<turnstile>p \<mapsto>n p' \<equiv> (p,p') \<in> {(x, y). step G x y}^n"
+
+abbreviation
+  steptr:: "[prog,term \<times> state,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>* _"[61,82,82] 81)
+  where "G\<turnstile>p \<mapsto>* p' \<equiv> (p,p') \<in> {(x, y). step G x y}\<^sup>*"
          
 (* Equivalenzen:
   Bigstep zu Smallstep komplett.
--- a/src/HOL/Bali/TypeRel.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/TypeRel.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -34,10 +34,6 @@
 (*subcls , by translation*)                  (*        subclass           *)
 (*subclseq, by translation*)                 (* subclass + identity       *)
   implmt1   :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{* direct implementation *}
-  implmt    :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{*        implementation *}
-  widen     :: "prog \<Rightarrow> (ty    \<times> ty   ) set" --{*        widening       *}
-  narrow    :: "prog \<Rightarrow> (ty    \<times> ty   ) set" --{*        narrowing      *}
-  cast     :: "prog \<Rightarrow> (ty    \<times> ty   ) set"  --{*        casting        *}
 
 syntax
 
@@ -49,12 +45,8 @@
  "@subcls"  :: "prog => [qtname, qtname] => bool" ("_|-_<:C _"[71,71,71] 70)
  *)
  "@implmt1" :: "prog => [qtname, qtname] => bool" ("_|-_~>1_"  [71,71,71] 70)
- "@implmt"  :: "prog => [qtname, qtname] => bool" ("_|-_~>_"   [71,71,71] 70)
- "@widen"   :: "prog => [ty   , ty   ] => bool" ("_|-_<=:_"  [71,71,71] 70)
- "@narrow"  :: "prog => [ty   , ty   ] => bool" ("_|-_:>_"   [71,71,71] 70)
- "@cast"    :: "prog => [ty   , ty   ] => bool" ("_|-_<=:? _"[71,71,71] 70)
 
-syntax (symbols)
+syntax (xsymbols)
 
   "@subint1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>I1_"  [71,71,71] 70)
   "@subint"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>I _"  [71,71,71] 70)
@@ -64,10 +56,6 @@
   "@subcls"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<prec>\<^sub>C _"  [71,71,71] 70)
   *)
   "@implmt1" :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<leadsto>1_"  [71,71,71] 70)
-  "@implmt"  :: "prog \<Rightarrow> [qtname, qtname] \<Rightarrow> bool" ("_\<turnstile>_\<leadsto>_"   [71,71,71] 70)
-  "@widen"   :: "prog \<Rightarrow> [ty   , ty   ] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>_"    [71,71,71] 70)
-  "@narrow"  :: "prog \<Rightarrow> [ty   , ty   ] \<Rightarrow> bool" ("_\<turnstile>_\<succ>_"    [71,71,71] 70)
-  "@cast"    :: "prog \<Rightarrow> [ty   , ty   ] \<Rightarrow> bool" ("_\<turnstile>_\<preceq>? _"  [71,71,71] 70)
 
 translations
 
@@ -78,10 +66,6 @@
 	"G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" 
 	*)
         "G\<turnstile>C \<leadsto>1 I" == "(C,I) \<in> implmt1 G"
-	"G\<turnstile>C \<leadsto>  I" == "(C,I) \<in> implmt  G"
-	"G\<turnstile>S \<preceq>   T" == "(S,T) \<in> widen   G"
-	"G\<turnstile>S \<succ>   T" == "(S,T) \<in> narrow  G"
-	"G\<turnstile>S \<preceq>?  T" == "(S,T) \<in> cast    G"
 
 
 section "subclass and subinterface relations"
@@ -367,11 +351,13 @@
 done
 
 
-inductive "implmt G" intros                    --{* cf. 8.1.4 *}
-
+inductive2 --{* implementation, cf. 8.1.4 *}
+  implmt :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" ("_\<turnstile>_\<leadsto>_" [71,71,71] 70)
+  for G :: prog
+where
   direct:         "G\<turnstile>C\<leadsto>1J     \<spacespace>\<spacespace>     \<Longrightarrow> G\<turnstile>C\<leadsto>J"
-  subint:        "\<lbrakk>G\<turnstile>C\<leadsto>1I; G\<turnstile>I\<preceq>I J\<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
-  subcls1:       "\<lbrakk>G\<turnstile>C\<prec>\<^sub>C\<^sub>1D; G\<turnstile>D\<leadsto>J \<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
+| subint:        "\<lbrakk>G\<turnstile>C\<leadsto>1I; G\<turnstile>I\<preceq>I J\<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
+| subcls1:       "\<lbrakk>G\<turnstile>C\<prec>\<^sub>C\<^sub>1D; G\<turnstile>D\<leadsto>J \<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
 
 lemma implmtD: "G\<turnstile>C\<leadsto>J \<Longrightarrow> (\<exists>I. G\<turnstile>C\<leadsto>1I \<and> G\<turnstile>I\<preceq>I J) \<or> (\<exists>D. G\<turnstile>C\<prec>\<^sub>C\<^sub>1D \<and> G\<turnstile>D\<leadsto>J)" 
 apply (erule implmt.induct)
@@ -399,17 +385,20 @@
 
 section "widening relation"
 
-inductive "widen G" intros
+inductive2
  --{*widening, viz. method invocation conversion, cf. 5.3
 			    i.e. kind of syntactic subtyping *}
+  widen :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<preceq>_" [71,71,71] 70)
+  for G :: prog
+where
   refl:    "G\<turnstile>T\<preceq>T" --{*identity conversion, cf. 5.1.1 *}
-  subint:  "G\<turnstile>I\<preceq>I J  \<Longrightarrow> G\<turnstile>Iface I\<preceq> Iface J" --{*wid.ref.conv.,cf. 5.1.4 *}
-  int_obj: "G\<turnstile>Iface I\<preceq> Class Object"
-  subcls:  "G\<turnstile>C\<preceq>\<^sub>C D  \<Longrightarrow> G\<turnstile>Class C\<preceq> Class D"
-  implmt:  "G\<turnstile>C\<leadsto>I   \<Longrightarrow> G\<turnstile>Class C\<preceq> Iface I"
-  null:    "G\<turnstile>NT\<preceq> RefT R"
-  arr_obj: "G\<turnstile>T.[]\<preceq> Class Object"
-  array:   "G\<turnstile>RefT S\<preceq>RefT T \<Longrightarrow> G\<turnstile>RefT S.[]\<preceq> RefT T.[]"
+| subint:  "G\<turnstile>I\<preceq>I J  \<Longrightarrow> G\<turnstile>Iface I\<preceq> Iface J" --{*wid.ref.conv.,cf. 5.1.4 *}
+| int_obj: "G\<turnstile>Iface I\<preceq> Class Object"
+| subcls:  "G\<turnstile>C\<preceq>\<^sub>C D  \<Longrightarrow> G\<turnstile>Class C\<preceq> Class D"
+| implmt:  "G\<turnstile>C\<leadsto>I   \<Longrightarrow> G\<turnstile>Class C\<preceq> Iface I"
+| null:    "G\<turnstile>NT\<preceq> RefT R"
+| arr_obj: "G\<turnstile>T.[]\<preceq> Class Object"
+| array:   "G\<turnstile>RefT S\<preceq>RefT T \<Longrightarrow> G\<turnstile>RefT S.[]\<preceq> RefT T.[]"
 
 declare widen.refl [intro!]
 declare widen.intros [simp]
@@ -418,14 +407,14 @@
 lemma widen_PrimT: "G\<turnstile>PrimT x\<preceq>T \<Longrightarrow> T = PrimT x"
 *)
 lemma widen_PrimT: "G\<turnstile>PrimT x\<preceq>T \<Longrightarrow> (\<exists>y. T = PrimT y)"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>PrimT x\<preceq>T")
 by auto
 
 (* too strong in general:
 lemma widen_PrimT2: "G\<turnstile>S\<preceq>PrimT x \<Longrightarrow> S = PrimT x"
 *)
 lemma widen_PrimT2: "G\<turnstile>S\<preceq>PrimT x \<Longrightarrow> \<exists>y. S = PrimT y"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>S\<preceq>PrimT x")
 by auto
 
 text {* 
@@ -435,37 +424,37 @@
  *}
 
 lemma widen_PrimT_strong: "G\<turnstile>PrimT x\<preceq>T \<Longrightarrow> T = PrimT x"
-by (ind_cases "G\<turnstile>S\<preceq>T") simp_all
+by (ind_cases2 "G\<turnstile>PrimT x\<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
+by (ind_cases2 "G\<turnstile>S\<preceq>PrimT x") 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
+by (ind_cases2 "G\<turnstile>PrimT Boolean\<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
+by (ind_cases2 "G\<turnstile>S\<preceq>PrimT Boolean") 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")
+apply (ind_cases2 "G\<turnstile>RefT R\<preceq>T")
 by auto
 
 lemma widen_RefT2: "G\<turnstile>S\<preceq>RefT R \<Longrightarrow> \<exists>t. S=RefT t"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>S\<preceq>RefT R")
 by auto
 
 lemma widen_Iface: "G\<turnstile>Iface I\<preceq>T \<Longrightarrow> T=Class Object \<or> (\<exists>J. T=Iface J)"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>Iface I\<preceq>T")
 by auto
 
 lemma widen_Iface2: "G\<turnstile>S\<preceq> Iface J \<Longrightarrow> S = NT \<or> (\<exists>I. S = Iface I) \<or> (\<exists>D. S = Class D)"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>S\<preceq> Iface J")
 by auto
 
 lemma widen_Iface_Iface: "G\<turnstile>Iface I\<preceq> Iface J \<Longrightarrow> G\<turnstile>I\<preceq>I J"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>Iface I\<preceq> Iface J")
 by auto
 
 lemma widen_Iface_Iface_eq [simp]: "G\<turnstile>Iface I\<preceq> Iface J = G\<turnstile>I\<preceq>I J"
@@ -475,15 +464,15 @@
 done
 
 lemma widen_Class: "G\<turnstile>Class C\<preceq>T \<Longrightarrow>  (\<exists>D. T=Class D) \<or> (\<exists>I. T=Iface I)"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>Class C\<preceq>T")
 by auto
 
 lemma widen_Class2: "G\<turnstile>S\<preceq> Class C \<Longrightarrow> C = Object \<or> S = NT \<or> (\<exists>D. S = Class D)"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>S\<preceq> Class C")
 by auto
 
 lemma widen_Class_Class: "G\<turnstile>Class C\<preceq> Class cm \<Longrightarrow> G\<turnstile>C\<preceq>\<^sub>C cm"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>Class C\<preceq> Class cm")
 by auto
 
 lemma widen_Class_Class_eq [simp]: "G\<turnstile>Class C\<preceq> Class cm = G\<turnstile>C\<preceq>\<^sub>C cm"
@@ -493,7 +482,7 @@
 done
 
 lemma widen_Class_Iface: "G\<turnstile>Class C\<preceq> Iface I \<Longrightarrow> G\<turnstile>C\<leadsto>I"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>Class C\<preceq> Iface I")
 by auto
 
 lemma widen_Class_Iface_eq [simp]: "G\<turnstile>Class C\<preceq> Iface I = G\<turnstile>C\<leadsto>I"
@@ -503,21 +492,21 @@
 done
 
 lemma widen_Array: "G\<turnstile>S.[]\<preceq>T \<Longrightarrow> T=Class Object \<or> (\<exists>T'. T=T'.[] \<and> G\<turnstile>S\<preceq>T')"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>S.[]\<preceq>T")
 by auto
 
 lemma widen_Array2: "G\<turnstile>S\<preceq>T.[] \<Longrightarrow> S = NT \<or> (\<exists>S'. S=S'.[] \<and> G\<turnstile>S'\<preceq>T)"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>S\<preceq>T.[]")
 by auto
 
 
 lemma widen_ArrayPrimT: "G\<turnstile>PrimT t.[]\<preceq>T \<Longrightarrow> T=Class Object \<or> T=PrimT t.[]"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>PrimT t.[]\<preceq>T")
 by auto
 
 lemma widen_ArrayRefT: 
   "G\<turnstile>RefT t.[]\<preceq>T \<Longrightarrow> T=Class Object \<or> (\<exists>s. T=RefT s.[] \<and> G\<turnstile>RefT t\<preceq>RefT s)"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>RefT t.[]\<preceq>T")
 by auto
 
 lemma widen_ArrayRefT_ArrayRefT_eq [simp]: 
@@ -538,7 +527,7 @@
 
 
 lemma widen_NT2: "G\<turnstile>S\<preceq>NT \<Longrightarrow> S = NT"
-apply (ind_cases "G\<turnstile>S\<preceq>T")
+apply (ind_cases2 "G\<turnstile>S\<preceq>NT")
 by auto
 
 lemma widen_Object:"\<lbrakk>isrtype G T;ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>RefT T \<preceq> Class Object"
@@ -621,35 +610,35 @@
 *)
 
 (* more detailed than necessary for type-safety, see above rules. *)
-inductive "narrow G" intros (* narrowing reference conversion, cf. 5.1.5 *)
-
+inductive2 --{* narrowing reference conversion, cf. 5.1.5 *}
+  narrow :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<succ>_" [71,71,71] 70)
+  for G :: prog
+where
   subcls:  "G\<turnstile>C\<preceq>\<^sub>C D \<Longrightarrow> G\<turnstile>     Class D\<succ>Class C"
-  implmt:  "\<not>G\<turnstile>C\<leadsto>I \<Longrightarrow> G\<turnstile>     Class C\<succ>Iface I"
-  obj_arr: "G\<turnstile>Class Object\<succ>T.[]"
-  int_cls: "G\<turnstile>     Iface I\<succ>Class C"
-  subint:  "imethds G I hidings imethds G J entails 
+| implmt:  "\<not>G\<turnstile>C\<leadsto>I \<Longrightarrow> G\<turnstile>     Class C\<succ>Iface I"
+| obj_arr: "G\<turnstile>Class Object\<succ>T.[]"
+| int_cls: "G\<turnstile>     Iface I\<succ>Class C"
+| subint:  "imethds G I hidings imethds G J entails 
 	   (\<lambda>(md, mh   ) (md',mh').\<spacespace>G\<turnstile>mrt mh\<preceq>mrt mh') \<Longrightarrow>
 	   \<not>G\<turnstile>I\<preceq>I J         \<spacespace>\<spacespace>\<spacespace>\<Longrightarrow> G\<turnstile>     Iface I\<succ>Iface J"
-  array:   "G\<turnstile>RefT S\<succ>RefT T   \<spacespace>\<Longrightarrow> G\<turnstile>   RefT S.[]\<succ>RefT T.[]"
+| array:   "G\<turnstile>RefT S\<succ>RefT T   \<spacespace>\<Longrightarrow> G\<turnstile>   RefT S.[]\<succ>RefT T.[]"
 
 (*unused*)
 lemma narrow_RefT: "G\<turnstile>RefT R\<succ>T \<Longrightarrow> \<exists>t. T=RefT t"
-apply (ind_cases "G\<turnstile>S\<succ>T")
+apply (ind_cases2 "G\<turnstile>RefT R\<succ>T")
 by auto
 
 lemma narrow_RefT2: "G\<turnstile>S\<succ>RefT R \<Longrightarrow> \<exists>t. S=RefT t"
-apply (ind_cases "G\<turnstile>S\<succ>T")
+apply (ind_cases2 "G\<turnstile>S\<succ>RefT R")
 by auto
 
 (*unused*)
 lemma narrow_PrimT: "G\<turnstile>PrimT pt\<succ>T \<Longrightarrow> \<exists>t. T=PrimT t"
-apply (ind_cases "G\<turnstile>S\<succ>T")
-by auto
+by (ind_cases2 "G\<turnstile>PrimT pt\<succ>T")
 
 lemma narrow_PrimT2: "G\<turnstile>S\<succ>PrimT pt \<Longrightarrow>  
 				  \<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
-apply (ind_cases "G\<turnstile>S\<succ>T")
-by auto
+by (ind_cases2 "G\<turnstile>S\<succ>PrimT pt")
 
 text {* 
    These narrowing lemmata hold in Bali but are to strong for ordinary
@@ -657,42 +646,44 @@
    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
+by (ind_cases2 "G\<turnstile>PrimT pt\<succ>T")
 
 lemma narrow_PrimT2_strong: "G\<turnstile>S\<succ>PrimT pt \<Longrightarrow> S=PrimT pt"
-by (ind_cases "G\<turnstile>S\<succ>T") simp_all
+by (ind_cases2 "G\<turnstile>S\<succ>PrimT pt")
 
 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
+by (ind_cases2 "G\<turnstile>PrimT Boolean\<succ>T")
 
 lemma narrow_Boolean2: "G\<turnstile>S\<succ>PrimT Boolean \<Longrightarrow> S=PrimT Boolean"
-by (ind_cases "G\<turnstile>S\<succ>T") simp_all
+by (ind_cases2 "G\<turnstile>S\<succ>PrimT Boolean")
 
 section "casting relation"
 
-inductive "cast G" intros --{* casting conversion, cf. 5.5 *}
-
+inductive2 --{* casting conversion, cf. 5.5 *}
+  cast :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<preceq>? _" [71,71,71] 70)
+  for G :: prog
+where
   widen:   "G\<turnstile>S\<preceq>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
-  narrow:  "G\<turnstile>S\<succ>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
+| narrow:  "G\<turnstile>S\<succ>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
 
 (*unused*)
 lemma cast_RefT: "G\<turnstile>RefT R\<preceq>? T \<Longrightarrow> \<exists>t. T=RefT t"
-apply (ind_cases "G\<turnstile>S\<preceq>? T")
+apply (ind_cases2 "G\<turnstile>RefT R\<preceq>? T")
 by (auto dest: widen_RefT narrow_RefT)
 
 lemma cast_RefT2: "G\<turnstile>S\<preceq>? RefT R \<Longrightarrow> \<exists>t. S=RefT t"
-apply (ind_cases "G\<turnstile>S\<preceq>? T")
+apply (ind_cases2 "G\<turnstile>S\<preceq>? RefT R")
 by (auto dest: widen_RefT2 narrow_RefT2)
 
 (*unused*)
 lemma cast_PrimT: "G\<turnstile>PrimT pt\<preceq>? T \<Longrightarrow> \<exists>t. T=PrimT t"
-apply (ind_cases "G\<turnstile>S\<preceq>? T")
+apply (ind_cases2 "G\<turnstile>PrimT pt\<preceq>? T")
 by (auto dest: widen_PrimT narrow_PrimT)
 
 lemma cast_PrimT2: "G\<turnstile>S\<preceq>? PrimT pt \<Longrightarrow> \<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
-apply (ind_cases "G\<turnstile>S\<preceq>? T")
+apply (ind_cases2 "G\<turnstile>S\<preceq>? PrimT pt")
 by (auto dest: widen_PrimT2 narrow_PrimT2)
 
 lemma cast_Boolean: 
--- a/src/HOL/Bali/TypeSafe.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/TypeSafe.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -226,7 +226,7 @@
 
 lemmas eval_gext = eval_gext_lemma [THEN conjunct1]
 
-lemma eval_gext': "G\<turnstile>(x1,s1) \<midarrow>t\<succ>\<rightarrow> (w,x2,s2) \<Longrightarrow> s1\<le>|s2"
+lemma eval_gext': "G\<turnstile>(x1,s1) \<midarrow>t\<succ>\<rightarrow> (w,(x2,s2)) \<Longrightarrow> s1\<le>|s2"
 apply (drule eval_gext)
 apply auto
 done
@@ -1735,7 +1735,7 @@
                  \<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 A) 
+    case (Abrupt xc s t 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) 
@@ -1751,7 +1751,7 @@
               (error_free (Norm s) = error_free (Norm s))"
       by (simp)
   next
-    case (Expr e s0 s1 v L accC T A)
+    case (Expr s0 e v s1 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)" .
@@ -1773,7 +1773,7 @@
           (error_free (Norm s0) = error_free s1)"
       by (simp)
   next
-    case (Lab c l s0 s1 L accC T A)
+    case (Lab s0 c s1 l L accC T A)
     have     hyp: "PROP ?TypeSafe (Norm s0) s1 (In1r c) \<diamondsuit>" .
     have conf_s0: "Norm s0\<Colon>\<preceq>(G, L)" .
     moreover
@@ -1796,7 +1796,7 @@
           (error_free (Norm s0) = error_free (abupd (absorb l) s1))"
       by (simp)
   next
-    case (Comp c1 c2 s0 s1 s2 L accC T A)
+    case (Comp s0 c1 s1 c2 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>" .
@@ -1842,7 +1842,7 @@
 	using wt by simp
     qed
   next
-    case (If b c1 c2 e s0 s1 s2 L accC T)
+    case (If s0 e b s1 c1 c2 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" .
     have hyp_e: "PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 b)" .
@@ -1917,7 +1917,7 @@
           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 A)
+    case (Loop s0 e b s1 c s2 l 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)" .
@@ -2048,7 +2048,7 @@
 	by simp
     qed
   next
-    case (Jmp j s L accC T A)
+    case (Jmp s j L accC T A)
     have "Norm s\<Colon>\<preceq>(G, L)" .
     moreover
     from Jmp.prems 
@@ -2062,7 +2062,7 @@
            (error_free (Norm s) = error_free (Some (Jump j), s))"
       by simp
   next
-    case (Throw a e s0 s1 L accC T A)
+    case (Throw s0 e a 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)" .
@@ -2090,7 +2090,7 @@
             (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 A)
+    case (Try s0 c1 s1 s2 catchC vn c2 s3 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>" .
@@ -2207,7 +2207,7 @@
       qed
     qed
   next
-    case (Fin c1 c2 s0 s1 s2 s3 x1 L accC T A)
+    case (Fin s0 c1 x1 s1 c2 s2 s3 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) 
@@ -2276,7 +2276,7 @@
 	by (cases s2) auto
     qed
   next
-    case (Init C c s0 s1 s2 s3 L accC T)
+    case (Init C c s0 s3 s1 s2 L accC T A)
     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" .
@@ -2382,7 +2382,7 @@
 	by simp
     qed
   next
-    case (NewC C a s0 s1 s2 L accC T A)
+    case (NewC s0 C s1 a 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>" .
@@ -2414,7 +2414,7 @@
           (error_free (Norm s0) = error_free s2)"
       by auto
   next
-    case (NewA elT a e i s0 s1 s2 s3 L accC T A)
+    case (NewA s0 elT s1 e i s2 a 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".
@@ -2485,7 +2485,7 @@
           (error_free (Norm s0) = error_free s3) "
       by simp
   next
-    case (Cast castT e s0 s1 s2 v L accC T A)
+    case (Cast s0 e v s1 s2 castT 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)" .
@@ -2531,7 +2531,7 @@
            (error_free (Norm s0) = error_free s2)"
       by blast
   next
-    case (Inst instT b e s0 s1 v L accC T A)
+    case (Inst s0 e v s1 b instT 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
@@ -2555,7 +2555,7 @@
       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 A)
+    case (UnOp s0 e v s1 unop 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" .
@@ -2582,7 +2582,7 @@
      error_free (Norm s0) = error_free s1"
       by simp
   next
-    case (BinOp binop e1 e2 s0 s1 s2 v1 v2 L accC T A)
+    case (BinOp s0 e1 v1 s1 binop e2 v2 s2 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)" .
@@ -2698,7 +2698,7 @@
            (error_free (Norm s) = error_free (Norm s))"
       by simp
   next
-    case (Acc upd s0 s1 w v L accC T A) 
+    case (Acc s0 v w upd s1 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
@@ -2738,7 +2738,7 @@
     with conf_s1 error_free_s1 show ?case
       by simp
   next
-    case (Ass e upd s0 s1 s2 v var w L accC T A)
+    case (Ass s0 var w upd s1 e v s2 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,upd))" .
@@ -2895,7 +2895,7 @@
       qed
     qed
   next
-    case (Cond b e0 e1 e2 s0 s1 s2 v L accC T A)
+    case (Cond s0 e0 b s1 e1 e2 v s2 L accC T A)
     have eval_e0: "G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1" .
     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)" .
@@ -2988,8 +2988,8 @@
       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 A)
+    case (Call s0 e a s1 args vs s2 invDeclC mode statT mn pTs' s3 s3' accC'
+           v s4 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 
@@ -3332,7 +3332,7 @@
       qed
     qed
   next
-    case (Methd D s0 s1 sig v L accC T A)
+    case (Methd s0 D sig v s1 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)" .
@@ -3357,7 +3357,7 @@
       using hyp [of _ _ "(Inl bodyT)"] conf_s0 
       by (auto simp add: Let_def body_def)
   next
-    case (Body D c s0 s1 s2 s3 L accC T A)
+    case (Body s0 D s1 c 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>" .
@@ -3481,7 +3481,7 @@
                  (error_free (Norm s) = error_free (Norm s))"
       by (simp add: lvar_def) 
   next
-    case (FVar a accC e fn s0 s1 s2 s2' s3 stat statDeclC v L accC' T A)
+    case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC 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" .
@@ -3597,7 +3597,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 A)
+    case (AVar s0 e1 a s1 e2 i s2 v s2' 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)" .
@@ -3697,7 +3697,7 @@
     then show ?case
       by (auto elim!: wt_elim_cases)
   next
-    case (Cons e es s0 s1 s2 v vs L accC T A)
+    case (Cons s0 e v s1 es vs s2 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)" .
@@ -3905,7 +3905,7 @@
   next
     case Skip from skip show ?case by simp
   next
-    case (Expr e s0 s1 v L accC T A) 
+    case (Expr s0 e v s1 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) 
@@ -3919,7 +3919,7 @@
     ultimately show ?case
       by (rule expr)
   next
-    case (Lab c l s0 s1 L accC T A)
+    case (Lab s0 c s1 l 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)
@@ -3933,7 +3933,7 @@
     ultimately show ?case
       by (rule lab)
   next
-    case (Comp c1 c2 s0 s1 s2 L accC T A) 
+    case (Comp s0 c1 s1 c2 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 
@@ -3976,7 +3976,7 @@
     show ?case
       by (rule comp) iprover+
   next
-    case (If b c1 c2 e s0 s1 s2 L accC T A)
+    case (If s0 e b s1 c1 c2 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
--- a/src/HOL/Bali/WellForm.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/WellForm.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -883,7 +883,7 @@
     show "?Overrides new old" 
       by (auto intro: overridesR.Direct stat_override_declclasses_relation) 
   next
-    case (Indirect inter new old)
+    case (Indirect new inter old)
     then show "?Overrides new old" 
       by (blast intro: overridesR.Indirect) 
   qed
@@ -1138,7 +1138,7 @@
       qed
     qed
   next
-    case (Indirect inter new old)
+    case (Indirect new inter old)
     assume accmodi_old: "accmodi old \<noteq> Package"
     assume "accmodi old \<noteq> Package \<Longrightarrow> G \<turnstile> inter overrides\<^sub>S old"
     with accmodi_old 
@@ -1244,7 +1244,7 @@
     then show "pid (declclass old) =  pid (declclass new)"
       by (auto simp add: inheritable_in_def)
   next
-    case (Indirect inter new old)
+    case (Indirect new inter old)
     assume accmodi_old: "accmodi old = Package" and
            accmodi_new: "accmodi new = Package" 
     assume "G \<turnstile> new overrides inter"
@@ -1280,7 +1280,7 @@
     show "?P new old"
       by (contradiction)
   next
-    case (Indirect inter new old)
+    case (Indirect new inter old)
     assume accmodi_old: "accmodi old = Package"
     assume outside_pack: "pid (declclass old) \<noteq> pid (declclass new)"
     assume override_new_inter: "G \<turnstile> new overrides inter"
@@ -2554,13 +2554,13 @@
   from stat_acc subclseq 
   show ?thesis (is "?Dyn_accessible m")
   proof (induct rule: accessible_fromR.induct)
-    case (Immediate statC m)
+    case (Immediate m statC)
     then show "?Dyn_accessible m"
       by (blast intro: dyn_accessible_fromR.Immediate
                        member_inI
                        permits_acc_inheritance)
   next
-    case (Overriding _ _ m)
+    case (Overriding m _ _)
     with wf show "?Dyn_accessible m"
       by (blast intro: dyn_accessible_fromR.Overriding
                        member_inI
@@ -2921,7 +2921,7 @@
   from dyn_acc priv
   show ?thesis
   proof (induct)
-    case (Immediate C m)
+    case (Immediate m C)
     have "G \<turnstile> m in C permits_acc_from accC" and "accmodi m = Private" .
     then show ?case
       by (simp add: permits_acc_def)
@@ -2946,14 +2946,14 @@
             \<Longrightarrow> pid accC = pid (declclass m)"
     (is "?Pack m \<Longrightarrow> ?P m")
   proof (induct rule: dyn_accessible_fromR.induct)
-    case (Immediate C m)
+    case (Immediate m C)
     assume "G\<turnstile>m member_in C"
            "G \<turnstile> m in C permits_acc_from accC"
            "accmodi m = Package"      
     then show "?P m"
       by (auto simp add: permits_acc_def)
   next
-    case (Overriding declC C new newm old Sup)
+    case (Overriding new C declC newm old Sup)
     assume member_new: "G \<turnstile> new member_in C" and
                   new: "new = (declC, mdecl newm)" and
              override: "G \<turnstile> (declC, newm) overrides old" and
@@ -2986,7 +2986,7 @@
   from dyn_acc pack field
   show ?thesis
   proof (induct)
-    case (Immediate C f)
+    case (Immediate f C)
     have "G \<turnstile> f in C permits_acc_from accC" and "accmodi f = Package" .
     then show ?case
       by (simp add: permits_acc_def)
@@ -3010,7 +3010,7 @@
   from dyn_acc prot field instance_field outside
   show ?thesis
   proof (induct)
-    case (Immediate C f)
+    case (Immediate f C)
     have "G \<turnstile> f in C permits_acc_from accC" .
     moreover 
     assume "accmodi f = Protected" and  "is_field f" and "\<not> is_static f" and
@@ -3035,7 +3035,7 @@
   from dyn_acc prot field static_field outside
   show ?thesis
   proof (induct)
-    case (Immediate C f)
+    case (Immediate f C)
     assume "accmodi f = Protected" and  "is_field f" and "is_static f" and
            "pid (declclass f) \<noteq> pid accC"
     moreover 
--- a/src/HOL/Bali/WellType.thy	Mon Dec 11 12:28:16 2006 +0100
+++ b/src/HOL/Bali/WellType.thy	Mon Dec 11 16:06:14 2006 +0100
@@ -245,33 +245,185 @@
 translations
   "tys"   <= (type) "ty + ty list"
 
-consts
-  wt    :: "(env\<spacespace> \<times> dyn_ty\<spacespace> \<times>  term \<times> tys) set"
-(*wt    :: " env \<Rightarrow> dyn_ty \<Rightarrow> (term \<times> tys) set" not feasible because of 
-					      \<spacespace>  changing env in Try stmt *)
+
+inductive2 
+  wt :: "env \<Rightarrow> dyn_ty \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>_"  [51,51,51,51] 50)
+  and wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow>  stmt       \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>\<surd>"  [51,51,51   ] 50)
+  and ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>-_" [51,51,51,51] 50)
+  and ty_var :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>=_" [51,51,51,51] 50)
+  and ty_exprs :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr list, ty   list] \<Rightarrow> bool"
+    ("_,_\<Turnstile>_\<Colon>\<doteq>_" [51,51,51,51] 50)
+where
+
+  "E,dt\<Turnstile>s\<Colon>\<surd> \<equiv> E,dt\<Turnstile>In1r s\<Colon>Inl (PrimT Void)"
+| "E,dt\<Turnstile>e\<Colon>-T \<equiv> E,dt\<Turnstile>In1l e\<Colon>Inl T"
+| "E,dt\<Turnstile>e\<Colon>=T \<equiv> E,dt\<Turnstile>In2  e\<Colon>Inl T"
+| "E,dt\<Turnstile>e\<Colon>\<doteq>T \<equiv> E,dt\<Turnstile>In3  e\<Colon>Inr T"
+
+--{* well-typed statements *}
+
+| Skip:					"E,dt\<Turnstile>Skip\<Colon>\<surd>"
+
+| Expr:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Expr e\<Colon>\<surd>"
+  --{* cf. 14.6 *}
+| Lab:  "E,dt\<Turnstile>c\<Colon>\<surd> \<Longrightarrow>                   
+                                         E,dt\<Turnstile>l\<bullet> c\<Colon>\<surd>" 
+
+| Comp:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; 
+	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>c1;; c2\<Colon>\<surd>"
+
+  --{* cf. 14.8 *}
+| If:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
+	  E,dt\<Turnstile>c1\<Colon>\<surd>;
+	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>If(e) c1 Else c2\<Colon>\<surd>"
+
+  --{* cf. 14.10 *}
+| Loop:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
+	  E,dt\<Turnstile>c\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>l\<bullet> While(e) c\<Colon>\<surd>"
+  --{* cf. 14.13, 14.15, 14.16 *}
+| Jmp:                                   "E,dt\<Turnstile>Jmp jump\<Colon>\<surd>"
 
-syntax
-wt      :: "env \<Rightarrow> dyn_ty \<Rightarrow> [term,tys]  \<Rightarrow> bool" ("_,_|=_::_" [51,51,51,51] 50)
-wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow>  stmt       \<Rightarrow> bool" ("_,_|=_:<>" [51,51,51   ] 50)
-ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_,_|=_:-_" [51,51,51,51] 50)
-ty_var  :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_,_|=_:=_" [51,51,51,51] 50)
-ty_exprs:: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr list,
-	         \<spacespace>        \<spacespace>  ty   list] \<Rightarrow> bool" ("_,_|=_:#_" [51,51,51,51] 50)
+  --{* cf. 14.16 *}
+| Throw: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class tn;
+	  prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Throw e\<Colon>\<surd>"
+  --{* cf. 14.18 *}
+| Try:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable;
+	  lcl E (VName vn)=None; E \<lparr>lcl := lcl E(VName vn\<mapsto>Class tn)\<rparr>,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk>
+          \<Longrightarrow>
+					 E,dt\<Turnstile>Try c1 Catch(tn vn) c2\<Colon>\<surd>"
+
+  --{* cf. 14.18 *}
+| Fin:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>c1 Finally c2\<Colon>\<surd>"
+
+| Init:	"\<lbrakk>is_class (prg E) C\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Init C\<Colon>\<surd>"
+  --{* @{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 *}
+
+  --{* cf. 15.8 *}
+| NewC:	"\<lbrakk>is_acc_class (prg E) (pkg E) C\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>NewC C\<Colon>-Class C"
+  --{* cf. 15.9 *}
+| NewA:	"\<lbrakk>is_acc_type (prg E) (pkg E) T;
+	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>New T[i]\<Colon>-T.[]"
+
+  --{* cf. 15.15 *}
+| Cast:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T; is_acc_type (prg E) (pkg E) T';
+	  prg E\<turnstile>T\<preceq>? T'\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Cast T' e\<Colon>-T'"
+
+  --{* cf. 15.19.2 *}
+| Inst:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT T; is_acc_type (prg E) (pkg E) (RefT T');
+	  prg E\<turnstile>RefT T\<preceq>? RefT T'\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>e InstOf T'\<Colon>-PrimT Boolean"
+
+  --{* cf. 15.7.1 *}
+| Lit:	"\<lbrakk>typeof dt x = Some T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Lit x\<Colon>-T"
 
-syntax (xsymbols)
-wt      :: "env \<Rightarrow> dyn_ty \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>_"  [51,51,51,51] 50)
-wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow>  stmt       \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>\<surd>"  [51,51,51   ] 50)
-ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>-_" [51,51,51,51] 50)
-ty_var  :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>=_" [51,51,51,51] 50)
-ty_exprs:: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr list,
-		  \<spacespace>        \<spacespace>  ty   list] \<Rightarrow> bool"("_,_\<Turnstile>_\<Colon>\<doteq>_" [51,51,51,51] 50)
+| UnOp: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Te; wt_unop unop Te; T=PrimT (unop_type unop)\<rbrakk> 
+          \<Longrightarrow>
+	  E,dt\<Turnstile>UnOp unop e\<Colon>-T"
+  
+| BinOp: "\<lbrakk>E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2; wt_binop (prg E) binop T1 T2; 
+           T=PrimT (binop_type binop)\<rbrakk> 
+           \<Longrightarrow>
+	   E,dt\<Turnstile>BinOp binop e1 e2\<Colon>-T"
+  
+  --{* cf. 15.10.2, 15.11.1 *}
+| Super: "\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object;
+	  class (prg E) C = Some c\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Super\<Colon>-Class (super c)"
+
+  --{* cf. 15.13.1, 15.10.1, 15.12 *}
+| Acc:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Acc va\<Colon>-T"
+
+  --{* cf. 15.25, 15.25.1 *}
+| Ass:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T; va \<noteq> LVar This;
+	  E,dt\<Turnstile>v \<Colon>-T';
+	  prg E\<turnstile>T'\<preceq>T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>va:=v\<Colon>-T'"
+
+  --{* cf. 15.24 *}
+| Cond:	"\<lbrakk>E,dt\<Turnstile>e0\<Colon>-PrimT Boolean;
+	  E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2;
+	  prg E\<turnstile>T1\<preceq>T2 \<and> T = T2  \<or>  prg E\<turnstile>T2\<preceq>T1 \<and> T = T1\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>e0 ? e1 : e2\<Colon>-T"
+
+  --{* cf. 15.11.1, 15.11.2, 15.11.3 *}
+| Call:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT;
+	  E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;
+	  max_spec (prg E) (cls E) statT \<lparr>name=mn,parTs=pTs\<rparr> 
+            = {((statDeclT,m),pTs')}
+         \<rbrakk> \<Longrightarrow>
+		   E,dt\<Turnstile>{cls E,statT,invmode m e}e\<cdot>mn({pTs'}ps)\<Colon>-(resTy m)"
 
-translations
-	"E,dt\<Turnstile>t\<Colon>T" == "(E,dt,t,T) \<in> wt"
-	"E,dt\<Turnstile>s\<Colon>\<surd>"  == "E,dt\<Turnstile>In1r s\<Colon>Inl (PrimT Void)"
-	"E,dt\<Turnstile>e\<Colon>-T" == "E,dt\<Turnstile>In1l e\<Colon>Inl T"
-	"E,dt\<Turnstile>e\<Colon>=T" == "E,dt\<Turnstile>In2  e\<Colon>Inl T"
-	"E,dt\<Turnstile>e\<Colon>\<doteq>T" == "E,dt\<Turnstile>In3  e\<Colon>Inr T"
+| Methd: "\<lbrakk>is_class (prg E) C;
+	  methd (prg E) C sig = Some m;
+	  E,dt\<Turnstile>Body (declclass m) (stmt (mbody (mthd m)))\<Colon>-T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>Methd C sig\<Colon>-T"
+ --{* The class @{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 @{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 *}
+
+  --{* cf. 15.13.1 *}
+| LVar:	"\<lbrakk>lcl E vn = Some T; is_acc_type (prg E) (pkg E) T\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>LVar vn\<Colon>=T"
+  --{* cf. 15.10.1 *}
+| FVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; 
+	  accfield (prg E) (cls E) C fn = Some (statDeclC,f)\<rbrakk> \<Longrightarrow>
+			 E,dt\<Turnstile>{cls E,statDeclC,is_static f}e..fn\<Colon>=(type f)"
+  --{* cf. 15.12 *}
+| AVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T.[]; 
+	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>e.[i]\<Colon>=T"
+
+
+--{* well-typed expression lists *}
+
+  --{* cf. 15.11.??? *}
+| Nil:					"E,dt\<Turnstile>[]\<Colon>\<doteq>[]"
+
+  --{* cf. 15.11.??? *}
+| Cons:	"\<lbrakk>E,dt\<Turnstile>e \<Colon>-T;
+	  E,dt\<Turnstile>es\<Colon>\<doteq>Ts\<rbrakk> \<Longrightarrow>
+					 E,dt\<Turnstile>e#es\<Colon>\<doteq>T#Ts"
+
 
 syntax (* for purely static typing *)
   wt_      :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_|-_::_" [51,51,51] 50)
@@ -297,174 +449,6 @@
 	"E\<turnstile>e\<Colon>\<doteq>T" == "E\<turnstile>In3  e\<Colon>Inr T"
 
 
-inductive wt intros 
-
-
---{* well-typed statements *}
-
-  Skip:					"E,dt\<Turnstile>Skip\<Colon>\<surd>"
-
-  Expr:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Expr e\<Colon>\<surd>"
-  --{* cf. 14.6 *}
-  Lab:  "E,dt\<Turnstile>c\<Colon>\<surd> \<Longrightarrow>                   
-                                         E,dt\<Turnstile>l\<bullet> c\<Colon>\<surd>" 
-
-  Comp:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; 
-	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>c1;; c2\<Colon>\<surd>"
-
-  --{* cf. 14.8 *}
-  If:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
-	  E,dt\<Turnstile>c1\<Colon>\<surd>;
-	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>If(e) c1 Else c2\<Colon>\<surd>"
-
-  --{* cf. 14.10 *}
-  Loop:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
-	  E,dt\<Turnstile>c\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>l\<bullet> While(e) c\<Colon>\<surd>"
-  --{* cf. 14.13, 14.15, 14.16 *}
-  Jmp:                                   "E,dt\<Turnstile>Jmp jump\<Colon>\<surd>"
-
-  --{* cf. 14.16 *}
-  Throw: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class tn;
-	  prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Throw e\<Colon>\<surd>"
-  --{* cf. 14.18 *}
-  Try:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable;
-	  lcl E (VName vn)=None; E \<lparr>lcl := lcl E(VName vn\<mapsto>Class tn)\<rparr>,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk>
-          \<Longrightarrow>
-					 E,dt\<Turnstile>Try c1 Catch(tn vn) c2\<Colon>\<surd>"
-
-  --{* cf. 14.18 *}
-  Fin:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>c1 Finally c2\<Colon>\<surd>"
-
-  Init:	"\<lbrakk>is_class (prg E) C\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Init C\<Colon>\<surd>"
-  --{* @{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 *}
-
-  --{* cf. 15.8 *}
-  NewC:	"\<lbrakk>is_acc_class (prg E) (pkg E) C\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>NewC C\<Colon>-Class C"
-  --{* cf. 15.9 *}
-  NewA:	"\<lbrakk>is_acc_type (prg E) (pkg E) T;
-	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>New T[i]\<Colon>-T.[]"
-
-  --{* cf. 15.15 *}
-  Cast:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T; is_acc_type (prg E) (pkg E) T';
-	  prg E\<turnstile>T\<preceq>? T'\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Cast T' e\<Colon>-T'"
-
-  --{* cf. 15.19.2 *}
-  Inst:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT T; is_acc_type (prg E) (pkg E) (RefT T');
-	  prg E\<turnstile>RefT T\<preceq>? RefT T'\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>e InstOf T'\<Colon>-PrimT Boolean"
-
-  --{* cf. 15.7.1 *}
-  Lit:	"\<lbrakk>typeof dt x = Some T\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Lit x\<Colon>-T"
-
-  UnOp: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Te; wt_unop unop Te; T=PrimT (unop_type unop)\<rbrakk> 
-          \<Longrightarrow>
-	  E,dt\<Turnstile>UnOp unop e\<Colon>-T"
-  
-  BinOp: "\<lbrakk>E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2; wt_binop (prg E) binop T1 T2; 
-           T=PrimT (binop_type binop)\<rbrakk> 
-           \<Longrightarrow>
-	   E,dt\<Turnstile>BinOp binop e1 e2\<Colon>-T"
-  
-  --{* cf. 15.10.2, 15.11.1 *}
-  Super: "\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object;
-	  class (prg E) C = Some c\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Super\<Colon>-Class (super c)"
-
-  --{* cf. 15.13.1, 15.10.1, 15.12 *}
-  Acc:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Acc va\<Colon>-T"
-
-  --{* cf. 15.25, 15.25.1 *}
-  Ass:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T; va \<noteq> LVar This;
-	  E,dt\<Turnstile>v \<Colon>-T';
-	  prg E\<turnstile>T'\<preceq>T\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>va:=v\<Colon>-T'"
-
-  --{* cf. 15.24 *}
-  Cond:	"\<lbrakk>E,dt\<Turnstile>e0\<Colon>-PrimT Boolean;
-	  E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2;
-	  prg E\<turnstile>T1\<preceq>T2 \<and> T = T2  \<or>  prg E\<turnstile>T2\<preceq>T1 \<and> T = T1\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>e0 ? e1 : e2\<Colon>-T"
-
-  --{* cf. 15.11.1, 15.11.2, 15.11.3 *}
-  Call:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT;
-	  E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;
-	  max_spec (prg E) (cls E) statT \<lparr>name=mn,parTs=pTs\<rparr> 
-            = {((statDeclT,m),pTs')}
-         \<rbrakk> \<Longrightarrow>
-		   E,dt\<Turnstile>{cls E,statT,invmode m e}e\<cdot>mn({pTs'}ps)\<Colon>-(resTy m)"
-
-  Methd: "\<lbrakk>is_class (prg E) C;
-	  methd (prg E) C sig = Some m;
-	  E,dt\<Turnstile>Body (declclass m) (stmt (mbody (mthd m)))\<Colon>-T\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Methd C sig\<Colon>-T"
- --{* 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 @{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 *}
-
-  --{* cf. 15.13.1 *}
-  LVar:	"\<lbrakk>lcl E vn = Some T; is_acc_type (prg E) (pkg E) T\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>LVar vn\<Colon>=T"
-  --{* cf. 15.10.1 *}
-  FVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; 
-	  accfield (prg E) (cls E) C fn = Some (statDeclC,f)\<rbrakk> \<Longrightarrow>
-			 E,dt\<Turnstile>{cls E,statDeclC,is_static f}e..fn\<Colon>=(type f)"
-  --{* cf. 15.12 *}
-  AVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T.[]; 
-	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>e.[i]\<Colon>=T"
-
-
---{* well-typed expression lists *}
-
-  --{* cf. 15.11.??? *}
-  Nil:					"E,dt\<Turnstile>[]\<Colon>\<doteq>[]"
-
-  --{* cf. 15.11.??? *}
-  Cons:	"\<lbrakk>E,dt\<Turnstile>e \<Colon>-T;
-	  E,dt\<Turnstile>es\<Colon>\<doteq>Ts\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>e#es\<Colon>\<doteq>T#Ts"
-
-
 declare not_None_eq [simp del] 
 declare split_if [split del] split_if_asm [split del]
 declare split_paired_All [simp del] split_paired_Ex [simp del]
@@ -472,7 +456,7 @@
 change_simpset (fn ss => ss delloop "split_all_tac")
 *}
 
-inductive_cases wt_elim_cases [cases set]:
+inductive_cases2 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"
@@ -610,25 +594,32 @@
     correlation. 
  *}
 
+lemma wt_expr_eq: "E,dt\<Turnstile>In1l t\<Colon>U = (\<exists>T. U=Inl T \<and> E,dt\<Turnstile>t\<Colon>-T)"
+  by (auto, frule wt_Inj_elim, auto)
+
+lemma wt_var_eq: "E,dt\<Turnstile>In2 t\<Colon>U = (\<exists>T. U=Inl T \<and> E,dt\<Turnstile>t\<Colon>=T)"
+  by (auto, frule wt_Inj_elim, auto)
+
+lemma wt_exprs_eq: "E,dt\<Turnstile>In3 t\<Colon>U = (\<exists>Ts. U=Inr Ts \<and> E,dt\<Turnstile>t\<Colon>\<doteq>Ts)"
+  by (auto, frule wt_Inj_elim, auto)
+
+lemma wt_stmt_eq: "E,dt\<Turnstile>In1r t\<Colon>U = (U=Inl(PrimT Void)\<and>E,dt\<Turnstile>t\<Colon>\<surd>)"
+  by (auto, frule wt_Inj_elim, auto, frule wt_Inj_elim, auto)
+
 ML {*
-fun wt_fun name inj rhs =
+fun wt_fun name lhs =
 let
-  val lhs = "E,dt\<Turnstile>" ^ inj ^ " t\<Colon>U"
-  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
-	(K [Auto_tac, ALLGOALS (ftac (thm "wt_Inj_elim")) THEN Auto_tac])
   fun is_Inj (Const (inj,_) $ _) = true
     | is_Inj _                   = false
-  fun pred (t as (_ $ (Const ("Pair",_) $
-     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
-       x))) $ _ )) = is_Inj x
+  fun pred (_ $ _ $ _ $ _ $ x) = is_Inj x
 in
   cond_simproc name lhs pred (thm name)
 end
 
-val wt_expr_proc  = wt_fun "wt_expr_eq"  "In1l" "\<exists>T.  U=Inl T  \<and> E,dt\<Turnstile>t\<Colon>-T"
-val wt_var_proc   = wt_fun "wt_var_eq"   "In2"  "\<exists>T.  U=Inl T  \<and> E,dt\<Turnstile>t\<Colon>=T"
-val wt_exprs_proc = wt_fun "wt_exprs_eq" "In3"  "\<exists>Ts. U=Inr Ts \<and> E,dt\<Turnstile>t\<Colon>\<doteq>Ts"
-val wt_stmt_proc  = wt_fun "wt_stmt_eq"  "In1r" "U=Inl(PrimT Void)\<and>E,dt\<Turnstile>t\<Colon>\<surd>"
+val wt_expr_proc  = wt_fun "wt_expr_eq"  "E,dt\<Turnstile>In1l t\<Colon>U";
+val wt_var_proc   = wt_fun "wt_var_eq"   "E,dt\<Turnstile>In2 t\<Colon>U";
+val wt_exprs_proc = wt_fun "wt_exprs_eq" "E,dt\<Turnstile>In3 t\<Colon>U";
+val wt_stmt_proc  = wt_fun "wt_stmt_eq"  "E,dt\<Turnstile>In1r t\<Colon>U";
 *}
 
 ML {*