merged
authorwenzelm
Sat, 17 Oct 2009 14:51:30 +0200
changeset 32963 fb05ae5cd343
parent 32962 4772d8dd18f8 (current diff)
parent 32960 69916a850301 (diff)
child 32964 2d7e1ab55037
merged
--- a/doc-src/Intro/bool.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/doc-src/Intro/bool.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,5 @@
 Bool = FOL +
-types 	bool 0
-arities bool 	:: term
-consts tt,ff	:: "bool"
+types   bool 0
+arities bool    :: term
+consts tt,ff    :: "bool"
 end
--- a/doc-src/Intro/list.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/doc-src/Intro/list.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,6 @@
 List = FOL +
-types 	list 1
-arities list	:: (term)term
-consts	Nil	:: "'a list"
-   	Cons	:: "['a, 'a list] => 'a list" 
+types   list 1
+arities list    :: (term)term
+consts  Nil     :: "'a list"
+        Cons    :: "['a, 'a list] => 'a list" 
 end
--- a/doc-src/Locales/Locales/Examples.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/doc-src/Locales/Locales/Examples.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -171,13 +171,13 @@
     proof (rule anti_sym)
       from inf' show "i \<sqsubseteq> i'"
       proof (rule is_inf_greatest)
-	from inf show "i \<sqsubseteq> x" ..
-	from inf show "i \<sqsubseteq> y" ..
+        from inf show "i \<sqsubseteq> x" ..
+        from inf show "i \<sqsubseteq> y" ..
       qed
       from inf show "i' \<sqsubseteq> i"
       proof (rule is_inf_greatest)
-	from inf' show "i' \<sqsubseteq> x" ..
-	from inf' show "i' \<sqsubseteq> y" ..
+        from inf' show "i' \<sqsubseteq> x" ..
+        from inf' show "i' \<sqsubseteq> y" ..
       qed
     qed
   qed
@@ -213,13 +213,13 @@
     proof (rule anti_sym)
       from sup show "s \<sqsubseteq> s'"
       proof (rule is_sup_least)
-	from sup' show "x \<sqsubseteq> s'" ..
-	from sup' show "y \<sqsubseteq> s'" ..
+        from sup' show "x \<sqsubseteq> s'" ..
+        from sup' show "y \<sqsubseteq> s'" ..
       qed
       from sup' show "s' \<sqsubseteq> s"
       proof (rule is_sup_least)
-	from sup show "x \<sqsubseteq> s" ..
-	from sup show "y \<sqsubseteq> s" ..
+        from sup show "x \<sqsubseteq> s" ..
+        from sup show "y \<sqsubseteq> s" ..
       qed
     qed
   qed
@@ -346,9 +346,9 @@
       show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> x" ..
       show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y"
       proof -
-	have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
-	also have "\<dots> \<sqsubseteq> y" ..
-	finally show ?thesis .
+        have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
+        also have "\<dots> \<sqsubseteq> y" ..
+        finally show ?thesis .
       qed
     qed
     show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> z"
@@ -362,19 +362,19 @@
     proof
       show "w \<sqsubseteq> x"
       proof -
-	have "w \<sqsubseteq> x \<sqinter> y" by fact
-	also have "\<dots> \<sqsubseteq> x" ..
-	finally show ?thesis .
+        have "w \<sqsubseteq> x \<sqinter> y" by fact
+        also have "\<dots> \<sqsubseteq> x" ..
+        finally show ?thesis .
       qed
       show "w \<sqsubseteq> y \<sqinter> z"
       proof
-	show "w \<sqsubseteq> y"
-	proof -
-	  have "w \<sqsubseteq> x \<sqinter> y" by fact
-	  also have "\<dots> \<sqsubseteq> y" ..
-	  finally show ?thesis .
-	qed
-	show "w \<sqsubseteq> z" by fact
+        show "w \<sqsubseteq> y"
+        proof -
+          have "w \<sqsubseteq> x \<sqinter> y" by fact
+          also have "\<dots> \<sqsubseteq> y" ..
+          finally show ?thesis .
+        qed
+        show "w \<sqsubseteq> z" by fact
       qed
     qed
   qed
@@ -402,9 +402,9 @@
       show "x \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
       show "y \<sqsubseteq> x \<squnion> (y \<squnion> z)"
       proof -
-	have "y \<sqsubseteq> y \<squnion> z" ..
-	also have "... \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
-	finally show ?thesis .
+        have "y \<sqsubseteq> y \<squnion> z" ..
+        also have "... \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
+        finally show ?thesis .
       qed
     qed
     show "z \<sqsubseteq> x \<squnion> (y \<squnion> z)"
@@ -418,19 +418,19 @@
     proof
       show "x \<sqsubseteq> w"
       proof -
-	have "x \<sqsubseteq> x \<squnion> y" ..
-	also have "\<dots> \<sqsubseteq> w" by fact
-	finally show ?thesis .
+        have "x \<sqsubseteq> x \<squnion> y" ..
+        also have "\<dots> \<sqsubseteq> w" by fact
+        finally show ?thesis .
       qed
       show "y \<squnion> z \<sqsubseteq> w"
       proof
-	show "y \<sqsubseteq> w"
-	proof -
-	  have "y \<sqsubseteq> x \<squnion> y" ..
-	  also have "... \<sqsubseteq> w" by fact
-	  finally show ?thesis .
-	qed
-	show "z \<sqsubseteq> w" by fact
+        show "y \<sqsubseteq> w"
+        proof -
+          have "y \<sqsubseteq> x \<squnion> y" ..
+          also have "... \<sqsubseteq> w" by fact
+          finally show ?thesis .
+        qed
+        show "z \<sqsubseteq> w" by fact
       qed
     qed
   qed
@@ -650,17 +650,17 @@
       txt {* Jacobson I, p.\ 462 *}
     proof -
       { assume c: "y \<sqsubseteq> x" "z \<sqsubseteq> x"
-	from c have "?l = y \<squnion> z"
-	  by (metis c join_connection2 join_related2 meet_related2 total)
-	also from c have "... = ?r" by (metis meet_related2)
-	finally have "?l = ?r" . }
+        from c have "?l = y \<squnion> z"
+          by (metis c join_connection2 join_related2 meet_related2 total)
+        also from c have "... = ?r" by (metis meet_related2)
+        finally have "?l = ?r" . }
       moreover
       { assume c: "x \<sqsubseteq> y \<or> x \<sqsubseteq> z"
-	from c have "?l = x"
-	  by (metis join_connection2 join_related2 meet_connection total trans)
-	also from c have "... = ?r"
-	  by (metis join_commute join_related2 meet_connection meet_related2 total)
-	finally have "?l = ?r" . }
+        from c have "?l = x"
+          by (metis join_connection2 join_related2 meet_connection total trans)
+        also from c have "... = ?r"
+          by (metis join_commute join_related2 meet_connection meet_related2 total)
+        finally have "?l = ?r" . }
       moreover note total
       ultimately show ?thesis by blast
     qed
--- a/doc-src/TutorialI/Overview/LNCS/Sets.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/doc-src/TutorialI/Overview/LNCS/Sets.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -149,7 +149,7 @@
     proof (rule converse_rtrancl_induct)
       show "t \<in> lfp ?F"
       proof (subst lfp_unfold[OF mono_ef])
-	show "t \<in> ?F(lfp ?F)" using tA by blast
+        show "t \<in> ?F(lfp ?F)" using tA by blast
       qed
     next
       fix s s'
@@ -157,7 +157,7 @@
          and IH: "s' \<in> lfp ?F"
       show "s \<in> lfp ?F"
       proof (subst lfp_unfold[OF mono_ef])
-	show "s \<in> ?F(lfp ?F)" using prems by blast
+        show "s \<in> ?F(lfp ?F)" using prems by blast
       qed
     qed
   qed
--- a/doc-src/TutorialI/Protocol/Event.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Event.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -22,7 +22,7 @@
         | Notes agent       msg
        
 consts 
-  bad    :: "agent set"				(*compromised agents*)
+  bad    :: "agent set"                         -- {* compromised agents *}
   knows  :: "agent => event list => msg set"
 
 
@@ -43,19 +43,19 @@
   knows_Cons:
     "knows A (ev # evs) =
        (if A = Spy then 
-	(case ev of
-	   Says A' B X => insert X (knows Spy evs)
-	 | Gets A' X => knows Spy evs
-	 | Notes A' X  => 
-	     if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
-	else
-	(case ev of
-	   Says A' B X => 
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Gets A' X    => 
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Notes A' X    => 
-	     if A'=A then insert X (knows A evs) else knows A evs))"
+        (case ev of
+           Says A' B X => insert X (knows Spy evs)
+         | Gets A' X => knows Spy evs
+         | Notes A' X  => 
+             if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
+        else
+        (case ev of
+           Says A' B X => 
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Gets A' X    => 
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Notes A' X    => 
+             if A'=A then insert X (knows A evs) else knows A evs))"
 
 (*
   Case A=Spy on the Gets event
@@ -71,10 +71,10 @@
 primrec
   used_Nil:   "used []         = (UN B. parts (initState B))"
   used_Cons:  "used (ev # evs) =
-		     (case ev of
-			Says A B X => parts {X} \<union> used evs
-		      | Gets A X   => used evs
-		      | Notes A X  => parts {X} \<union> used evs)"
+                     (case ev of
+                        Says A B X => parts {X} \<union> used evs
+                      | Gets A X   => used evs
+                      | Notes A X  => parts {X} \<union> used evs)"
     --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
         follows @{term Says} in real protocols.  Seems difficult to change.
         See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
--- a/doc-src/TutorialI/Protocol/Message.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Message.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -61,8 +61,8 @@
      msg = Agent  agent
          | Nonce  nat
          | Key    key
-	 | MPair  msg msg
-	 | Crypt  key msg
+         | MPair  msg msg
+         | Crypt  key msg
 
 text {*
 \noindent
@@ -855,8 +855,8 @@
     (Fake_insert_simp_tac ss 1
      THEN
      IF_UNSOLVED (Blast.depth_tac
-		  (cs addIs [analz_insertI,
-				   impOfSubs analz_subset_parts]) 4 1))
+                  (cs addIs [analz_insertI,
+                                   impOfSubs analz_subset_parts]) 4 1))
 
 fun spy_analz_tac (cs,ss) i =
   DETERM
--- a/doc-src/TutorialI/Protocol/NS_Public.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/doc-src/TutorialI/Protocol/NS_Public.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -221,8 +221,8 @@
 lemma A_trusts_NS2_lemma [rule_format]:
    "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
     \<Longrightarrow> Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
-	Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
-	Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
+        Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
+        Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake, NS1*)
 apply (blast dest: Spy_not_see_NA)+
@@ -240,8 +240,8 @@
 lemma B_trusts_NS1 [rule_format]:
      "evs \<in> ns_public
       \<Longrightarrow> Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
-	  Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
-	  Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
+          Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
+          Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake*)
 apply (blast intro!: analz_insertI)
--- a/doc-src/TutorialI/Protocol/Public.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Public.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -25,11 +25,11 @@
 primrec
         (*Agents know their private key and all public keys*)
   initState_Server:  "initState Server     =    
- 		         insert (Key (priK Server)) (Key ` range pubK)"
+                         insert (Key (priK Server)) (Key ` range pubK)"
   initState_Friend:  "initState (Friend i) =    
- 		         insert (Key (priK (Friend i))) (Key ` range pubK)"
+                         insert (Key (priK (Friend i))) (Key ` range pubK)"
   initState_Spy:     "initState Spy        =    
- 		         (Key`invKey`pubK`bad) Un (Key ` range pubK)"
+                         (Key`invKey`pubK`bad) Un (Key ` range pubK)"
 (*>*)
 
 text {*
--- a/doc-src/TutorialI/Rules/Basic.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/doc-src/TutorialI/Rules/Basic.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,4 +1,3 @@
-(* ID:         $Id$ *)
 theory Basic imports Main begin
 
 lemma conj_rule: "\<lbrakk> P; Q \<rbrakk> \<Longrightarrow> P \<and> (Q \<and> P)"
@@ -149,9 +148,9 @@
 
 lemma "\<lbrakk>\<not>(P\<longrightarrow>Q); \<not>(R\<longrightarrow>Q)\<rbrakk> \<Longrightarrow> R"
 apply (erule_tac Q="R\<longrightarrow>Q" in contrapos_np)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (intro impI)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 by (erule notE)
 
 text {*
@@ -161,11 +160,11 @@
 
 lemma "(P \<or> Q) \<and> R \<Longrightarrow> P \<or> Q \<and> R"
 apply (intro disjCI conjI)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 
 apply (elim conjE disjE)
  apply assumption
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 
 by (erule contrapos_np, rule conjI)
 text{*
@@ -241,18 +240,18 @@
 text{*rename_tac*}
 lemma "x < y \<Longrightarrow> \<forall>x y. P x (f y)"
 apply (intro allI)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (rename_tac v w)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 oops
 
 
 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
 apply (frule spec)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (drule mp, assumption)
 apply (drule spec)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 by (drule mp)
 
 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
@@ -276,11 +275,11 @@
 
 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
 apply (frule spec)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (drule mp, assumption)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (drule_tac x = "h a" in spec)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 by (drule mp)
 
 text {*
@@ -290,15 +289,15 @@
 
 lemma mult_dvd_mono: "\<lbrakk>i dvd m; j dvd n\<rbrakk> \<Longrightarrow> i*j dvd (m*n :: nat)"
 apply (simp add: dvd_def)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (erule exE) 
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (erule exE) 
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (rename_tac l)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (rule_tac x="k*l" in exI) 
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply simp
 done
 
--- a/src/FOL/FOL.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/FOL/FOL.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -174,13 +174,13 @@
   structure Blast = Blast
   (
     val thy = @{theory}
-    type claset	= Cla.claset
+    type claset = Cla.claset
     val equality_name = @{const_name "op ="}
     val not_name = @{const_name Not}
-    val notE	= @{thm notE}
-    val ccontr	= @{thm ccontr}
+    val notE = @{thm notE}
+    val ccontr = @{thm ccontr}
     val contr_tac = Cla.contr_tac
-    val dup_intr	= Cla.dup_intr
+    val dup_intr = Cla.dup_intr
     val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
     val rep_cs = Cla.rep_cs
     val cla_modifiers = Cla.cla_modifiers
--- a/src/FOL/ex/Classical.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/FOL/ex/Classical.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -418,7 +418,7 @@
 by fast
 
 text{*Halting problem: Formulation of Li Dafa (AAR Newsletter 27, Oct 1994.)
-	author U. Egly*}
+  author U. Egly*}
 lemma "((\<exists>x. A(x) & (\<forall>y. C(y) --> (\<forall>z. D(x,y,z)))) -->                
    (\<exists>w. C(w) & (\<forall>y. C(y) --> (\<forall>z. D(w,y,z)))))                   
   &                                                                      
--- a/src/FOL/simpdata.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/FOL/simpdata.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -27,7 +27,7 @@
 
 (*Congruence rules for = or <-> (instead of ==)*)
 fun mk_meta_cong rl =
-  standard(mk_meta_eq (mk_meta_prems rl))
+  Drule.standard (mk_meta_eq (mk_meta_prems rl))
   handle THM _ =>
   error("Premises and conclusion of congruence rules must use =-equality or <->");
 
--- a/src/FOLP/simp.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/FOLP/simp.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -519,7 +519,7 @@
 (* Compute Congruence rules for individual constants using the substition
    rules *)
 
-val subst_thms = map standard subst_thms;
+val subst_thms = map Drule.standard subst_thms;
 
 
 fun exp_app(0,t) = t
--- a/src/HOL/Algebra/Coset.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Algebra/Coset.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -609,7 +609,7 @@
     proof (simp add: r_congruent_def sym_def, clarify)
       fix x y
       assume [simp]: "x \<in> carrier G" "y \<in> carrier G" 
-	 and "inv x \<otimes> y \<in> H"
+         and "inv x \<otimes> y \<in> H"
       hence "inv (inv x \<otimes> y) \<in> H" by (simp add: m_inv_closed) 
       thus "inv y \<otimes> x \<in> H" by (simp add: inv_mult_group)
     qed
@@ -618,10 +618,10 @@
     proof (simp add: r_congruent_def trans_def, clarify)
       fix x y z
       assume [simp]: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
-	 and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H"
+         and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H"
       hence "(inv x \<otimes> y) \<otimes> (inv y \<otimes> z) \<in> H" by simp
       hence "inv x \<otimes> (y \<otimes> inv y) \<otimes> z \<in> H"
-	by (simp add: m_assoc del: r_inv Units_r_inv) 
+        by (simp add: m_assoc del: r_inv Units_r_inv) 
       thus "inv x \<otimes> z \<in> H" by simp
     qed
   qed
--- a/src/HOL/Algebra/Divisibility.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Algebra/Divisibility.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -2579,33 +2579,33 @@
            by force
 
       from this obtain p'
-	  where "p' \<in> set (as@bs)"
-	  and pp': "p \<sim> p'" by auto
+          where "p' \<in> set (as@bs)"
+          and pp': "p \<sim> p'" by auto
 
       hence "p' \<in> set as \<or> p' \<in> set bs" by simp
       moreover
       {
-	assume p'elem: "p' \<in> set as"
-	with ascarr have [simp]: "p' \<in> carrier G" by fast
-
-	note pp'
-	also from afac p'elem
-	     have "p' divides a" by (rule factors_dividesI) fact+
-	finally
-	     have "p divides a" by simp
+        assume p'elem: "p' \<in> set as"
+        with ascarr have [simp]: "p' \<in> carrier G" by fast
+
+        note pp'
+        also from afac p'elem
+             have "p' divides a" by (rule factors_dividesI) fact+
+        finally
+             have "p divides a" by simp
       }
       moreover
       {
-	assume p'elem: "p' \<in> set bs"
-	with bscarr have [simp]: "p' \<in> carrier G" by fast
-
-	note pp'
-	also from bfac
-	     have "p' divides b" by (rule factors_dividesI) fact+
-	finally have "p divides b" by simp
+        assume p'elem: "p' \<in> set bs"
+        with bscarr have [simp]: "p' \<in> carrier G" by fast
+
+        note pp'
+        also from bfac
+             have "p' divides b" by (rule factors_dividesI) fact+
+        finally have "p divides b" by simp
       }
       ultimately
-	  show "p divides a \<or> p divides b" by fast
+          show "p divides a \<or> p divides b" by fast
     qed
   qed
 qed
@@ -3176,7 +3176,7 @@
   have "c = c \<otimes> \<one>" by simp
   also from abrelprime[symmetric]
        have "\<dots> \<sim> c \<otimes> somegcd G a b"
-	 by (rule assoc_subst) (simp add: mult_cong_r)+
+         by (rule assoc_subst) (simp add: mult_cong_r)+
   also have "\<dots> \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by (rule gcd_mult) fact+
   finally
        have c: "c \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by simp
@@ -3188,13 +3188,13 @@
   have "somegcd G a (b \<otimes> c) \<sim> somegcd G a (c \<otimes> b)" by (simp add: m_comm)
   also from a
        have "\<dots> \<sim> somegcd G (somegcd G a (c \<otimes> a)) (c \<otimes> b)"
-	 by (rule assoc_subst) (simp add: gcd_cong_l)+
+         by (rule assoc_subst) (simp add: gcd_cong_l)+
   also from gcd_assoc
        have "\<dots> \<sim> somegcd G a (somegcd G (c \<otimes> a) (c \<otimes> b))"
        by (rule assoc_subst) simp+
   also from c[symmetric]
        have "\<dots> \<sim> somegcd G a c"
-	 by (rule assoc_subst) (simp add: gcd_cong_r)+
+         by (rule assoc_subst) (simp add: gcd_cong_r)+
   also note acrelprime
   finally
        show "somegcd G a (b \<otimes> c) \<sim> \<one>" by simp
--- a/src/HOL/Algebra/FiniteProduct.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Algebra/FiniteProduct.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -423,17 +423,17 @@
       then have "finprod G f A = finprod G f (insert x B)" by simp
       also from insert have "... = f x \<otimes> finprod G f B"
       proof (intro finprod_insert)
-	show "finite B" by fact
+        show "finite B" by fact
       next
-	show "x ~: B" by fact
+        show "x ~: B" by fact
       next
-	assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
-	  "g \<in> insert x B \<rightarrow> carrier G"
-	thus "f \<in> B -> carrier G" by fastsimp
+        assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
+          "g \<in> insert x B \<rightarrow> carrier G"
+        thus "f \<in> B -> carrier G" by fastsimp
       next
-	assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
-	  "g \<in> insert x B \<rightarrow> carrier G"
-	thus "f x \<in> carrier G" by fastsimp
+        assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
+          "g \<in> insert x B \<rightarrow> carrier G"
+        thus "f x \<in> carrier G" by fastsimp
       qed
       also from insert have "... = g x \<otimes> finprod G g B" by fastsimp
       also from insert have "... = finprod G g (insert x B)"
--- a/src/HOL/Algebra/Group.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Algebra/Group.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -785,16 +785,16 @@
       assume H: "H \<in> A"
       with L have subgroupH: "subgroup H G" by auto
       from subgroupH have groupH: "group (G (| carrier := H |))" (is "group ?H")
-	by (rule subgroup_imp_group)
+        by (rule subgroup_imp_group)
       from groupH have monoidH: "monoid ?H"
-	by (rule group.is_monoid)
+        by (rule group.is_monoid)
       from H have Int_subset: "?Int \<subseteq> H" by fastsimp
       then show "le ?L ?Int H" by simp
     next
       fix H
       assume H: "H \<in> Lower ?L A"
       with L Int_subgroup show "le ?L H ?Int"
-	by (fastsimp simp: Lower_def intro: Inter_greatest)
+        by (fastsimp simp: Lower_def intro: Inter_greatest)
     next
       show "A \<subseteq> carrier ?L" by (rule L)
     next
--- a/src/HOL/Algebra/Lattice.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Algebra/Lattice.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -533,22 +533,22 @@
       assume y: "y \<in> Upper L (insert x A)"
       show "s \<sqsubseteq> y"
       proof (rule least_le [OF least_s], rule Upper_memI)
-	fix z
-	assume z: "z \<in> {a, x}"
-	then show "z \<sqsubseteq> y"
-	proof
+        fix z
+        assume z: "z \<in> {a, x}"
+        then show "z \<sqsubseteq> y"
+        proof
           have y': "y \<in> Upper L A"
             apply (rule subsetD [where A = "Upper L (insert x A)"])
              apply (rule Upper_antimono)
-	     apply blast
-	    apply (rule y)
+             apply blast
+            apply (rule y)
             done
           assume "z = a"
           with y' least_a show ?thesis by (fast dest: least_le)
-	next
-	  assume "z \<in> {x}"  (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
+        next
+          assume "z \<in> {x}"  (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
           with y L show ?thesis by blast
-	qed
+        qed
       qed (rule Upper_closed [THEN subsetD, OF y])
     next
       from L show "insert x A \<subseteq> carrier L" by simp
@@ -569,9 +569,9 @@
     case True
     with insert show ?thesis
       by simp (simp add: least_cong [OF weak_sup_of_singleton]
-	sup_of_singleton_closed sup_of_singletonI)
-	(* The above step is hairy; least_cong can make simp loop.
-	Would want special version of simp to apply least_cong. *)
+        sup_of_singleton_closed sup_of_singletonI)
+        (* The above step is hairy; least_cong can make simp loop.
+        Would want special version of simp to apply least_cong. *)
   next
     case False
     with insert have "least L (\<Squnion>A) (Upper L A)" by simp
@@ -774,22 +774,22 @@
       assume y: "y \<in> Lower L (insert x A)"
       show "y \<sqsubseteq> i"
       proof (rule greatest_le [OF greatest_i], rule Lower_memI)
-	fix z
-	assume z: "z \<in> {a, x}"
-	then show "y \<sqsubseteq> z"
-	proof
+        fix z
+        assume z: "z \<in> {a, x}"
+        then show "y \<sqsubseteq> z"
+        proof
           have y': "y \<in> Lower L A"
             apply (rule subsetD [where A = "Lower L (insert x A)"])
             apply (rule Lower_antimono)
-	     apply blast
-	    apply (rule y)
+             apply blast
+            apply (rule y)
             done
           assume "z = a"
           with y' greatest_a show ?thesis by (fast dest: greatest_le)
-	next
+        next
           assume "z \<in> {x}"
           with y L show ?thesis by blast
-	qed
+        qed
       qed (rule Lower_closed [THEN subsetD, OF y])
     next
       from L show "insert x A \<subseteq> carrier L" by simp
@@ -809,7 +809,7 @@
     case True
     with insert show ?thesis
       by simp (simp add: greatest_cong [OF weak_inf_of_singleton]
-	inf_of_singleton_closed inf_of_singletonI)
+        inf_of_singleton_closed inf_of_singletonI)
   next
     case False
     from insert show ?thesis
@@ -1291,7 +1291,7 @@
   proof
     from B show "greatest ?L (\<Inter> B \<inter> A) (Lower ?L B)"
       txt {* @{term "\<Inter> B"} is not the infimum of @{term B}:
-	@{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
+        @{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
       by (fastsimp intro!: greatest_LowerI simp: Lower_def)
   qed
 qed
--- a/src/HOL/Algebra/UnivPoly.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Algebra/UnivPoly.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -349,19 +349,19 @@
       fix nn assume Succ: "n = Suc nn"
       have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = coeff P p (Suc nn)"
       proof -
-	have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = (\<Oplus>i\<in>{..Suc nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" using R by simp
-	also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>) \<oplus> (\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))"
-	  using finsum_Suc [of "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "nn"] unfolding Pi_def using R by simp
-	also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"
-	proof -
-	  have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
-	    using finsum_cong [of "{..nn}" "{..nn}" "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "(\<lambda>i::nat. \<zero>)"] using R 
-	    unfolding Pi_def by simp
-	  also have "\<dots> = \<zero>" by simp
-	  finally show ?thesis using r_zero R by simp
-	qed
-	also have "\<dots> = coeff P p (Suc nn)" using R by simp
-	finally show ?thesis by simp
+        have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = (\<Oplus>i\<in>{..Suc nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" using R by simp
+        also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>) \<oplus> (\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))"
+          using finsum_Suc [of "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "nn"] unfolding Pi_def using R by simp
+        also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"
+        proof -
+          have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
+            using finsum_cong [of "{..nn}" "{..nn}" "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "(\<lambda>i::nat. \<zero>)"] using R 
+            unfolding Pi_def by simp
+          also have "\<dots> = \<zero>" by simp
+          finally show ?thesis using r_zero R by simp
+        qed
+        also have "\<dots> = coeff P p (Suc nn)" using R by simp
+        finally show ?thesis by simp
       qed
       then show ?thesis using Succ by simp
     }
@@ -627,11 +627,11 @@
     then show "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k)"
     proof -
       have lhs: "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
-	unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..
+        unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..
       note cl = monom_closed [OF R.one_closed, of 1]
       note clk = monom_closed [OF R.one_closed, of k]
       have rhs: "monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
-	unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc  [OF cl clk cl]] ..
+        unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc  [OF cl clk cl]] ..
       from lhs rhs show ?thesis by simp
     qed
   }
@@ -670,25 +670,25 @@
     case True 
     {
       show ?thesis
-	unfolding True [symmetric]
-	  coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"] 
-	  coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]
-	using R.finsum_cong [of "{.. n + m}" "{.. n + m}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = n + m - i then b else \<zero>))" 
-	  "(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]
-	  a_in_R b_in_R
-	unfolding simp_implies_def
-	using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
-	unfolding Pi_def by auto
+        unfolding True [symmetric]
+          coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"] 
+          coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]
+        using R.finsum_cong [of "{.. n + m}" "{.. n + m}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = n + m - i then b else \<zero>))" 
+          "(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]
+          a_in_R b_in_R
+        unfolding simp_implies_def
+        using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
+        unfolding Pi_def by auto
     }
   next
     case False
     {
       show ?thesis
-	unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)
-	unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]
-	unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False
-	using R.finsum_cong [of "{..k}" "{..k}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = k - i then b else \<zero>))" "(\<lambda>i. \<zero>)"]
-	unfolding Pi_def simp_implies_def using a_in_R b_in_R by force
+        unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)
+        unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]
+        unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False
+        using R.finsum_cong [of "{..k}" "{..k}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = k - i then b else \<zero>))" "(\<lambda>i. \<zero>)"]
+        unfolding Pi_def simp_implies_def using a_in_R b_in_R by force
     }
   qed
 qed (simp_all add: a_in_R b_in_R)
@@ -1517,7 +1517,7 @@
       then have max_sl: "max (deg R p) (deg R q) < m" by simp
       then have "deg R (p \<oplus>\<^bsub>P\<^esub> q) < m" using deg_add [OF p_in_P q_in_P] by arith
       with deg_R_p deg_R_q show ?thesis using coeff_add [OF p_in_P q_in_P, of m]
-	using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp 
+        using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp 
     qed
   qed (simp add: p_in_P q_in_P)
   moreover have deg_ne: "deg R (p \<oplus>\<^bsub>P\<^esub> q) \<noteq> deg R r"
@@ -1582,114 +1582,114 @@
       (*JE: we now apply the induction hypothesis with some additional facts required*)
       from f_in_P deg_g_le_deg_f show ?thesis
       proof (induct n \<equiv> "deg R f" arbitrary: "f" rule: nat_less_induct)
-	fix n f
-	assume hypo: "\<forall>m<n. \<forall>x. x \<in> carrier P \<longrightarrow>
+        fix n f
+        assume hypo: "\<forall>m<n. \<forall>x. x \<in> carrier P \<longrightarrow>
           deg R g \<le> deg R x \<longrightarrow> 
-	  m = deg R x \<longrightarrow>
-	  (\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> x = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"
-	  and prem: "n = deg R f" and f_in_P [simp]: "f \<in> carrier P"
-	  and deg_g_le_deg_f: "deg R g \<le> deg R f"
-	let ?k = "1::nat" and ?r = "(g \<otimes>\<^bsub>P\<^esub> (monom P (lcoeff f) (deg R f - deg R g))) \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)"
-	  and ?q = "monom P (lcoeff f) (deg R f - deg R g)"
-	show "\<exists> q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"
-	proof -
-	  (*JE: we first extablish the existence of a triple satisfying the previous equation. 
-	    Then we will have to prove the second part of the predicate.*)
-	  have exist: "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r"
-	    using minus_add
-	    using sym [OF a_assoc [of "g \<otimes>\<^bsub>P\<^esub> ?q" "\<ominus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "lcoeff g \<odot>\<^bsub>P\<^esub> f"]]
-	    using r_neg by auto
-	  show ?thesis
-	  proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")
-	    (*JE: if the degree of the remainder satisfies the statement property we are done*)
-	    case True
-	    {
-	      show ?thesis
-	      proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)
-		show "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r" using exist by simp
-		show "\<ominus>\<^bsub>P\<^esub> ?r = \<zero>\<^bsub>P\<^esub> \<or> deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g" using True by simp
-	      qed (simp_all)
-	    }
-	  next
-	    case False note n_deg_r_l_deg_g = False
-	    {
-	      (*JE: otherwise, we verify the conditions of the induction hypothesis.*)
-	      show ?thesis
-	      proof (cases "deg R f = 0")
-		(*JE: the solutions are different if the degree of f is zero or not*)
-		case True
-		{
-		  have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp
-		  have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"
-		    unfolding deg_g apply simp
-		    unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]
-		    using deg_zero_impl_monom [OF g_in_P deg_g] by simp
-		  then show ?thesis using f_in_P by blast
-		}
-	      next
-		case False note deg_f_nzero = False
-		{
-		  (*JE: now it only remains the case where the induction hypothesis can be used.*)
-		  (*JE: we first prove that the degree of the remainder is smaller than the one of f*)
-		  have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n"
-		  proof -
-		    have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp
-		    also have "\<dots> < deg R f"
-		    proof (rule deg_lcoeff_cancel)
-		      show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"
-			using deg_smult_ring [of "lcoeff g" f] using prem
-			using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp
-		      show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"
-			using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f
-			by simp
-		      show "coeff P (g \<otimes>\<^bsub>P\<^esub> ?q) (deg R f) = \<ominus> coeff P (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) (deg R f)"
-			unfolding coeff_mult [OF g_in_P monom_closed [OF lcoeff_closed [OF f_in_P], of "deg R f - deg R g"], of "deg R f"]
-			unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]
-			using R.finsum_cong' [of "{..deg R f}" "{..deg R f}" 
-			  "(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))" 
-			  "(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]
-			using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
-			unfolding Pi_def using deg_g_le_deg_f by force
-		    qed (simp_all add: deg_f_nzero)
-		    finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n" unfolding prem .
-		  qed
-		  moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp
-		  moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto
-		  moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp
-		    (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)
-		  ultimately obtain q' r' k'
-		    where rem_desc: "lcoeff g (^) (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?r) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"and rem_deg: "(r' = \<zero>\<^bsub>P\<^esub> \<or> deg R r' < deg R g)"
-		    and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
-		    using hypo by blast
-		      (*JE: we now prove that the new quotient, remainder and exponent can be used to get 
-		      the quotient, remainder and exponent of the long division theorem*)
-		  show ?thesis
-		  proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
-		    show "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"
-		    proof -
-		      have "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r)" 
-			using smult_assoc1 exist by simp
-		      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?r))"
-			using UP_smult_r_distr by simp
-		      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"
-			using rem_desc by simp
-		      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
-			using sym [OF a_assoc [of "lcoeff g (^) k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]
-			using q'_in_carrier r'_in_carrier by simp
-		      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
-			using q'_in_carrier by (auto simp add: m_comm)
-		      also have "\<dots> = (((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'" 
-			using smult_assoc2 q'_in_carrier by auto
-		      also have "\<dots> = ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
-			using sym [OF l_distr] and q'_in_carrier by auto
-		      finally show ?thesis using m_comm q'_in_carrier by auto
-		    qed
-		  qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)
-		}
-	      qed
-	    }
-	  qed
-	qed
+          m = deg R x \<longrightarrow>
+          (\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> x = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"
+          and prem: "n = deg R f" and f_in_P [simp]: "f \<in> carrier P"
+          and deg_g_le_deg_f: "deg R g \<le> deg R f"
+        let ?k = "1::nat" and ?r = "(g \<otimes>\<^bsub>P\<^esub> (monom P (lcoeff f) (deg R f - deg R g))) \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)"
+          and ?q = "monom P (lcoeff f) (deg R f - deg R g)"
+        show "\<exists> q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"
+        proof -
+          (*JE: we first extablish the existence of a triple satisfying the previous equation. 
+            Then we will have to prove the second part of the predicate.*)
+          have exist: "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r"
+            using minus_add
+            using sym [OF a_assoc [of "g \<otimes>\<^bsub>P\<^esub> ?q" "\<ominus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "lcoeff g \<odot>\<^bsub>P\<^esub> f"]]
+            using r_neg by auto
+          show ?thesis
+          proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")
+            (*JE: if the degree of the remainder satisfies the statement property we are done*)
+            case True
+            {
+              show ?thesis
+              proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)
+                show "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r" using exist by simp
+                show "\<ominus>\<^bsub>P\<^esub> ?r = \<zero>\<^bsub>P\<^esub> \<or> deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g" using True by simp
+              qed (simp_all)
+            }
+          next
+            case False note n_deg_r_l_deg_g = False
+            {
+              (*JE: otherwise, we verify the conditions of the induction hypothesis.*)
+              show ?thesis
+              proof (cases "deg R f = 0")
+                (*JE: the solutions are different if the degree of f is zero or not*)
+                case True
+                {
+                  have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp
+                  have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"
+                    unfolding deg_g apply simp
+                    unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]
+                    using deg_zero_impl_monom [OF g_in_P deg_g] by simp
+                  then show ?thesis using f_in_P by blast
+                }
+              next
+                case False note deg_f_nzero = False
+                {
+                  (*JE: now it only remains the case where the induction hypothesis can be used.*)
+                  (*JE: we first prove that the degree of the remainder is smaller than the one of f*)
+                  have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n"
+                  proof -
+                    have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp
+                    also have "\<dots> < deg R f"
+                    proof (rule deg_lcoeff_cancel)
+                      show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"
+                        using deg_smult_ring [of "lcoeff g" f] using prem
+                        using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp
+                      show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"
+                        using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f
+                        by simp
+                      show "coeff P (g \<otimes>\<^bsub>P\<^esub> ?q) (deg R f) = \<ominus> coeff P (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) (deg R f)"
+                        unfolding coeff_mult [OF g_in_P monom_closed [OF lcoeff_closed [OF f_in_P], of "deg R f - deg R g"], of "deg R f"]
+                        unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]
+                        using R.finsum_cong' [of "{..deg R f}" "{..deg R f}" 
+                          "(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))" 
+                          "(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]
+                        using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
+                        unfolding Pi_def using deg_g_le_deg_f by force
+                    qed (simp_all add: deg_f_nzero)
+                    finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n" unfolding prem .
+                  qed
+                  moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp
+                  moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto
+                  moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp
+                    (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)
+                  ultimately obtain q' r' k'
+                    where rem_desc: "lcoeff g (^) (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?r) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"and rem_deg: "(r' = \<zero>\<^bsub>P\<^esub> \<or> deg R r' < deg R g)"
+                    and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
+                    using hypo by blast
+                      (*JE: we now prove that the new quotient, remainder and exponent can be used to get 
+                      the quotient, remainder and exponent of the long division theorem*)
+                  show ?thesis
+                  proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
+                    show "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"
+                    proof -
+                      have "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r)" 
+                        using smult_assoc1 exist by simp
+                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?r))"
+                        using UP_smult_r_distr by simp
+                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"
+                        using rem_desc by simp
+                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
+                        using sym [OF a_assoc [of "lcoeff g (^) k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]
+                        using q'_in_carrier r'_in_carrier by simp
+                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
+                        using q'_in_carrier by (auto simp add: m_comm)
+                      also have "\<dots> = (((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'" 
+                        using smult_assoc2 q'_in_carrier by auto
+                      also have "\<dots> = ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
+                        using sym [OF l_distr] and q'_in_carrier by auto
+                      finally show ?thesis using m_comm q'_in_carrier by auto
+                    qed
+                  qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)
+                }
+              qed
+            }
+          qed
+        qed
       qed
     }
   qed
--- a/src/HOL/Algebra/poly/UnivPoly2.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Algebra/poly/UnivPoly2.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -100,7 +100,7 @@
 begin
 
 definition
-  up_add_def:	"p + q = Abs_UP (%n. Rep_UP p n + Rep_UP q n)"
+  up_add_def: "p + q = Abs_UP (%n. Rep_UP p n + Rep_UP q n)"
 
 instance ..
 
@@ -133,7 +133,7 @@
 
 definition
   up_mult_def:  "p * q = Abs_UP (%n::nat. setsum
-		     (%i. Rep_UP p i * Rep_UP q (n-i)) {..n})"
+                     (%i. Rep_UP p i * Rep_UP q (n-i)) {..n})"
 
 instance ..
 
@@ -201,18 +201,18 @@
     have "(%i. f i + g i) : UP"
     proof -
       from fup obtain n where boundn: "bound n f"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
       from gup obtain m where boundm: "bound m g"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
       have "bound (max n m) (%i. (f i + g i))"
       proof
-	fix i
-	assume "max n m < i"
-	with boundn and boundm show "f i + g i = 0"
+        fix i
+        assume "max n m < i"
+        with boundn and boundm show "f i + g i = 0"
           by (fastsimp simp add: algebra_simps)
       qed
       then show "(%i. (f i + g i)) : UP"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
     qed
   }
   then show ?thesis
@@ -228,30 +228,30 @@
     have "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
     proof -
       from fup obtain n where "bound n f"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
       from gup obtain m where "bound m g"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
       have "bound (n + m) (%n. setsum (%i. f i * g (n-i)) {..n})"
       proof
-	fix k
-	assume bound: "n + m < k"
-	{
-	  fix i
-	  have "f i * g (k-i) = 0"
-	  proof cases
-	    assume "n < i"
-	    with `bound n f` show ?thesis by (auto simp add: algebra_simps)
-	  next
-	    assume "~ (n < i)"
-	    with bound have "m < k-i" by arith
-	    with `bound m g` show ?thesis by (auto simp add: algebra_simps)
-	  qed
-	}
-	then show "setsum (%i. f i * g (k-i)) {..k} = 0"
-	  by (simp add: algebra_simps)
+        fix k
+        assume bound: "n + m < k"
+        {
+          fix i
+          have "f i * g (k-i) = 0"
+          proof cases
+            assume "n < i"
+            with `bound n f` show ?thesis by (auto simp add: algebra_simps)
+          next
+            assume "~ (n < i)"
+            with bound have "m < k-i" by arith
+            with `bound m g` show ?thesis by (auto simp add: algebra_simps)
+          qed
+        }
+        then show "setsum (%i. f i * g (k-i)) {..k} = 0"
+          by (simp add: algebra_simps)
       qed
       then show "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
     qed
   }
   then show ?thesis
@@ -290,17 +290,17 @@
     {
       fix k and a b c :: "nat=>'a::ring"
       have "k <= n ==> 
-	setsum (%j. setsum (%i. a i * b (j-i)) {..j} * c (n-j)) {..k} = 
-	setsum (%j. a j * setsum  (%i. b i * c (n-j-i)) {..k-j}) {..k}"
-	(is "_ ==> ?eq k")
+        setsum (%j. setsum (%i. a i * b (j-i)) {..j} * c (n-j)) {..k} = 
+        setsum (%j. a j * setsum  (%i. b i * c (n-j-i)) {..k-j}) {..k}"
+        (is "_ ==> ?eq k")
       proof (induct k)
-	case 0 show ?case by simp
+        case 0 show ?case by simp
       next
-	case (Suc k)
-	then have "k <= n" by arith
-	then have "?eq k" by (rule Suc)
-	then show ?case
-	  by (simp add: Suc_diff_le natsum_ldistr)
+        case (Suc k)
+        then have "k <= n" by arith
+        then have "?eq k" by (rule Suc)
+        then show ?case
+          by (simp add: Suc_diff_le natsum_ldistr)
       qed
     }
     then show "coeff ((p * q) * r) n = coeff (p * (q * r)) n"
@@ -326,13 +326,13 @@
       fix k
       fix a b :: "nat=>'a::ring"
       have "k <= n ==> 
-	setsum (%i. a i * b (n-i)) {..k} =
-	setsum (%i. a (k-i) * b (i+n-k)) {..k}"
-	(is "_ ==> ?eq k")
+        setsum (%i. a i * b (n-i)) {..k} =
+        setsum (%i. a (k-i) * b (i+n-k)) {..k}"
+        (is "_ ==> ?eq k")
       proof (induct k)
-	case 0 show ?case by simp
+        case 0 show ?case by simp
       next
-	case (Suc k) then show ?case by (subst natsum_Suc2) simp
+        case (Suc k) then show ?case by (subst natsum_Suc2) simp
       qed
     }
     then show "coeff (p * q) n = coeff (q * p) n"
--- a/src/HOL/Auth/CertifiedEmail.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/CertifiedEmail.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -43,7 +43,7 @@
 
 | FakeSSL: --{*The Spy may open SSL sessions with TTP, who is the only agent
     equipped with the necessary credentials to serve as an SSL server.*}
-	 "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
+         "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
           ==> Notes TTP {|Agent Spy, Agent TTP, X|} # evsfssl \<in> certified_mail"
 
 | CM1: --{*The sender approaches the recipient.  The message is a number.*}
@@ -54,14 +54,14 @@
     hs = Hash{|Number cleartext, Nonce q, response S R q, Crypt K (Number m)|};
     S2TTP = Crypt(pubEK TTP) {|Agent S, Number BothAuth, Key K, Agent R, hs|}|]
   ==> Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number BothAuth, 
-		 Number cleartext, Nonce q, S2TTP|} # evs1 
-	\<in> certified_mail"
+                 Number cleartext, Nonce q, S2TTP|} # evs1 
+        \<in> certified_mail"
 
 | CM2: --{*The recipient records @{term S2TTP} while transmitting it and her
      password to @{term TTP} over an SSL channel.*}
  "[|evs2 \<in> certified_mail;
     Gets R {|Agent S, Agent TTP, em, Number BothAuth, Number cleartext, 
-	     Nonce q, S2TTP|} \<in> set evs2;
+             Nonce q, S2TTP|} \<in> set evs2;
     TTP \<noteq> R;  
     hr = Hash {|Number cleartext, Nonce q, response S R q, em|} |]
   ==> 
@@ -289,14 +289,14 @@
  "evs \<in> certified_mail ==>
     Key K \<notin> analz (spies evs) -->
     (\<forall>AO. Crypt (pubEK TTP)
-	   {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
+           {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
     (\<exists>m ctxt q. 
         hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
-	Says S R
-	   {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
-	     Number ctxt, Nonce q,
-	     Crypt (pubEK TTP)
-	      {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))" 
+        Says S R
+           {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
+             Number ctxt, Nonce q,
+             Crypt (pubEK TTP)
+              {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))" 
 apply (erule certified_mail.induct, analz_mono_contra)
 apply (drule_tac [5] CM2_S2TTP_parts_knows_Spy, simp)
 apply (simp add: used_Nil Crypt_notin_initState, simp_all)
@@ -322,11 +322,11 @@
     evs \<in> certified_mail|]
   ==> \<exists>m ctxt q. 
         hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
-	Says S R
-	   {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
-	     Number ctxt, Nonce q,
-	     Crypt (pubEK TTP)
-	      {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs" 
+        Says S R
+           {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
+             Number ctxt, Nonce q,
+             Crypt (pubEK TTP)
+              {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs" 
 by (blast intro: S2TTP_sender_lemma) 
 
 
@@ -401,7 +401,7 @@
                      Number cleartext, Nonce q, S2TTP|} \<in> set evs;
          S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
          Key K \<in> analz (spies evs);
-	 evs \<in> certified_mail;
+         evs \<in> certified_mail;
          S\<noteq>Spy|]
       ==> R \<in> bad & Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
 apply (erule rev_mp)
@@ -428,7 +428,7 @@
       "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
                      Number cleartext, Nonce q, S2TTP|} \<in> set evs;
          S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
-	 evs \<in> certified_mail;
+         evs \<in> certified_mail;
          S\<noteq>Spy; R \<notin> bad|]
       ==> Key K \<notin> analz(spies evs)"
 by (blast dest: S_fairness_bad_R) 
@@ -438,10 +438,10 @@
  until @{term S} has access to the return receipt.*} 
 theorem S_guarantee:
      "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
-		    Number cleartext, Nonce q, S2TTP|} \<in> set evs;
-	S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
-	Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
-	S\<noteq>Spy;  evs \<in> certified_mail|]
+                    Number cleartext, Nonce q, S2TTP|} \<in> set evs;
+        S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
+        Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
+        S\<noteq>Spy;  evs \<in> certified_mail|]
      ==> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
 apply (erule rev_mp)
 apply (erule ssubst)
--- a/src/HOL/Auth/Event.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/Event.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -22,7 +22,7 @@
         | Notes agent       msg
        
 consts 
-  bad    :: "agent set"				(*compromised agents*)
+  bad    :: "agent set"                         -- {* compromised agents *}
   knows  :: "agent => event list => msg set"
 
 
@@ -43,19 +43,19 @@
   knows_Cons:
     "knows A (ev # evs) =
        (if A = Spy then 
-	(case ev of
-	   Says A' B X => insert X (knows Spy evs)
-	 | Gets A' X => knows Spy evs
-	 | Notes A' X  => 
-	     if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
-	else
-	(case ev of
-	   Says A' B X => 
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Gets A' X    => 
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Notes A' X    => 
-	     if A'=A then insert X (knows A evs) else knows A evs))"
+        (case ev of
+           Says A' B X => insert X (knows Spy evs)
+         | Gets A' X => knows Spy evs
+         | Notes A' X  => 
+             if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
+        else
+        (case ev of
+           Says A' B X => 
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Gets A' X    => 
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Notes A' X    => 
+             if A'=A then insert X (knows A evs) else knows A evs))"
 
 (*
   Case A=Spy on the Gets event
@@ -71,10 +71,10 @@
 primrec
   used_Nil:   "used []         = (UN B. parts (initState B))"
   used_Cons:  "used (ev # evs) =
-		     (case ev of
-			Says A B X => parts {X} \<union> used evs
-		      | Gets A X   => used evs
-		      | Notes A X  => parts {X} \<union> used evs)"
+                     (case ev of
+                        Says A B X => parts {X} \<union> used evs
+                      | Gets A X   => used evs
+                      | Notes A X  => parts {X} \<union> used evs)"
     --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
         follows @{term Says} in real protocols.  Seems difficult to change.
         See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
--- a/src/HOL/Auth/KerberosIV.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/KerberosIV.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -170,18 +170,18 @@
             B \<noteq> Tgs;  authK \<in> symKeys;
             Says A' Tgs \<lbrace>
              (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
-				 Number Ta\<rbrace>),
+                                 Number Ta\<rbrace>),
              (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
-	        \<in> set evs4;
+                \<in> set evs4;
             \<not> expiredAK Ta evs4;
             \<not> expiredA T2 evs4;
             servKlife + (CT evs4) <= authKlife + Ta
          \<rbrakk>
           \<Longrightarrow> Says Tgs A
                 (Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
-			       Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
-		 			        Number (CT evs4)\<rbrace> \<rbrace>)
-	        # evs4 \<in> kerbIV"
+                               Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
+                                                Number (CT evs4)\<rbrace> \<rbrace>)
+                # evs4 \<in> kerbIV"
 (* Tgs creates a new session key per each request for a service, without
    checking if there is still a fresh one for that service.
    The cipher under Tgs' key is the authTicket, the cipher under B's key
@@ -196,14 +196,14 @@
  | K5:  "\<lbrakk> evs5 \<in> kerbIV; authK \<in> symKeys; servK \<in> symKeys;
             Says A Tgs
                 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
-		  Agent B\<rbrace>
+                  Agent B\<rbrace>
               \<in> set evs5;
             Says Tgs' A
              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
                 \<in> set evs5;
             valid Ts wrt T2 \<rbrakk>
           \<Longrightarrow> Says A B \<lbrace>servTicket,
-			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+                         Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
                # evs5 \<in> kerbIV"
 (* Checks similar to those in K3. *)
 
@@ -609,7 +609,7 @@
      evs \<in> kerbIV \<rbrakk>
   \<Longrightarrow> servK \<notin> range shrK &
       (\<exists>A. servTicket =
-	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
+              Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
        | servTicket \<in> analz (spies evs)"
 apply (frule Says_imp_spies [THEN analz.Inj], auto)
  apply (force dest!: servTicket_form)
@@ -1336,15 +1336,15 @@
 
 lemma Confidentiality_lemma [rule_format]:
      "\<lbrakk> Says Tgs A
-	    (Crypt authK
-	       \<lbrace>Key servK, Agent B, Number Ts,
-		 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
-	   \<in> set evs;
-	Key authK \<notin> analz (spies evs);
+            (Crypt authK
+               \<lbrace>Key servK, Agent B, Number Ts,
+                 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+           \<in> set evs;
+        Key authK \<notin> analz (spies evs);
         servK \<in> symKeys;
-	A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+        A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
       \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
-	  expiredSK Ts evs"
+          expiredSK Ts evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule kerbIV.induct)
--- a/src/HOL/Auth/KerberosIV_Gets.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/KerberosIV_Gets.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -161,18 +161,18 @@
             B \<noteq> Tgs;  authK \<in> symKeys;
             Gets Tgs \<lbrace>
              (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
-				 Number Ta\<rbrace>),
+                                 Number Ta\<rbrace>),
              (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
-	        \<in> set evs4;
+                \<in> set evs4;
             \<not> expiredAK Ta evs4;
             \<not> expiredA T2 evs4;
             servKlife + (CT evs4) <= authKlife + Ta
          \<rbrakk>
           \<Longrightarrow> Says Tgs A
                 (Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
-			       Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
-		 			        Number (CT evs4)\<rbrace> \<rbrace>)
-	        # evs4 \<in> kerbIV_gets"
+                               Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
+                                                Number (CT evs4)\<rbrace> \<rbrace>)
+                # evs4 \<in> kerbIV_gets"
 (* Tgs creates a new session key per each request for a service, without
    checking if there is still a fresh one for that service.
    The cipher under Tgs' key is the authTicket, the cipher under B's key
@@ -187,14 +187,14 @@
  | K5:  "\<lbrakk> evs5 \<in> kerbIV_gets; authK \<in> symKeys; servK \<in> symKeys;
             Says A Tgs
                 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
-		  Agent B\<rbrace>
+                  Agent B\<rbrace>
               \<in> set evs5;
             Gets A
              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
                 \<in> set evs5;
             valid Ts wrt T2 \<rbrakk>
           \<Longrightarrow> Says A B \<lbrace>servTicket,
-			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+                         Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
                # evs5 \<in> kerbIV_gets"
 (* Checks similar to those in K3. *)
 
@@ -495,7 +495,7 @@
      evs \<in> kerbIV_gets \<rbrakk>
   \<Longrightarrow> servK \<notin> range shrK &
       (\<exists>A. servTicket =
-	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
+              Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
        | servTicket \<in> analz (spies evs)"
 apply (frule Gets_imp_knows_Spy [THEN analz.Inj], auto)
  apply (force dest!: servTicket_form)
@@ -1231,15 +1231,15 @@
 
 lemma Confidentiality_lemma [rule_format]:
      "\<lbrakk> Says Tgs A
-	    (Crypt authK
-	       \<lbrace>Key servK, Agent B, Number Ts,
-		 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
-	   \<in> set evs;
-	Key authK \<notin> analz (spies evs);
+            (Crypt authK
+               \<lbrace>Key servK, Agent B, Number Ts,
+                 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+           \<in> set evs;
+        Key authK \<notin> analz (spies evs);
         servK \<in> symKeys;
-	A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+        A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
       \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
-	  expiredSK Ts evs"
+          expiredSK Ts evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule kerbIV_gets.induct)
--- a/src/HOL/Auth/KerberosV.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/KerberosV.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -137,9 +137,9 @@
             B \<noteq> Tgs;  authK \<in> symKeys;
             Says A' Tgs \<lbrace>
              (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
-				 Number Ta\<rbrace>),
+                                 Number Ta\<rbrace>),
              (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
-	        \<in> set evs4;
+                \<in> set evs4;
             \<not> expiredAK Ta evs4;
             \<not> expiredA T2 evs4;
             servKlife + (CT evs4) <= authKlife + Ta
@@ -155,14 +155,14 @@
             A \<noteq> Kas; A \<noteq> Tgs;
             Says A Tgs
                 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
-		  Agent B\<rbrace>
+                  Agent B\<rbrace>
               \<in> set evs5;
             Says Tgs' A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
                           servTicket\<rbrace>
                 \<in> set evs5;
             valid Ts wrt T2 \<rbrakk>
           \<Longrightarrow> Says A B \<lbrace>servTicket,
-			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+                         Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
                # evs5 \<in> kerbV"
 
   | KV6:  "\<lbrakk> evs6 \<in> kerbV; B \<noteq> Kas; B \<noteq> Tgs;
@@ -1081,14 +1081,14 @@
 
 lemma Confidentiality_lemma [rule_format]:
      "\<lbrakk> Says Tgs A
-	    \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
-	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
-	   \<in> set evs;
-	Key authK \<notin> analz (spies evs);
+            \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+              Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
+           \<in> set evs;
+        Key authK \<notin> analz (spies evs);
         servK \<in> symKeys;
-	A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
+        A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
       \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
-	  expiredSK Ts evs"
+          expiredSK Ts evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule kerbV.induct)
--- a/src/HOL/Auth/Kerberos_BAN.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/Kerberos_BAN.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -77,45 +77,45 @@
    Nil:  "[] \<in> bankerberos"
 
  | Fake: "\<lbrakk> evsf \<in> bankerberos;  X \<in> synth (analz (spies evsf)) \<rbrakk>
-	  \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
+          \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
 
 
  | BK1:  "\<lbrakk> evs1 \<in> bankerberos \<rbrakk>
-	  \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
-		\<in>  bankerberos"
+          \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
+                \<in>  bankerberos"
 
 
  | BK2:  "\<lbrakk> evs2 \<in> bankerberos;  Key K \<notin> used evs2; K \<in> symKeys;
-	     Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
-	  \<Longrightarrow> Says Server A
-		(Crypt (shrK A)
-		   \<lbrace>Number (CT evs2), Agent B, Key K,
-		    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
-		# evs2 \<in> bankerberos"
+             Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+          \<Longrightarrow> Says Server A
+                (Crypt (shrK A)
+                   \<lbrace>Number (CT evs2), Agent B, Key K,
+                    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
+                # evs2 \<in> bankerberos"
 
 
  | BK3:  "\<lbrakk> evs3 \<in> bankerberos;
-	     Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
-	       \<in> set evs3;
-	     Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
-	     \<not> expiredK Tk evs3 \<rbrakk>
-	  \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
-	       # evs3 \<in> bankerberos"
+             Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+               \<in> set evs3;
+             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+             \<not> expiredK Tk evs3 \<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
+               # evs3 \<in> bankerberos"
 
 
  | BK4:  "\<lbrakk> evs4 \<in> bankerberos;
-	     Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
-			 (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
-	     \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
-	  \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
-		\<in> bankerberos"
+             Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
+                         (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
+             \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
+          \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
+                \<in> bankerberos"
 
-	(*Old session keys may become compromised*)
+        (*Old session keys may become compromised*)
  | Oops: "\<lbrakk> evso \<in> bankerberos;
          Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
-	       \<in> set evso;
-	     expiredK Tk evso \<rbrakk>
-	  \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
+               \<in> set evso;
+             expiredK Tk evso \<rbrakk>
+          \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
 
 
 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
--- a/src/HOL/Auth/Kerberos_BAN_Gets.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -69,47 +69,47 @@
    Nil:  "[] \<in> bankerb_gets"
 
  | Fake: "\<lbrakk> evsf \<in> bankerb_gets;  X \<in> synth (analz (knows Spy evsf)) \<rbrakk>
-	  \<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
+          \<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
 
  | Reception: "\<lbrakk> evsr\<in> bankerb_gets; Says A B X \<in> set evsr \<rbrakk>
                 \<Longrightarrow> Gets B X # evsr \<in> bankerb_gets"
 
  | BK1:  "\<lbrakk> evs1 \<in> bankerb_gets \<rbrakk>
-	  \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
-		\<in>  bankerb_gets"
+          \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
+                \<in>  bankerb_gets"
 
 
  | BK2:  "\<lbrakk> evs2 \<in> bankerb_gets;  Key K \<notin> used evs2; K \<in> symKeys;
-	     Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
-	  \<Longrightarrow> Says Server A
-		(Crypt (shrK A)
-		   \<lbrace>Number (CT evs2), Agent B, Key K,
-		    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
-		# evs2 \<in> bankerb_gets"
+             Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+          \<Longrightarrow> Says Server A
+                (Crypt (shrK A)
+                   \<lbrace>Number (CT evs2), Agent B, Key K,
+                    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
+                # evs2 \<in> bankerb_gets"
 
 
  | BK3:  "\<lbrakk> evs3 \<in> bankerb_gets;
-	     Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
-	       \<in> set evs3;
-	     Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
-	     \<not> expiredK Tk evs3 \<rbrakk>
-	  \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
-	       # evs3 \<in> bankerb_gets"
+             Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+               \<in> set evs3;
+             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+             \<not> expiredK Tk evs3 \<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
+               # evs3 \<in> bankerb_gets"
 
 
  | BK4:  "\<lbrakk> evs4 \<in> bankerb_gets;
-	     Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
-			 (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
-	     \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
-	  \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
-		\<in> bankerb_gets"
+             Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
+                         (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
+             \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
+          \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
+                \<in> bankerb_gets"
 
-	(*Old session keys may become compromised*)
+        (*Old session keys may become compromised*)
  | Oops: "\<lbrakk> evso \<in> bankerb_gets;
          Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
-	       \<in> set evso;
-	     expiredK Tk evso \<rbrakk>
-	  \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"
+               \<in> set evso;
+             expiredK Tk evso \<rbrakk>
+          \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"
 
 
 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
@@ -359,7 +359,7 @@
    "\<lbrakk> Says B A (Crypt K (Number Ta)) \<in> set evs;
       B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>
   \<Longrightarrow> \<exists> Tk. Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
-	            Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+                    Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
 apply (erule rev_mp)
 apply (erule bankerb_gets.induct)
 apply auto
--- a/src/HOL/Auth/Message.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/Message.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -40,13 +40,13 @@
   agent = Server | Friend nat | Spy
 
 datatype
-     msg = Agent  agent	    --{*Agent names*}
+     msg = Agent  agent     --{*Agent names*}
          | Number nat       --{*Ordinary integers, timestamps, ...*}
          | Nonce  nat       --{*Unguessable nonces*}
          | Key    key       --{*Crypto keys*}
-	 | Hash   msg       --{*Hashing*}
-	 | MPair  msg msg   --{*Compound messages*}
-	 | Crypt  key msg   --{*Encryption, public- or shared-key*}
+         | Hash   msg       --{*Hashing*}
+         | MPair  msg msg   --{*Compound messages*}
+         | Crypt  key msg   --{*Encryption, public- or shared-key*}
 
 
 text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*}
@@ -873,8 +873,8 @@
     (Fake_insert_simp_tac ss 1
      THEN
      IF_UNSOLVED (Blast.depth_tac
-		  (cs addIs [@{thm analz_insertI},
-				   impOfSubs @{thm analz_subset_parts}]) 4 1))
+                  (cs addIs [@{thm analz_insertI},
+                                   impOfSubs @{thm analz_subset_parts}]) 4 1))
 
 fun spy_analz_tac (cs,ss) i =
   DETERM
--- a/src/HOL/Auth/NS_Public.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/NS_Public.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -110,8 +110,8 @@
 lemma A_trusts_NS2_lemma [rule_format]: 
    "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
     \<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (spies evs) \<longrightarrow>
-	Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
-	Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
+        Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
+        Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake, NS1*)
 apply (blast dest: Spy_not_see_NA)+
@@ -129,8 +129,8 @@
 lemma B_trusts_NS1 [rule_format]:
      "evs \<in> ns_public                                         
       \<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
-	  Nonce NA \<notin> analz (spies evs) \<longrightarrow>
-	  Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
+          Nonce NA \<notin> analz (spies evs) \<longrightarrow>
+          Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake*)
 apply (blast intro!: analz_insertI)
--- a/src/HOL/Auth/NS_Public_Bad.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/NS_Public_Bad.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -116,8 +116,8 @@
 lemma A_trusts_NS2_lemma [rule_format]: 
    "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
     \<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
-	Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
-	Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
+        Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
+        Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct)
 apply (auto dest: Spy_not_see_NA unique_NA)
 done
@@ -134,8 +134,8 @@
 lemma B_trusts_NS1 [rule_format]:
      "evs \<in> ns_public                                         
       \<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
-	  Nonce NA \<notin> analz (spies evs) \<longrightarrow>
-	  Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
+          Nonce NA \<notin> analz (spies evs) \<longrightarrow>
+          Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake*)
 apply (blast intro!: analz_insertI)
--- a/src/HOL/Auth/NS_Shared.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/NS_Shared.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -27,61 +27,61 @@
 
 inductive_set ns_shared :: "event list set"
  where
-	(*Initial trace is empty*)
+        (*Initial trace is empty*)
   Nil:  "[] \<in> ns_shared"
-	(*The spy MAY say anything he CAN say.  We do not expect him to
-	  invent new nonces here, but he can also use NS1.  Common to
-	  all similar protocols.*)
+        (*The spy MAY say anything he CAN say.  We do not expect him to
+          invent new nonces here, but he can also use NS1.  Common to
+          all similar protocols.*)
 | Fake: "\<lbrakk>evsf \<in> ns_shared;  X \<in> synth (analz (spies evsf))\<rbrakk>
-	 \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
+         \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
 
-	(*Alice initiates a protocol run, requesting to talk to any B*)
+        (*Alice initiates a protocol run, requesting to talk to any B*)
 | NS1:  "\<lbrakk>evs1 \<in> ns_shared;  Nonce NA \<notin> used evs1\<rbrakk>
-	 \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1  \<in>  ns_shared"
+         \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1  \<in>  ns_shared"
 
-	(*Server's response to Alice's message.
-	  !! It may respond more than once to A's request !!
-	  Server doesn't know who the true sender is, hence the A' in
-	      the sender field.*)
+        (*Server's response to Alice's message.
+          !! It may respond more than once to A's request !!
+          Server doesn't know who the true sender is, hence the A' in
+              the sender field.*)
 | NS2:  "\<lbrakk>evs2 \<in> ns_shared;  Key KAB \<notin> used evs2;  KAB \<in> symKeys;
-	  Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
-	 \<Longrightarrow> Says Server A
-	       (Crypt (shrK A)
-		  \<lbrace>Nonce NA, Agent B, Key KAB,
-		    (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>)
-	       # evs2 \<in> ns_shared"
+          Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
+         \<Longrightarrow> Says Server A
+               (Crypt (shrK A)
+                  \<lbrace>Nonce NA, Agent B, Key KAB,
+                    (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>)
+               # evs2 \<in> ns_shared"
 
-	 (*We can't assume S=Server.  Agent A "remembers" her nonce.
-	   Need A \<noteq> Server because we allow messages to self.*)
+         (*We can't assume S=Server.  Agent A "remembers" her nonce.
+           Need A \<noteq> Server because we allow messages to self.*)
 | NS3:  "\<lbrakk>evs3 \<in> ns_shared;  A \<noteq> Server;
-	  Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
-	  Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
-	 \<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
+          Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
+          Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
+         \<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
 
-	(*Bob's nonce exchange.  He does not know who the message came
-	  from, but responds to A because she is mentioned inside.*)
+        (*Bob's nonce exchange.  He does not know who the message came
+          from, but responds to A because she is mentioned inside.*)
 | NS4:  "\<lbrakk>evs4 \<in> ns_shared;  Nonce NB \<notin> used evs4;  K \<in> symKeys;
-	  Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
-	 \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
+          Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
+         \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
 
-	(*Alice responds with Nonce NB if she has seen the key before.
-	  Maybe should somehow check Nonce NA again.
-	  We do NOT send NB-1 or similar as the Spy cannot spoof such things.
-	  Letting the Spy add or subtract 1 lets him send all nonces.
-	  Instead we distinguish the messages by sending the nonce twice.*)
+        (*Alice responds with Nonce NB if she has seen the key before.
+          Maybe should somehow check Nonce NA again.
+          We do NOT send NB-1 or similar as the Spy cannot spoof such things.
+          Letting the Spy add or subtract 1 lets him send all nonces.
+          Instead we distinguish the messages by sending the nonce twice.*)
 | NS5:  "\<lbrakk>evs5 \<in> ns_shared;  K \<in> symKeys;
-	  Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
-	  Says S  A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
-	    \<in> set evs5\<rbrakk>
-	 \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared"
+          Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
+          Says S  A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
+            \<in> set evs5\<rbrakk>
+         \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared"
 
-	(*This message models possible leaks of session keys.
-	  The two Nonces identify the protocol run: the rule insists upon
-	  the true senders in order to make them accurate.*)
+        (*This message models possible leaks of session keys.
+          The two Nonces identify the protocol run: the rule insists upon
+          the true senders in order to make them accurate.*)
 | Oops: "\<lbrakk>evso \<in> ns_shared;  Says B A (Crypt K (Nonce NB)) \<in> set evso;
-	  Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
-	      \<in> set evso\<rbrakk>
-	 \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
+          Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
+              \<in> set evso\<rbrakk>
+         \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
 
 
 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
@@ -98,7 +98,7 @@
 apply (intro exI bexI)
 apply (rule_tac [2] ns_shared.Nil
        [THEN ns_shared.NS1, THEN ns_shared.NS2, THEN ns_shared.NS3,
-	THEN ns_shared.NS4, THEN ns_shared.NS5])
+        THEN ns_shared.NS4, THEN ns_shared.NS5])
 apply (possibility, simp add: used_Cons)
 done
 
@@ -275,7 +275,7 @@
 apply blast
 txt{*NS3*}
 apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2
-	     dest:  Says_imp_knows_Spy analz.Inj unique_session_keys)
+             dest:  Says_imp_knows_Spy analz.Inj unique_session_keys)
 txt{*Oops*}
 apply (blast dest: unique_session_keys)
 done
@@ -355,8 +355,8 @@
   "\<lbrakk>B \<notin> bad;  evs \<in> ns_shared\<rbrakk> \<Longrightarrow>
      Key K \<notin> analz (spies evs) \<longrightarrow>
      Says Server A
-	  (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
-			    Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
+          (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
+                            Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
      Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
      Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs"
 apply (erule ns_shared.induct, force)
@@ -366,7 +366,7 @@
 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
 txt{*NS5*}
 apply (blast dest!: A_trusts_NS2
-	     dest: Says_imp_knows_Spy [THEN analz.Inj]
+             dest: Says_imp_knows_Spy [THEN analz.Inj]
                    unique_session_keys Crypt_Spy_analz_bad)
 done
 
--- a/src/HOL/Auth/OtwayRees.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/OtwayRees.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -61,7 +61,7 @@
                  # evs3 : otway"
 
          (*Bob receives the Server's (?) message and compares the Nonces with
-	   those in the message he previously sent the Server.
+           those in the message he previously sent the Server.
            Need B \<noteq> Server because we allow messages to self.*)
  | OR4:  "[| evs4 \<in> otway;  B \<noteq> Server;
              Says B Server {|Nonce NA, Agent A, Agent B, X',
--- a/src/HOL/Auth/OtwayReesBella.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/OtwayReesBella.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -22,50 +22,50 @@
   Nil:  "[]\<in> orb"
 
 | Fake: "\<lbrakk>evsa\<in> orb;  X\<in> synth (analz (knows Spy evsa))\<rbrakk>
- 	 \<Longrightarrow> Says Spy B X  # evsa \<in> orb"
+         \<Longrightarrow> Says Spy B X  # evsa \<in> orb"
 
 | Reception: "\<lbrakk>evsr\<in> orb;  Says A B X \<in> set evsr\<rbrakk>
-	      \<Longrightarrow> Gets B X # evsr \<in> orb"
+              \<Longrightarrow> Gets B X # evsr \<in> orb"
 
 | OR1:  "\<lbrakk>evs1\<in> orb;  Nonce NA \<notin> used evs1\<rbrakk>
-	 \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, 
-		   Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> 
-	       # evs1 \<in> orb"
+         \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, 
+                   Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> 
+               # evs1 \<in> orb"
 
 | OR2:  "\<lbrakk>evs2\<in> orb;  Nonce NB \<notin> used evs2;
-	   Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
-	\<Longrightarrow> Says B Server 
-		\<lbrace>Nonce M, Agent A, Agent B, X, 
-	   Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
-	       # evs2 \<in> orb"
+           Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
+        \<Longrightarrow> Says B Server 
+                \<lbrace>Nonce M, Agent A, Agent B, X, 
+           Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+               # evs2 \<in> orb"
 
 | OR3:  "\<lbrakk>evs3\<in> orb;  Key KAB \<notin> used evs3;
-	  Gets Server 
-	     \<lbrace>Nonce M, Agent A, Agent B, 
-	       Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>, 
-	       Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
-	  \<in> set evs3\<rbrakk>
-	\<Longrightarrow> Says Server B \<lbrace>Nonce M,
-		    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
-				      Nonce NB, Key KAB\<rbrace>\<rbrace>
-	       # evs3 \<in> orb"
+          Gets Server 
+             \<lbrace>Nonce M, Agent A, Agent B, 
+               Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>, 
+               Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+          \<in> set evs3\<rbrakk>
+        \<Longrightarrow> Says Server B \<lbrace>Nonce M,
+                    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
+                                      Nonce NB, Key KAB\<rbrace>\<rbrace>
+               # evs3 \<in> orb"
 
   (*B can only check that the message he is bouncing is a ciphertext*)
   (*Sending M back is omitted*)   
 | OR4:  "\<lbrakk>evs4\<in> orb; B \<noteq> Server; \<forall> p q. X \<noteq> \<lbrace>p, q\<rbrace>; 
-	  Says B Server \<lbrace>Nonce M, Agent A, Agent B, X', 
-		Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
-	    \<in> set evs4;
-	  Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
-	    \<in> set evs4\<rbrakk>
-	\<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
+          Says B Server \<lbrace>Nonce M, Agent A, Agent B, X', 
+                Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+            \<in> set evs4;
+          Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
+            \<in> set evs4\<rbrakk>
+        \<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
 
 
 | Oops: "\<lbrakk>evso\<in> orb;  
-	   Says Server B \<lbrace>Nonce M,
-		    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
-				      Nonce NB, Key KAB\<rbrace>\<rbrace> 
-	     \<in> set evso\<rbrakk>
+           Says Server B \<lbrace>Nonce M,
+                    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
+                                      Nonce NB, Key KAB\<rbrace>\<rbrace> 
+             \<in> set evso\<rbrakk>
  \<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB, Key KAB\<rbrace> # evso 
      \<in> orb"
 
--- a/src/HOL/Auth/OtwayRees_AN.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/OtwayRees_AN.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -35,7 +35,7 @@
         
  | Reception: --{*A message that has been sent can be received by the
                   intended recipient.*}
-	      "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
+              "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
                ==> Gets B X # evsr \<in> otway"
 
  | OR1:  --{*Alice initiates a protocol run*}
@@ -43,14 +43,14 @@
           ==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 \<in> otway"
 
  | OR2:  --{*Bob's response to Alice's message.*}
-	 "[| evs2 \<in> otway;
+         "[| evs2 \<in> otway;
              Gets B {|Agent A, Agent B, Nonce NA|} \<in>set evs2 |]
           ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
                  # evs2 \<in> otway"
 
  | OR3:  --{*The Server receives Bob's message.  Then he sends a new
            session key to Bob with a packet for forwarding to Alice.*}
-	 "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
+         "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
              Gets Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
                \<in>set evs3 |]
           ==> Says Server B
@@ -59,9 +59,9 @@
               # evs3 \<in> otway"
 
  | OR4:  --{*Bob receives the Server's (?) message and compares the Nonces with
-	     those in the message he previously sent the Server.
+             those in the message he previously sent the Server.
              Need @{term "B \<noteq> Server"} because we allow messages to self.*}
-	 "[| evs4 \<in> otway;  B \<noteq> Server;
+         "[| evs4 \<in> otway;  B \<noteq> Server;
              Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} \<in>set evs4;
              Gets B {|X, Crypt(shrK B){|Nonce NB,Agent A,Agent B,Key K|}|}
                \<in>set evs4 |]
@@ -69,7 +69,7 @@
 
  | Oops: --{*This message models possible leaks of session keys.  The nonces
              identify the protocol run.*}
-	 "[| evso \<in> otway;
+         "[| evso \<in> otway;
              Says Server B
                       {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|},
                         Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key K|}|}
--- a/src/HOL/Auth/OtwayRees_Bad.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/OtwayRees_Bad.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -32,18 +32,18 @@
         
  | Reception: --{*A message that has been sent can be received by the
                   intended recipient.*}
-	      "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
+              "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
                ==> Gets B X # evsr \<in> otway"
 
  | OR1:  --{*Alice initiates a protocol run*}
-	 "[| evs1 \<in> otway;  Nonce NA \<notin> used evs1 |]
+         "[| evs1 \<in> otway;  Nonce NA \<notin> used evs1 |]
           ==> Says A B {|Nonce NA, Agent A, Agent B,
                          Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |}
                  # evs1 \<in> otway"
 
  | OR2:  --{*Bob's response to Alice's message.
              This variant of the protocol does NOT encrypt NB.*}
-	 "[| evs2 \<in> otway;  Nonce NB \<notin> used evs2;
+         "[| evs2 \<in> otway;  Nonce NB \<notin> used evs2;
              Gets B {|Nonce NA, Agent A, Agent B, X|} \<in> set evs2 |]
           ==> Says B Server
                   {|Nonce NA, Agent A, Agent B, X, Nonce NB,
@@ -53,7 +53,7 @@
  | OR3:  --{*The Server receives Bob's message and checks that the three NAs
            match.  Then he sends a new session key to Bob with a packet for
            forwarding to Alice.*}
-	 "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
+         "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
              Gets Server
                   {|Nonce NA, Agent A, Agent B,
                     Crypt (shrK A) {|Nonce NA, Agent A, Agent B|},
@@ -67,9 +67,9 @@
                  # evs3 \<in> otway"
 
  | OR4:  --{*Bob receives the Server's (?) message and compares the Nonces with
-	     those in the message he previously sent the Server.
+             those in the message he previously sent the Server.
              Need @{term "B \<noteq> Server"} because we allow messages to self.*}
-	 "[| evs4 \<in> otway;  B \<noteq> Server;
+         "[| evs4 \<in> otway;  B \<noteq> Server;
              Says B Server {|Nonce NA, Agent A, Agent B, X', Nonce NB,
                              Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
                \<in> set evs4;
@@ -79,7 +79,7 @@
 
  | Oops: --{*This message models possible leaks of session keys.  The nonces
              identify the protocol run.*}
-	 "[| evso \<in> otway;
+         "[| evso \<in> otway;
              Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
                \<in> set evso |]
           ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \<in> otway"
--- a/src/HOL/Auth/Recur.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/Recur.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -101,7 +101,7 @@
      etc.
 
    Oops:  "[| evso \<in> recur;  Says Server B RB \<in> set evso;
-	      RB \<in> responses evs';  Key K \<in> parts {RB} |]
+              RB \<in> responses evs';  Key K \<in> parts {RB} |]
            ==> Notes Spy {|Key K, RB|} # evso \<in> recur"
   *)
 
@@ -140,10 +140,10 @@
 apply (rule_tac [2] 
           recur.Nil
            [THEN recur.RA1 [of _ NA], 
-	    THEN recur.RA2 [of _ NB],
-	    THEN recur.RA3 [OF _ _ respond.One 
+            THEN recur.RA2 [of _ NB],
+            THEN recur.RA3 [OF _ _ respond.One 
                                      [THEN respond.Cons [of _ _ K _ K']]],
-	    THEN recur.RA4], possibility)
+            THEN recur.RA4], possibility)
 apply (auto simp add: used_Cons)
 done
 
@@ -241,7 +241,7 @@
                    (K \<in> KK | Key K \<in> analz (insert RB H))"
 apply (erule responses.induct)
 apply (simp_all del: image_insert
-	        add: analz_image_freshK_simps, auto)
+                add: analz_image_freshK_simps, auto)
 done 
 
 
--- a/src/HOL/Auth/Smartcard/EventSC.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/Smartcard/EventSC.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -38,13 +38,13 @@
 specification (stolen)
   (*The server's card is secure by assumption\<dots>*)
   Card_Server_not_stolen [iff]: "Card Server \<notin> stolen"
-  Card_Spy_not_stolen  	 [iff]: "Card Spy \<notin> stolen"
+  Card_Spy_not_stolen    [iff]: "Card Spy \<notin> stolen"
   apply blast done
 
 specification (cloned)
   (*\<dots> the spy's card is secure because she already can use it freely*)
   Card_Server_not_cloned [iff]: "Card Server \<notin> cloned"
-  Card_Spy_not_cloned  	 [iff]: "Card Spy \<notin> cloned"
+  Card_Spy_not_cloned    [iff]: "Card Spy \<notin> cloned"
   apply blast done
 
 
@@ -52,28 +52,28 @@
           assumption of secure means*)
   knows_Nil:   "knows A [] = initState A"
   knows_Cons:  "knows A (ev # evs) =
-	 (case ev of
-	    Says A' B X => 
+         (case ev of
+            Says A' B X => 
                 if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
-	  | Notes A' X  => 
-	        if (A=A' | (A=Spy & A'\<in>bad)) then insert X (knows A evs) 
+          | Notes A' X  => 
+                if (A=A' | (A=Spy & A'\<in>bad)) then insert X (knows A evs) 
                                              else knows A evs 
           | Gets A' X   =>
-		if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
+                if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
                                      else knows A evs
           | Inputs A' C X =>
-	      if secureM then
+              if secureM then
                 if A=A' then insert X (knows A evs) else knows A evs
-	      else
-	        if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
+              else
+                if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
           | C_Gets C X   => knows A evs
           | Outpts C A' X =>
-	      if secureM then
+              if secureM then
                 if A=A' then insert X (knows A evs) else knows A evs
               else
-	        if A=Spy then insert X (knows A evs) else knows A evs
+                if A=Spy then insert X (knows A evs) else knows A evs
           | A_Gets A' X   =>
-		if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
+                if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
                                      else knows A evs)"
 
 
@@ -86,14 +86,14 @@
 primrec
   used_Nil:   "used []         = (UN B. parts (initState B))"
   used_Cons:  "used (ev # evs) =
-	         (case ev of
-		    Says A B X => parts {X} \<union> (used evs)
-		  | Notes A X  => parts {X} \<union> (used evs)
-		  | Gets A X   => used evs
+                 (case ev of
+                    Says A B X => parts {X} \<union> (used evs)
+                  | Notes A X  => parts {X} \<union> (used evs)
+                  | Gets A X   => used evs
                   | Inputs A C X  => parts{X} \<union> (used evs) 
-		  | C_Gets C X   => used evs
+                  | C_Gets C X   => used evs
                   | Outpts C A X  => parts{X} \<union> (used evs)
-		  | A_Gets A X   => used evs)"
+                  | A_Gets A X   => used evs)"
     --{*@{term Gets} always follows @{term Says} in real protocols. 
        Likewise, @{term C_Gets} will always have to follow @{term Inputs}
        and @{term A_Gets} will always have to follow @{term Outpts}*}
--- a/src/HOL/Auth/Smartcard/Smartcard.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -386,8 +386,8 @@
 
 val analz_image_freshK_ss = 
      @{simpset} delsimps [image_insert, image_Un]
-	       delsimps [@{thm imp_disjL}]    (*reduces blow-up*)
-	       addsimps @{thms analz_image_freshK_simps}
+               delsimps [@{thm imp_disjL}]    (*reduces blow-up*)
+               addsimps @{thms analz_image_freshK_simps}
 end
 *}
 
--- a/src/HOL/Auth/TLS.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/TLS.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -36,7 +36,7 @@
 Proofs would be simpler if ClientKeyExch included A's name within
 Crypt KB (Nonce PMS).  As things stand, there is much overlap between proofs
 about that message (which B receives) and the stronger event
-	Notes A {|Agent B, Nonce PMS|}.
+Notes A {|Agent B, Nonce PMS|}.
 *)
 
 header{*The TLS Protocol: Transport Layer Security*}
@@ -112,30 +112,30 @@
  | SpyKeys: --{*The spy may apply @{term PRF} and @{term sessionK}
                 to available nonces*}
          "[| evsSK \<in> tls;
-	     {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
+             {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
           ==> Notes Spy {| Nonce (PRF(M,NA,NB)),
-			   Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
+                           Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
 
  | ClientHello:
-	 --{*(7.4.1.2)
-	   PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
-	   It is uninterpreted but will be confirmed in the FINISHED messages.
-	   NA is CLIENT RANDOM, while SID is @{text SESSION_ID}.
+         --{*(7.4.1.2)
+           PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
+           It is uninterpreted but will be confirmed in the FINISHED messages.
+           NA is CLIENT RANDOM, while SID is @{text SESSION_ID}.
            UNIX TIME is omitted because the protocol doesn't use it.
            May assume @{term "NA \<notin> range PRF"} because CLIENT RANDOM is 
            28 bytes while MASTER SECRET is 48 bytes*}
          "[| evsCH \<in> tls;  Nonce NA \<notin> used evsCH;  NA \<notin> range PRF |]
           ==> Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
-	        # evsCH  \<in>  tls"
+                # evsCH  \<in>  tls"
 
  | ServerHello:
          --{*7.4.1.3 of the TLS Internet-Draft
-	   PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
+           PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
            SERVER CERTIFICATE (7.4.2) is always present.
            @{text CERTIFICATE_REQUEST} (7.4.4) is implied.*}
          "[| evsSH \<in> tls;  Nonce NB \<notin> used evsSH;  NB \<notin> range PRF;
              Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
-	       \<in> set evsSH |]
+               \<in> set evsSH |]
           ==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH  \<in>  tls"
 
  | Certificate:
@@ -148,28 +148,28 @@
            She encrypts PMS using the supplied KB, which ought to be pubK B.
            We assume @{term "PMS \<notin> range PRF"} because a clash betweem the PMS
            and another MASTER SECRET is highly unlikely (even though
-	   both items have the same length, 48 bytes).
+           both items have the same length, 48 bytes).
            The Note event records in the trace that she knows PMS
                (see REMARK at top). *}
          "[| evsCX \<in> tls;  Nonce PMS \<notin> used evsCX;  PMS \<notin> range PRF;
              Says B' A (certificate B KB) \<in> set evsCX |]
           ==> Says A B (Crypt KB (Nonce PMS))
-	      # Notes A {|Agent B, Nonce PMS|}
-	      # evsCX  \<in>  tls"
+              # Notes A {|Agent B, Nonce PMS|}
+              # evsCX  \<in>  tls"
 
  | CertVerify:
-	--{*The optional Certificate Verify (7.4.8) message contains the
+        --{*The optional Certificate Verify (7.4.8) message contains the
           specific components listed in the security analysis, F.1.1.2.
           It adds the pre-master-secret, which is also essential!
           Checking the signature, which is the only use of A's certificate,
           assures B of A's presence*}
          "[| evsCV \<in> tls;
              Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCV;
-	     Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
+             Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
           ==> Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|}))
               # evsCV  \<in>  tls"
 
-	--{*Finally come the FINISHED messages (7.4.8), confirming PA and PB
+        --{*Finally come the FINISHED messages (7.4.8), confirming PA and PB
           among other things.  The master-secret is PRF(PMS,NA,NB).
           Either party may send its message first.*}
 
@@ -181,60 +181,60 @@
           could simply put @{term "A\<noteq>Spy"} into the rule, but one should not
           expect the spy to be well-behaved.*}
          "[| evsCF \<in> tls;
-	     Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}
-	       \<in> set evsCF;
+             Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}
+               \<in> set evsCF;
              Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCF;
              Notes A {|Agent B, Nonce PMS|} \<in> set evsCF;
-	     M = PRF(PMS,NA,NB) |]
+             M = PRF(PMS,NA,NB) |]
           ==> Says A B (Crypt (clientK(NA,NB,M))
-			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A,
-			       Nonce NB, Number PB, Agent B|}))
+                        (Hash{|Number SID, Nonce M,
+                               Nonce NA, Number PA, Agent A,
+                               Nonce NB, Number PB, Agent B|}))
               # evsCF  \<in>  tls"
 
  | ServerFinished:
-	--{*Keeping A' and A'' distinct means B cannot even check that the
+        --{*Keeping A' and A'' distinct means B cannot even check that the
           two messages originate from the same source. *}
          "[| evsSF \<in> tls;
-	     Says A' B  {|Agent A, Nonce NA, Number SID, Number PA|}
-	       \<in> set evsSF;
-	     Says B  A  {|Nonce NB, Number SID, Number PB|} \<in> set evsSF;
-	     Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
-	     M = PRF(PMS,NA,NB) |]
+             Says A' B  {|Agent A, Nonce NA, Number SID, Number PA|}
+               \<in> set evsSF;
+             Says B  A  {|Nonce NB, Number SID, Number PB|} \<in> set evsSF;
+             Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
+             M = PRF(PMS,NA,NB) |]
           ==> Says B A (Crypt (serverK(NA,NB,M))
-			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A,
-			       Nonce NB, Number PB, Agent B|}))
+                        (Hash{|Number SID, Nonce M,
+                               Nonce NA, Number PA, Agent A,
+                               Nonce NB, Number PB, Agent B|}))
               # evsSF  \<in>  tls"
 
  | ClientAccepts:
-	--{*Having transmitted ClientFinished and received an identical
+        --{*Having transmitted ClientFinished and received an identical
           message encrypted with serverK, the client stores the parameters
           needed to resume this session.  The "Notes A ..." premise is
           used to prove @{text Notes_master_imp_Crypt_PMS}.*}
          "[| evsCA \<in> tls;
-	     Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
-	     M = PRF(PMS,NA,NB);
-	     X = Hash{|Number SID, Nonce M,
-	               Nonce NA, Number PA, Agent A,
-		       Nonce NB, Number PB, Agent B|};
+             Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
+             M = PRF(PMS,NA,NB);
+             X = Hash{|Number SID, Nonce M,
+                       Nonce NA, Number PA, Agent A,
+                       Nonce NB, Number PB, Agent B|};
              Says A  B (Crypt (clientK(NA,NB,M)) X) \<in> set evsCA;
              Says B' A (Crypt (serverK(NA,NB,M)) X) \<in> set evsCA |]
           ==>
              Notes A {|Number SID, Agent A, Agent B, Nonce M|} # evsCA  \<in>  tls"
 
  | ServerAccepts:
-	--{*Having transmitted ServerFinished and received an identical
+        --{*Having transmitted ServerFinished and received an identical
           message encrypted with clientK, the server stores the parameters
           needed to resume this session.  The "Says A'' B ..." premise is
           used to prove @{text Notes_master_imp_Crypt_PMS}.*}
          "[| evsSA \<in> tls;
-	     A \<noteq> B;
+             A \<noteq> B;
              Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSA;
-	     M = PRF(PMS,NA,NB);
-	     X = Hash{|Number SID, Nonce M,
-	               Nonce NA, Number PA, Agent A,
-		       Nonce NB, Number PB, Agent B|};
+             M = PRF(PMS,NA,NB);
+             X = Hash{|Number SID, Nonce M,
+                       Nonce NA, Number PA, Agent A,
+                       Nonce NB, Number PB, Agent B|};
              Says B  A (Crypt (serverK(NA,NB,M)) X) \<in> set evsSA;
              Says A' B (Crypt (clientK(NA,NB,M)) X) \<in> set evsSA |]
           ==>
@@ -244,27 +244,27 @@
          --{*If A recalls the @{text SESSION_ID}, then she sends a FINISHED
              message using the new nonces and stored MASTER SECRET.*}
          "[| evsCR \<in> tls;
-	     Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
+             Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
              Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCR;
              Notes A {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsCR |]
           ==> Says A B (Crypt (clientK(NA,NB,M))
-			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A,
-			       Nonce NB, Number PB, Agent B|}))
+                        (Hash{|Number SID, Nonce M,
+                               Nonce NA, Number PA, Agent A,
+                               Nonce NB, Number PB, Agent B|}))
               # evsCR  \<in>  tls"
 
  | ServerResume:
          --{*Resumption (7.3):  If B finds the @{text SESSION_ID} then he can 
              send a FINISHED message using the recovered MASTER SECRET*}
          "[| evsSR \<in> tls;
-	     Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
-	     Says B  A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
+             Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
+             Says B  A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
              Notes B {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsSR |]
           ==> Says B A (Crypt (serverK(NA,NB,M))
-			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A,
-			       Nonce NB, Number PB, Agent B|})) # evsSR
-	        \<in>  tls"
+                        (Hash{|Number SID, Nonce M,
+                               Nonce NA, Number PA, Agent A,
+                               Nonce NB, Number PB, Agent B|})) # evsSR
+                \<in>  tls"
 
  | Oops:
          --{*The most plausible compromise is of an old session key.  Losing
@@ -273,7 +273,7 @@
            otherwise the Spy could learn session keys merely by 
            replaying messages!*}
          "[| evso \<in> tls;  A \<noteq> Spy;
-	     Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
+             Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
           ==> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso  \<in>  tls"
 
 (*
@@ -328,7 +328,7 @@
 
 
 (** These proofs assume that the Nonce_supply nonces
-	(which have the form  @ N. Nonce N \<notin> used evs)
+        (which have the form  @ N. Nonce N \<notin> used evs)
     lie outside the range of PRF.  It seems reasonable, but as it is needed
     only for the possibility theorems, it is not taken as an axiom.
 **)
@@ -381,11 +381,11 @@
           \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF;
           A \<noteq> B |]
       ==> \<exists>NA PA NB PB X. \<exists>evs \<in> tls.
-		X = Hash{|Number SID, Nonce M,
-			  Nonce NA, Number PA, Agent A,
-			  Nonce NB, Number PB, Agent B|}  &
-		Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evs  &
-		Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evs"
+                X = Hash{|Number SID, Nonce M,
+                          Nonce NA, Number PA, Agent A,
+                          Nonce NB, Number PB, Agent B|}  &
+                Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evs  &
+                Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evs"
 apply (intro exI bexI)
 apply (rule_tac [2] tls.ClientHello
                     [THEN tls.ServerHello,
@@ -570,7 +570,7 @@
           (priK B \<in> KK | B \<in> bad)"
 apply (erule tls.induct)
 apply (simp_all (no_asm_simp)
-		del: image_insert
+                del: image_insert
                 add: image_Un [THEN sym]
                      insert_Key_image Un_assoc [THEN sym])
 txt{*Fake*}
@@ -598,16 +598,16 @@
 lemma analz_image_keys [rule_format]:
      "evs \<in> tls ==>
       \<forall>KK. KK <= range sessionK -->
-	      (Nonce N \<in> analz (Key`KK Un (spies evs))) =
-	      (Nonce N \<in> analz (spies evs))"
+              (Nonce N \<in> analz (Key`KK Un (spies evs))) =
+              (Nonce N \<in> analz (spies evs))"
 apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
 apply (safe del: iffI)
 apply (safe del: impI iffI intro!: analz_image_keys_lemma)
 apply (simp_all (no_asm_simp)               (*faster*)
                 del: image_insert imp_disjL (*reduces blow-up*)
-		add: image_Un [THEN sym]  Un_assoc [THEN sym]
-		     insert_Key_singleton
-		     range_sessionkeys_not_priK analz_image_priK)
+                add: image_Un [THEN sym]  Un_assoc [THEN sym]
+                     insert_Key_singleton
+                     range_sessionkeys_not_priK analz_image_priK)
 apply (simp_all add: insert_absorb)
 txt{*Fake*}
 apply spy_analz
@@ -901,11 +901,11 @@
            down to 433s on albatross*)
 
 (*5/5/01: conversion to Isar script
-	  loads in 137s (perch)
+          loads in 137s (perch)
           the last ML version loaded in 122s on perch, a 600MHz machine:
-		twice as fast as pike.  No idea why it's so much slower!
-	  The Isar script is slower still, perhaps because simp_all simplifies
-	  the assumptions be default.
+                twice as fast as pike.  No idea why it's so much slower!
+          The Isar script is slower still, perhaps because simp_all simplifies
+          the assumptions be default.
 *)
 
 end
--- a/src/HOL/Auth/Yahalom.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/Yahalom.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -58,7 +58,7 @@
            uses the new session key to send Bob his Nonce.  The premise
            @{term "A \<noteq> Server"} is needed for @{text Says_Server_not_range}.
            Alice can check that K is symmetric by its length.*}
-	 "[| evs4 \<in> yahalom;  A \<noteq> Server;  K \<in> symKeys;
+         "[| evs4 \<in> yahalom;  A \<noteq> Server;  K \<in> symKeys;
              Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
                 \<in> set evs4;
              Says A B {|Agent A, Nonce NA|} \<in> set evs4 |]
@@ -481,9 +481,9 @@
 text{*A vital theorem for B, that nonce NB remains secure from the Spy.*}
 lemma Spy_not_see_NB :
      "[| Says B Server
-	        {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
-	   \<in> set evs;
-	 (\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \<notin> set evs);
+                {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
+           \<in> set evs;
+         (\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \<notin> set evs);
          A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
       ==> Nonce NB \<notin> analz (knows Spy evs)"
 apply (erule rev_mp, erule rev_mp)
--- a/src/HOL/Auth/Yahalom2.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/Yahalom2.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -51,7 +51,7 @@
            Both agents are quoted in the 2nd certificate to prevent attacks!*)
  | YM3:  "[| evs3 \<in> yahalom;  Key KAB \<notin> used evs3;
              Gets Server {|Agent B, Nonce NB,
-			   Crypt (shrK B) {|Agent A, Nonce NA|}|}
+                           Crypt (shrK B) {|Agent A, Nonce NA|}|}
                \<in> set evs3 |]
           ==> Says Server A
                {|Nonce NB,
--- a/src/HOL/Auth/ZhouGollmann.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Auth/ZhouGollmann.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -35,7 +35,7 @@
   Nil:  "[] \<in> zg"
 
 | Fake: "[| evsf \<in> zg;  X \<in> synth (analz (spies evsf)) |]
-	 ==> Says Spy B X  # evsf \<in> zg"
+         ==> Says Spy B X  # evsf \<in> zg"
 
 | Reception:  "[| evsr \<in> zg; Says A B X \<in> set evsr |] ==> Gets B X # evsr \<in> zg"
 
@@ -44,26 +44,26 @@
     We just assume that the protocol's objective is to deliver K fairly,
     rather than to keep M secret.*)
 | ZG1: "[| evs1 \<in> zg;  Nonce L \<notin> used evs1; C = Crypt K (Number m);
-	   K \<in> symKeys;
-	   NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
+           K \<in> symKeys;
+           NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
        ==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} # evs1 \<in> zg"
 
   (*B must check that NRO is A's signature to learn the sender's name*)
 | ZG2: "[| evs2 \<in> zg;
-	   Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
-	   NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
-	   NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
+           Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
+           NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
+           NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
        ==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} # evs2  \<in>  zg"
 
   (*A must check that NRR is B's signature to learn the sender's name;
     without spy, the matching label would be enough*)
 | ZG3: "[| evs3 \<in> zg; C = Crypt K M; K \<in> symKeys;
-	   Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
-	   Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
-	   NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
-	   sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|]
+           Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
+           Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
+           NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
+           sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|]
        ==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
-	     # evs3 \<in> zg"
+             # evs3 \<in> zg"
 
  (*TTP checks that sub_K is A's signature to learn who issued K, then
    gives credentials to A and B.  The Notes event models the availability of
@@ -72,15 +72,15 @@
    also allowing lemma @{text Crypt_used_imp_spies} to omit the condition
    @{term "K \<noteq> priK TTP"}. *)
 | ZG4: "[| evs4 \<in> zg; K \<in> symKeys;
-	   Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
-	     \<in> set evs4;
-	   sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
-	   con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
-				      Nonce L, Key K|}|]
+           Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
+             \<in> set evs4;
+           sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
+           con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
+                                      Nonce L, Key K|}|]
        ==> Says TTP Spy con_K
            #
-	   Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
-	   # evs4 \<in> zg"
+           Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
+           # evs4 \<in> zg"
 
 
 declare Says_imp_knows_Spy [THEN analz.Inj, dest]
--- a/src/HOL/Bali/AxCompl.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/AxCompl.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -319,13 +319,13 @@
       case 0
       with is_cls
       show ?thesis
-	by - (rule ax_impossible [THEN conseq1],fastsimp dest: nyinitcls_emptyD)
+        by - (rule ax_impossible [THEN conseq1],fastsimp dest: nyinitcls_emptyD)
     next
       case (Suc m)
       with mgf_hyp have mgf_hyp': "\<And> t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>}"
-	by simp
+        by simp
       from is_cls obtain c where c: "the (class G C) = c"
-	by auto
+        by auto
       let ?Q= "(\<lambda>Y s' (x,s) . 
           G\<turnstile> (x,init_class_obj G C s) 
              \<midarrow> (if C=Object then Skip else Init (super (the (class G C))))\<rightarrow> s'
@@ -333,45 +333,45 @@
       from c
       show ?thesis
       proof (rule ax_derivs.Init [where ?Q="?Q"])
-	let ?P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s 
+        let ?P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s 
                            \<and> normal s \<and> \<not> initd C s) \<and>. G\<turnstile>init\<le>m)" 
-	show "G,A\<turnstile>{Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
+        show "G,A\<turnstile>{Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
                   .(if C = Object then Skip else Init (super c)). 
                   {?Q}"
-	proof (rule conseq1 [where ?P'="?P'"])
-	  show "G,A\<turnstile>{?P'} .(if C = Object then Skip else Init (super c)). {?Q}"
-	  proof (cases "C=Object")
-	    case True
-	    have "G,A\<turnstile>{?P'} .Skip. {?Q}"
-	      by (rule ax_derivs.Skip [THEN conseq1])
-	         (auto simp add: True intro: eval.Skip)
+        proof (rule conseq1 [where ?P'="?P'"])
+          show "G,A\<turnstile>{?P'} .(if C = Object then Skip else Init (super c)). {?Q}"
+          proof (cases "C=Object")
+            case True
+            have "G,A\<turnstile>{?P'} .Skip. {?Q}"
+              by (rule ax_derivs.Skip [THEN conseq1])
+                 (auto simp add: True intro: eval.Skip)
             with True show ?thesis 
-	      by simp
-	  next
-	    case False
-	    from mgf_hyp'
-	    have "G,A\<turnstile>{?P'} .Init (super c). {?Q}"
-	      by (rule MGFnD' [THEN conseq12]) (fastsimp simp add: False c)
-	    with False show ?thesis
-	      by simp
-	  qed
-	next
-	  from Suc is_cls
-	  show "Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))
+              by simp
+          next
+            case False
+            from mgf_hyp'
+            have "G,A\<turnstile>{?P'} .Init (super c). {?Q}"
+              by (rule MGFnD' [THEN conseq12]) (fastsimp simp add: False c)
+            with False show ?thesis
+              by simp
+          qed
+        next
+          from Suc is_cls
+          show "Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))
                 \<Rightarrow> ?P'"
-	    by (fastsimp elim: nyinitcls_le_SucD)
-	qed
+            by (fastsimp elim: nyinitcls_le_SucD)
+        qed
       next
-	from mgf_hyp'
-	show "\<forall>l. G,A\<turnstile>{?Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty} 
+        from mgf_hyp'
+        show "\<forall>l. G,A\<turnstile>{?Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty} 
                       .init c.
                       {set_lvars l .; ?R}"
-	  apply (rule MGFnD' [THEN conseq12, THEN allI])
-	  apply (clarsimp simp add: split_paired_all)
-	  apply (rule eval.Init [OF c])
-	  apply (insert c)
-	  apply auto
-	  done
+          apply (rule MGFnD' [THEN conseq12, THEN allI])
+          apply (clarsimp simp add: split_paired_all)
+          apply (rule eval.Init [OF c])
+          apply (insert c)
+          apply auto
+          done
       qed
     qed
     thus "G,A\<turnstile>{Normal ?P  \<and>. Not \<circ> initd C} .Init C. {?R}"
@@ -399,7 +399,7 @@
                  mode: "mode = invmode statM e" and
                     T: "T =Inl (resTy statM)" and
         eq_accC_accC': "accC=accC'"
-	by cases fastsimp+
+        by cases fastsimp+
   let ?Q="(\<lambda>Y s1 (x,s) . x = None \<and> 
               (\<exists>a. G\<turnstile>Norm s \<midarrow>e-\<succ>a\<rightarrow> s1 \<and> 
                    (normal s1 \<longrightarrow> G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT)
@@ -435,29 +435,29 @@
               (\<exists>P. \<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P))
             \<and> s1\<Colon>\<preceq>(G, L)"
       proof -
-	from da obtain C where
-	  da_e:  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
+        from da obtain C where
+          da_e:  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
                     dom (locals (store ((Norm s0)::state)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
-	  da_ps: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> E" 
-	  by cases (simp add: eq_accC_accC')
-	from eval_e conf_s0 wt_e da_e wf
-	obtain "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT)"
-	  and  "s1\<Colon>\<preceq>(G, L)"
-	  by (rule eval_type_soundE) simp
-	moreover
-	{
-	  assume normal_s1: "normal s1"
-	  have "\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
-	  proof -
-	    from eval_e wt_e da_e wf normal_s1
-	    have "nrm C \<subseteq>  dom (locals (store s1))"
-	      by (cases rule: da_good_approxE') iprover
-	    with da_ps show ?thesis
-	      by (rule da_weakenE) iprover
-	  qed
-	}
-	ultimately show ?thesis
-	  using eq_accC_accC' by simp
+          da_ps: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> E" 
+          by cases (simp add: eq_accC_accC')
+        from eval_e conf_s0 wt_e da_e wf
+        obtain "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT)"
+          and  "s1\<Colon>\<preceq>(G, L)"
+          by (rule eval_type_soundE) simp
+        moreover
+        {
+          assume normal_s1: "normal s1"
+          have "\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
+          proof -
+            from eval_e wt_e da_e wf normal_s1
+            have "nrm C \<subseteq>  dom (locals (store s1))"
+              by (cases rule: da_good_approxE') iprover
+            with da_ps show ?thesis
+              by (rule da_weakenE) iprover
+          qed
+        }
+        ultimately show ?thesis
+          using eq_accC_accC' by simp
       qed
     qed
   next
@@ -467,36 +467,36 @@
       show "?PS a"
       proof (rule MGFnD' [OF mgf_ps, THEN conseq12],
              clarsimp simp add: eq_accC_accC' [symmetric])
-	fix s0 s1 s2 vs
-	assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
-	assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
-	assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
-	assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
-	assume da_ps: "abrupt s1 = None \<longrightarrow> 
+        fix s0 s1 s2 vs
+        assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
+        assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+        assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+        assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
+        assume da_ps: "abrupt s1 = None \<longrightarrow> 
                        (\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
                                dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P)"
-	show "(\<exists>s1. G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
+        show "(\<exists>s1. G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
                 (abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT) \<and>
                 G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2) \<and>
               s2\<Colon>\<preceq>(G, L)"
-	proof (cases "normal s1")
-	  case True
-	  with da_ps obtain P where
-	   "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
-	    by auto
-	  from eval_ps conf_s1 wt_args this wf
-	  have "s2\<Colon>\<preceq>(G, L)"
-	    by (rule eval_type_soundE)
-	  with eval_e conf_a eval_ps 
-	  show ?thesis 
-	    by auto
-	next
-	  case False
-	  with eval_ps have "s2=s1" by auto
-	  with eval_e conf_a eval_ps conf_s1 
-	  show ?thesis 
-	    by auto
-	qed
+        proof (cases "normal s1")
+          case True
+          with da_ps obtain P where
+           "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
+            by auto
+          from eval_ps conf_s1 wt_args this wf
+          have "s2\<Colon>\<preceq>(G, L)"
+            by (rule eval_type_soundE)
+          with eval_e conf_a eval_ps 
+          show ?thesis 
+            by auto
+        next
+          case False
+          with eval_ps have "s2=s1" by auto
+          with eval_e conf_a eval_ps conf_s1 
+          show ?thesis 
+            by auto
+        qed
       qed
     qed
   next
@@ -517,52 +517,52 @@
       from mgf_methds [rule_format]
       show "?METHD a vs invC declC l"
       proof (rule MGFnD' [THEN conseq12],clarsimp)
-	fix s4 s2 s1::state
-	fix s0 v
-	let ?D= "invocation_declclass G mode (store s2) a statT 
+        fix s4 s2 s1::state
+        fix s0 v
+        let ?D= "invocation_declclass G mode (store s2) a statT 
                     \<lparr>name=mn,parTs=pTs'\<rparr>"
-	let ?s3= "init_lvars G ?D \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs s2"
-	assume inv_prop: "abrupt ?s3=None 
+        let ?s3= "init_lvars G ?D \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs s2"
+        assume inv_prop: "abrupt ?s3=None 
              \<longrightarrow> G\<turnstile>mode\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
-	assume conf_s2: "s2\<Colon>\<preceq>(G, L)"
-	assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
-	assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
-	assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
-	assume eval_mthd: "G\<turnstile>?s3 \<midarrow>Methd ?D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
-	show "G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v
+        assume conf_s2: "s2\<Colon>\<preceq>(G, L)"
+        assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+        assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+        assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
+        assume eval_mthd: "G\<turnstile>?s3 \<midarrow>Methd ?D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+        show "G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v
                         \<rightarrow> (set_lvars (locals (store s2))) s4"
-	proof -
-	  obtain D where D: "D=?D" by simp
-	  obtain s3 where s3: "s3=?s3" by simp
-	  obtain s3' where 
-	    s3': "s3' = check_method_access G accC statT mode 
+        proof -
+          obtain D where D: "D=?D" by simp
+          obtain s3 where s3: "s3=?s3" by simp
+          obtain s3' where 
+            s3': "s3' = check_method_access G accC statT mode 
                            \<lparr>name=mn,parTs=pTs'\<rparr> a s3"
-	    by simp
-	  have eq_s3'_s3: "s3'=s3"
-	  proof -
-	    from inv_prop s3 mode
-	    have "normal s3 \<Longrightarrow> 
+            by simp
+          have eq_s3'_s3: "s3'=s3"
+          proof -
+            from inv_prop s3 mode
+            have "normal s3 \<Longrightarrow> 
              G\<turnstile>invmode statM e\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
-	      by auto
-	    with eval_ps wt_e statM conf_s2 conf_a [rule_format] 
-	    have "check_method_access G accC statT (invmode statM e)
+              by auto
+            with eval_ps wt_e statM conf_s2 conf_a [rule_format] 
+            have "check_method_access G accC statT (invmode statM e)
                       \<lparr>name=mn,parTs=pTs'\<rparr> a s3 = s3"
-	      by (rule error_free_call_access) (auto simp add: s3 mode wf)
-	    thus ?thesis 
-	      by (simp add: s3' mode)
-	  qed
-	  with eval_mthd D s3
-	  have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
-	    by simp
-	  with eval_e eval_ps D _ s3' 
-	  show ?thesis
-	    by (rule eval_Call) (auto simp add: s3 mode D)
-	qed
+              by (rule error_free_call_access) (auto simp add: s3 mode wf)
+            thus ?thesis 
+              by (simp add: s3' mode)
+          qed
+          with eval_mthd D s3
+          have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+            by simp
+          with eval_e eval_ps D _ s3' 
+          show ?thesis
+            by (rule eval_Call) (auto simp add: s3 mode D)
+        qed
       qed
     qed
   qed
 qed
-	  	  
+                  
 lemma eval_expression_no_jump':
   assumes eval: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<rightarrow> s1"
   and   no_jmp: "abrupt s0 \<noteq> Some (Jump j)"
@@ -610,36 +610,36 @@
       case True
       with normal_t eval_e normal_termination
       show ?thesis
-	by (auto intro: eval.Loop)
+        by (auto intro: eval.Loop)
     next
       case False
       note abrupt_s' = this
       from eval_e _ wt wf
       have no_cont: "abrupt s' \<noteq> Some (Jump (Cont l))"
-	by (rule eval_expression_no_jump') (insert normal_t,simp)
+        by (rule eval_expression_no_jump') (insert normal_t,simp)
       have
-	"if the_Bool v 
+        "if the_Bool v 
              then (G\<turnstile>s' \<midarrow>c\<rightarrow> s' \<and> 
                    G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s')
-	     else s' = s'"
+             else s' = s'"
       proof (cases "the_Bool v")
-	case False thus ?thesis by simp
+        case False thus ?thesis by simp
       next
-	case True
-	with abrupt_s' have "G\<turnstile>s' \<midarrow>c\<rightarrow> s'" by auto
-	moreover from abrupt_s' no_cont 
-	have no_absorb: "(abupd (absorb (Cont l)) s')=s'"
-	  by (cases s') (simp add: absorb_def split: split_if)
-	moreover
-	from no_absorb abrupt_s'
-	have "G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
-	  by auto
-	ultimately show ?thesis
-	  using True by simp
+        case True
+        with abrupt_s' have "G\<turnstile>s' \<midarrow>c\<rightarrow> s'" by auto
+        moreover from abrupt_s' no_cont 
+        have no_absorb: "(abupd (absorb (Cont l)) s')=s'"
+          by (cases s') (simp add: absorb_def split: split_if)
+        moreover
+        from no_absorb abrupt_s'
+        have "G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+          by auto
+        ultimately show ?thesis
+          using True by simp
       qed
       with eval_e 
       show ?thesis
-	using normal_t by (auto intro: eval.Loop)
+        using normal_t by (auto intro: eval.Loop)
     qed
   qed
 next
@@ -690,49 +690,49 @@
                                   Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s')) 
                               \<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>}"
       proof (rule ax_derivs.Loop)
-	from mfg_e
-	show "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n} 
+        from mfg_e
+        show "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n} 
                    e-\<succ>
                   {(\<lambda>Y s' s. (\<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and> 
                                      Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s')) 
                    \<and>. G\<turnstile>init\<le>n}"
-	proof (rule MGFnD' [THEN conseq12],clarsimp)
-	  fix s Z s' v
-	  assume "(Z, s) \<in> (unroll G l e c)\<^sup>*"
-	  moreover
-	  assume "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s'"
-	  ultimately
-	  show "\<exists>t. (Z, t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'"
-	    by blast
-	qed
+        proof (rule MGFnD' [THEN conseq12],clarsimp)
+          fix s Z s' v
+          assume "(Z, s) \<in> (unroll G l e c)\<^sup>*"
+          moreover
+          assume "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s'"
+          ultimately
+          show "\<exists>t. (Z, t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'"
+            by blast
+        qed
       next
-	from mfg_c
-	show "G,A\<turnstile>{Normal (((\<lambda>Y s' s. \<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
+        from mfg_c
+        show "G,A\<turnstile>{Normal (((\<lambda>Y s' s. \<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
                                        Y = \<lfloor>b\<rfloor>\<^sub>e \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s') 
                           \<and>. G\<turnstile>init\<le>n)\<leftarrow>=True)}
                   .c.
                   {abupd (absorb (Cont l)) .;
                    ((\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n)}"
-	proof (rule MGFnD' [THEN conseq12],clarsimp)
-	  fix Z s' s v t
-	  assume unroll: "(Z, t) \<in> (unroll G l e c)\<^sup>*"
-	  assume eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> Norm s" 
-	  assume true: "the_Bool v"
-	  assume eval_c: "G\<turnstile>Norm s \<midarrow>c\<rightarrow> s'"
-	  show "(Z, abupd (absorb (Cont l)) s') \<in> (unroll G l e c)\<^sup>*"
-	  proof -
-	    note unroll
-	    also
-	    from eval_e true eval_c
-	    have "(t,abupd (absorb (Cont l)) s') \<in> unroll G l e c" 
-	      by (unfold unroll_def) force
-	    ultimately show ?thesis ..
-	  qed
-	qed
+        proof (rule MGFnD' [THEN conseq12],clarsimp)
+          fix Z s' s v t
+          assume unroll: "(Z, t) \<in> (unroll G l e c)\<^sup>*"
+          assume eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> Norm s" 
+          assume true: "the_Bool v"
+          assume eval_c: "G\<turnstile>Norm s \<midarrow>c\<rightarrow> s'"
+          show "(Z, abupd (absorb (Cont l)) s') \<in> (unroll G l e c)\<^sup>*"
+          proof -
+            note unroll
+            also
+            from eval_e true eval_c
+            have "(t,abupd (absorb (Cont l)) s') \<in> unroll G l e c" 
+              by (unfold unroll_def) force
+            ultimately show ?thesis ..
+          qed
+        qed
       qed
     next
       show 
-	"\<forall>Y s Z.
+        "\<forall>Y s Z.
          (Normal ((\<lambda>Y' s' s. s' = s \<and> normal s) \<and>. G\<turnstile>init\<le>n)) Y s Z 
          \<longrightarrow> (\<forall>Y' s'.
                (\<forall>Y Z'. 
@@ -742,28 +742,28 @@
                      \<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>) Y' s' Z') 
                \<longrightarrow> G\<turnstile>Z \<midarrow>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ>\<rightarrow> (Y', s'))"
       proof (clarsimp)
-	fix Y' s' s
-	assume asm:
-	  "\<forall>Z'. (Z', Norm s) \<in> (unroll G l e c)\<^sup>* 
+        fix Y' s' s
+        assume asm:
+          "\<forall>Z'. (Z', Norm s) \<in> (unroll G l e c)\<^sup>* 
                  \<longrightarrow> card (nyinitcls G s') \<le> n \<and>
                      (\<exists>v. (\<exists>t. (Z', t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s') \<and>
                      (fst s' = None \<longrightarrow> \<not> the_Bool v)) \<and> Y' = \<diamondsuit>"
-	show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
-	proof -
-	  from asm obtain v t where 
-	    -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}  
-	    unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
+        show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+        proof -
+          from asm obtain v t where 
+            -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}  
+            unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
             eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'" and
             normal_termination: "normal s' \<longrightarrow> \<not> the_Bool v" and
-	     Y': "Y' = \<diamondsuit>"
-	    by auto
-	  from unroll eval_e normal_termination wt_e wf
-	  have "G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
-	    by (rule unroll_while)
-	  with Y' 
-	  show ?thesis
-	    by simp
-	qed
+             Y': "Y' = \<diamondsuit>"
+            by auto
+          from unroll eval_e normal_termination wt_e wf
+          have "G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+            by (rule unroll_while)
+          with Y' 
+          show ?thesis
+            by simp
+        qed
       qed
     qed
   qed
@@ -810,39 +810,39 @@
       show "(\<exists>E. \<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E) \<and>
             s'\<Colon>\<preceq>(G, L)"
       proof -
-	from da 
-	obtain E where
-	  "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals s) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
-	  by cases simp
-	moreover
-	from eval_init
-	have "dom (locals s) \<subseteq> dom (locals (store s'))"
-	  by (rule dom_locals_eval_mono [elim_format]) simp
-	ultimately obtain E' where
-	  "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
-	  by (rule da_weakenE)
-	moreover
-	have "s'\<Colon>\<preceq>(G, L)"
-	proof -
-	  have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
-	  proof -
-	    from wf wt_e 
-	    have iscls_statC: "is_class G statC"
-	      by (auto dest: ty_expr_is_type type_is_class)
-	    with wf accfield 
-	    have iscls_statDeclC: "is_class G statDeclC"
-	      by (auto dest!: accfield_fields dest: fields_declC)
-	    thus ?thesis by simp
-	  qed
-	  obtain I where 
-	    da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        from da 
+        obtain E where
+          "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals s) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+          by cases simp
+        moreover
+        from eval_init
+        have "dom (locals s) \<subseteq> dom (locals (store s'))"
+          by (rule dom_locals_eval_mono [elim_format]) simp
+        ultimately obtain E' where
+          "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
+          by (rule da_weakenE)
+        moreover
+        have "s'\<Colon>\<preceq>(G, L)"
+        proof -
+          have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
+          proof -
+            from wf wt_e 
+            have iscls_statC: "is_class G statC"
+              by (auto dest: ty_expr_is_type type_is_class)
+            with wf accfield 
+            have iscls_statDeclC: "is_class G statDeclC"
+              by (auto dest!: accfield_fields dest: fields_declC)
+            thus ?thesis by simp
+          qed
+          obtain I where 
+            da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                \<turnstile> dom (locals (store ((Norm s)::state))) \<guillemotright>\<langle>Init statDeclC\<rangle>\<^sub>s\<guillemotright> I"
-	    by (auto intro: da_Init [simplified] assigned.select_convs)
-	  from eval_init conf_s wt_init da_init  wf
-	  show ?thesis
-	    by (rule eval_type_soundE)
-	qed
-	ultimately show ?thesis by iprover
+            by (auto intro: da_Init [simplified] assigned.select_convs)
+          from eval_init conf_s wt_init da_init  wf
+          show ?thesis
+            by (rule eval_type_soundE)
+        qed
+        ultimately show ?thesis by iprover
       qed
     qed
   next
@@ -857,30 +857,30 @@
       assume da_e: "\<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
       show "G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>fst ?fvar\<rightarrow> snd ?fvar"
       proof -
-	obtain v s2' where
-	  v: "v=fst ?fvar" and s2': "s2'=snd ?fvar"
-	  by simp
-	obtain s3 where
-	  s3: "s3= check_field_access G accC' statDeclC fn stat a s2'"
-	  by simp
-	have eq_s3_s2': "s3=s2'"
-	proof -
-	  from eval_e conf_s1 wt_e da_e wf obtain
-	    conf_s2: "s2\<Colon>\<preceq>(G, L)"  and
-	    conf_a: "normal s2 \<Longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
-	    by (rule eval_type_soundE) simp
-	  from accfield wt_e eval_init eval_e conf_s2 conf_a _ wf
-	  show ?thesis
-	    by (rule  error_free_field_access 
+        obtain v s2' where
+          v: "v=fst ?fvar" and s2': "s2'=snd ?fvar"
+          by simp
+        obtain s3 where
+          s3: "s3= check_field_access G accC' statDeclC fn stat a s2'"
+          by simp
+        have eq_s3_s2': "s3=s2'"
+        proof -
+          from eval_e conf_s1 wt_e da_e wf obtain
+            conf_s2: "s2\<Colon>\<preceq>(G, L)"  and
+            conf_a: "normal s2 \<Longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+            by (rule eval_type_soundE) simp
+          from accfield wt_e eval_init eval_e conf_s2 conf_a _ wf
+          show ?thesis
+            by (rule  error_free_field_access 
                       [where ?v=v and ?s2'=s2',elim_format])
-	       (simp add: s3 v s2' stat)+
+               (simp add: s3 v s2' stat)+
         qed
-	from eval_init eval_e 
-	show ?thesis
-	  apply (rule eval.FVar [where ?s2'=s2'])
-	  apply  (simp add: s2')
-	  apply  (simp add: s3 [symmetric]   eq_s3_s2' eq_accC s2' [symmetric])
-	  done
+        from eval_init eval_e 
+        show ?thesis
+          apply (rule eval.FVar [where ?s2'=s2'])
+          apply  (simp add: s2')
+          apply  (simp add: s3 [symmetric]   eq_s3_s2' eq_accC s2' [symmetric])
+          done
       qed
     qed
   qed
@@ -918,7 +918,7 @@
       fix s0
       assume "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1 Finally c2\<rangle>\<^sub>s\<guillemotright> C"
       thus "\<exists>C1. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1"
-	by cases (auto simp add: inj_term_simps)
+        by cases (auto simp add: inj_term_simps)
     qed
   next
     from mgf_c2
@@ -933,25 +933,25 @@
       show "G\<turnstile>Norm s0 \<midarrow>c1 Finally c2
                \<rightarrow> abupd (abrupt_if (\<exists>y. abrupt s1 = Some y) (abrupt s1)) s2"
       proof -
-	obtain abr1 str1 where s1: "s1=(abr1,str1)"
-	  by (cases s1)
-	with eval_c1 eval_c2 obtain
-	  eval_c1': "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (abr1,str1)" and
-	  eval_c2': "G\<turnstile>Norm str1 \<midarrow>c2\<rightarrow> s2"
-	  by simp
-	obtain s3 where 
-	  s3: "s3 = (if \<exists>err. abr1 = Some (Error err) 
-	                then (abr1, str1)
+        obtain abr1 str1 where s1: "s1=(abr1,str1)"
+          by (cases s1)
+        with eval_c1 eval_c2 obtain
+          eval_c1': "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (abr1,str1)" and
+          eval_c2': "G\<turnstile>Norm str1 \<midarrow>c2\<rightarrow> s2"
+          by simp
+        obtain s3 where 
+          s3: "s3 = (if \<exists>err. abr1 = Some (Error err) 
+                        then (abr1, str1)
                         else abupd (abrupt_if (abr1 \<noteq> None) abr1) s2)"
-	  by simp
-	from eval_c1' conf_s0 wt_c1 _ wf 
-	have "error_free (abr1,str1)"
-	  by (rule eval_type_soundE) (insert da_c1,auto)
-	with s3 have eq_s3: "s3=abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
-	  by (simp add: error_free_def)
-	from eval_c1' eval_c2' s3
-	show ?thesis
-	  by (rule eval.Fin [elim_format]) (simp add: s1 eq_s3)
+          by simp
+        from eval_c1' conf_s0 wt_c1 _ wf 
+        have "error_free (abr1,str1)"
+          by (rule eval_type_soundE) (insert da_c1,auto)
+        with s3 have eq_s3: "s3=abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
+          by (simp add: error_free_def)
+        from eval_c1' eval_c2' s3
+        show ?thesis
+          by (rule eval.Fin [elim_format]) (simp add: s1 eq_s3)
       qed
     qed 
   qed
@@ -1009,7 +1009,7 @@
       fix s0
       assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>Body D c\<rangle>\<^sub>e\<guillemotright> E"
       thus "jumpNestingOkS {Ret} c"
-	by cases simp
+        by cases simp
     qed
   next
     from mgf_c
@@ -1022,22 +1022,22 @@
       show "G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
               \<rightarrow> abupd (absorb Ret) s2"
       proof -
-	from wt obtain d where 
+        from wt obtain d where 
           d: "class G D=Some d" and
           wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
-	  by cases auto
-	obtain s3 where 
-	  s3: "s3= (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
+          by cases auto
+        obtain s3 where 
+          s3: "s3= (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
                            fst s2 = Some (Jump (Cont l))
                        then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 
                        else s2)"
-	  by simp
-	from eval_init eval_c nestingOk wt_c d wf
-	have eq_s3_s2: "s3=s2"
-	  by (rule Body_no_break [elim_format]) (simp add: s3)
-	from eval_init eval_c s3
-	show ?thesis
-	  by (rule eval.Body [elim_format]) (simp add: eq_s3_s2)
+          by simp
+        from eval_init eval_c nestingOk wt_c d wf
+        have eq_s3_s2: "s3=s2"
+          by (rule Body_no_break [elim_format]) (simp add: s3)
+        from eval_init eval_c s3
+        show ?thesis
+          by (rule eval.Body [elim_format]) (simp add: eq_s3_s2)
       qed
     qed
   qed
@@ -1062,312 +1062,312 @@
     proof (induct rule: var_expr_stmt.inducts)
       case (LVar v)
       show "G,A\<turnstile>{=:n} \<langle>LVar v\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.LVar [THEN conseq1])
-	apply (clarsimp)
-	apply (rule eval.LVar)
-	done
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.LVar [THEN conseq1])
+        apply (clarsimp)
+        apply (rule eval.LVar)
+        done
     next
       case (FVar accC statDeclC stat e fn)
       from MGFn_Init [OF hyp] and `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}` and wf
       show ?case
-	by (rule MGFn_FVar)
+        by (rule MGFn_FVar)
     next
       case (AVar e1 e2)
       note mgf_e1 = `G,A\<turnstile>{=:n} \<langle>e1\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
       note mgf_e2 = `G,A\<turnstile>{=:n} \<langle>e2\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
       show "G,A\<turnstile>{=:n} \<langle>e1.[e2]\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.AVar)
-	apply  (rule MGFnD [OF mgf_e1, THEN ax_NormalD])
-	apply (rule allI)
-	apply (rule MGFnD' [OF mgf_e2, THEN conseq12])
-	apply (fastsimp intro: eval.AVar)
-	done
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.AVar)
+        apply  (rule MGFnD [OF mgf_e1, THEN ax_NormalD])
+        apply (rule allI)
+        apply (rule MGFnD' [OF mgf_e2, THEN conseq12])
+        apply (fastsimp intro: eval.AVar)
+        done
     next
       case (InsInitV c v)
       show ?case
-	by (rule MGFn_NormalI) (rule ax_derivs.InsInitV)
+        by (rule MGFn_NormalI) (rule ax_derivs.InsInitV)
     next
       case (NewC C)
       show ?case
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.NewC)
-	apply (rule MGFn_InitD [OF hyp, THEN conseq2])
-	apply (fastsimp intro: eval.NewC)
-	done
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.NewC)
+        apply (rule MGFn_InitD [OF hyp, THEN conseq2])
+        apply (fastsimp intro: eval.NewC)
+        done
     next
       case (NewA T e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI) 
-	apply (rule ax_derivs.NewA 
+        apply -
+        apply (rule MGFn_NormalI) 
+        apply (rule ax_derivs.NewA 
                [where ?Q = "(\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>In1r (init_comp_ty T) 
                               \<succ>\<rightarrow> (Y',s')) \<and>. G\<turnstile>init\<le>n"])
-	apply  (simp add: init_comp_ty_def split add: split_if)
-	apply  (rule conjI, clarsimp)
-	apply   (rule MGFn_InitD [OF hyp, THEN conseq2])
-	apply   (clarsimp intro: eval.Init)
-	apply  clarsimp
-	apply  (rule ax_derivs.Skip [THEN conseq1])
-	apply  (clarsimp intro: eval.Skip)
-	apply (erule MGFnD' [THEN conseq12])
-	apply (fastsimp intro: eval.NewA)
-	done
+        apply  (simp add: init_comp_ty_def split add: split_if)
+        apply  (rule conjI, clarsimp)
+        apply   (rule MGFn_InitD [OF hyp, THEN conseq2])
+        apply   (clarsimp intro: eval.Init)
+        apply  clarsimp
+        apply  (rule ax_derivs.Skip [THEN conseq1])
+        apply  (clarsimp intro: eval.Skip)
+        apply (erule MGFnD' [THEN conseq12])
+        apply (fastsimp intro: eval.NewA)
+        done
     next
       case (Cast C e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast])
-	apply (fastsimp intro: eval.Cast)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast])
+        apply (fastsimp intro: eval.Cast)
+        done
     next
       case (Inst e C)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst])
-	apply (fastsimp intro: eval.Inst)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst])
+        apply (fastsimp intro: eval.Inst)
+        done
     next
       case (Lit v)
       show ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Lit [THEN conseq1])
-	apply (fastsimp intro: eval.Lit)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Lit [THEN conseq1])
+        apply (fastsimp intro: eval.Lit)
+        done
     next
       case (UnOp unop e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.UnOp)
-	apply (erule MGFnD' [THEN conseq12])
-	apply (fastsimp intro: eval.UnOp)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.UnOp)
+        apply (erule MGFnD' [THEN conseq12])
+        apply (fastsimp intro: eval.UnOp)
+        done
     next
       case (BinOp binop e1 e2)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.BinOp)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (case_tac "need_second_arg binop v1")
-	apply  simp
-	apply  (erule MGFnD' [THEN conseq12])
-	apply  (fastsimp intro: eval.BinOp)
-	apply simp
-	apply (rule ax_Normal_cases)
-	apply  (rule ax_derivs.Skip [THEN conseq1])
-	apply  clarsimp
-	apply  (rule eval_BinOp_arg2_indepI)
-	apply   simp
-	apply  simp
-	apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
-	apply (fastsimp intro: eval.BinOp)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.BinOp)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (case_tac "need_second_arg binop v1")
+        apply  simp
+        apply  (erule MGFnD' [THEN conseq12])
+        apply  (fastsimp intro: eval.BinOp)
+        apply simp
+        apply (rule ax_Normal_cases)
+        apply  (rule ax_derivs.Skip [THEN conseq1])
+        apply  clarsimp
+        apply  (rule eval_BinOp_arg2_indepI)
+        apply   simp
+        apply  simp
+        apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
+        apply (fastsimp intro: eval.BinOp)
+        done
     next
       case Super
       show ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Super [THEN conseq1])
-	apply (fastsimp intro: eval.Super)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Super [THEN conseq1])
+        apply (fastsimp intro: eval.Super)
+        done
     next
       case (Acc v)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc])
-	apply (fastsimp intro: eval.Acc simp add: split_paired_all)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc])
+        apply (fastsimp intro: eval.Acc simp add: split_paired_all)
+        done
     next
       case (Ass v e)
       thus "G,A\<turnstile>{=:n} \<langle>v:=e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Ass)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (erule MGFnD'[THEN conseq12])
-	apply (fastsimp intro: eval.Ass simp add: split_paired_all)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Ass)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (erule MGFnD'[THEN conseq12])
+        apply (fastsimp intro: eval.Ass simp add: split_paired_all)
+        done
     next
       case (Cond e1 e2 e3)
       thus "G,A\<turnstile>{=:n} \<langle>e1 ? e2 : e3\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Cond)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (rule ax_Normal_cases)
-	prefer 2
-	apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
-	apply  (fastsimp intro: eval.Cond)
-	apply (case_tac "b")
-	apply  simp
-	apply  (erule MGFnD'[THEN conseq12])
-	apply  (fastsimp intro: eval.Cond)
-	apply simp
-	apply (erule MGFnD'[THEN conseq12])
-	apply (fastsimp intro: eval.Cond)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Cond)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (rule ax_Normal_cases)
+        prefer 2
+        apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
+        apply  (fastsimp intro: eval.Cond)
+        apply (case_tac "b")
+        apply  simp
+        apply  (erule MGFnD'[THEN conseq12])
+        apply  (fastsimp intro: eval.Cond)
+        apply simp
+        apply (erule MGFnD'[THEN conseq12])
+        apply (fastsimp intro: eval.Cond)
+        done
     next
       case (Call accC statT mode e mn pTs' ps)
       note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
       note mgf_ps = `G,A\<turnstile>{=:n} \<langle>ps\<rangle>\<^sub>l\<succ> {G\<rightarrow>}`
       from mgf_methds mgf_e mgf_ps wf
       show "G,A\<turnstile>{=:n} \<langle>{accC,statT,mode}e\<cdot>mn({pTs'}ps)\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Call)
+        by (rule MGFn_Call)
     next
       case (Methd D mn)
       from mgf_methds
       show "G,A\<turnstile>{=:n} \<langle>Methd D mn\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	by simp
+        by simp
     next
       case (Body D c)
       note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
       from wf MGFn_Init [OF hyp] mgf_c
       show "G,A\<turnstile>{=:n} \<langle>Body D c\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Body)
+        by (rule MGFn_Body)
     next
       case (InsInitE c e)
       show ?case
-	by (rule MGFn_NormalI) (rule ax_derivs.InsInitE)
+        by (rule MGFn_NormalI) (rule ax_derivs.InsInitE)
     next
       case (Callee l e)
       show ?case
-	by (rule MGFn_NormalI) (rule ax_derivs.Callee)
+        by (rule MGFn_NormalI) (rule ax_derivs.Callee)
     next
       case Skip
       show ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Skip [THEN conseq1])
-	apply (fastsimp intro: eval.Skip)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Skip [THEN conseq1])
+        apply (fastsimp intro: eval.Skip)
+        done
     next
       case (Expr e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr])
-	apply (fastsimp intro: eval.Expr)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr])
+        apply (fastsimp intro: eval.Expr)
+        done
     next
       case (Lab l c)
       thus "G,A\<turnstile>{=:n} \<langle>l\<bullet> c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
-	apply (fastsimp intro: eval.Lab)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
+        apply (fastsimp intro: eval.Lab)
+        done
     next
       case (Comp c1 c2)
       thus "G,A\<turnstile>{=:n} \<langle>c1;; c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Comp)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (erule MGFnD' [THEN conseq12])
-	apply (fastsimp intro: eval.Comp) 
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Comp)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (erule MGFnD' [THEN conseq12])
+        apply (fastsimp intro: eval.Comp) 
+        done
     next
       case (If' e c1 c2)
       thus "G,A\<turnstile>{=:n} \<langle>If(e) c1 Else c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.If)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (rule ax_Normal_cases)
-	prefer 2
-	apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
-	apply  (fastsimp intro: eval.If)
-	apply (case_tac "b")
-	apply  simp
-	apply  (erule MGFnD' [THEN conseq12])
-	apply  (fastsimp intro: eval.If)
-	apply simp
-	apply (erule MGFnD' [THEN conseq12])
-	apply (fastsimp intro: eval.If)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.If)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (rule ax_Normal_cases)
+        prefer 2
+        apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
+        apply  (fastsimp intro: eval.If)
+        apply (case_tac "b")
+        apply  simp
+        apply  (erule MGFnD' [THEN conseq12])
+        apply  (fastsimp intro: eval.If)
+        apply simp
+        apply (erule MGFnD' [THEN conseq12])
+        apply (fastsimp intro: eval.If)
+        done
     next
       case (Loop l e c)
       note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
       note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
       from mgf_e mgf_c wf
       show "G,A\<turnstile>{=:n} \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Loop)
+        by (rule MGFn_Loop)
     next
       case (Jmp j)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Jmp [THEN conseq1])
-	apply (auto intro: eval.Jmp simp add: abupd_def2)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Jmp [THEN conseq1])
+        apply (auto intro: eval.Jmp simp add: abupd_def2)
+        done
     next
       case (Throw e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
-	apply (fastsimp intro: eval.Throw)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
+        apply (fastsimp intro: eval.Throw)
+        done
     next
       case (TryC c1 C vn c2)
       thus "G,A\<turnstile>{=:n} \<langle>Try c1 Catch(C vn) c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Try [where 
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Try [where 
           ?Q = " (\<lambda>Y' s' s. normal s \<and> (\<exists>s''. G\<turnstile>s \<midarrow>\<langle>c1\<rangle>\<^sub>s\<succ>\<rightarrow> (Y',s'') \<and> 
                             G\<turnstile>s'' \<midarrow>sxalloc\<rightarrow> s')) \<and>. G\<turnstile>init\<le>n"])
-	apply   (erule MGFnD [THEN ax_NormalD, THEN conseq2])
-	apply   (fastsimp elim: sxalloc_gext [THEN card_nyinitcls_gext])
-	apply  (erule MGFnD'[THEN conseq12])
-	apply  (fastsimp intro: eval.Try)
-	apply (fastsimp intro: eval.Try)
-	done
+        apply   (erule MGFnD [THEN ax_NormalD, THEN conseq2])
+        apply   (fastsimp elim: sxalloc_gext [THEN card_nyinitcls_gext])
+        apply  (erule MGFnD'[THEN conseq12])
+        apply  (fastsimp intro: eval.Try)
+        apply (fastsimp intro: eval.Try)
+        done
     next
       case (Fin c1 c2)
       note mgf_c1 = `G,A\<turnstile>{=:n} \<langle>c1\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
       note mgf_c2 = `G,A\<turnstile>{=:n} \<langle>c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
       from wf mgf_c1 mgf_c2
       show "G,A\<turnstile>{=:n} \<langle>c1 Finally c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Fin)
+        by (rule MGFn_Fin)
     next
       case (FinA abr c)
       show ?case
-	by (rule MGFn_NormalI) (rule ax_derivs.FinA)
+        by (rule MGFn_NormalI) (rule ax_derivs.FinA)
     next
       case (Init C)
       from hyp
       show "G,A\<turnstile>{=:n} \<langle>Init C\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Init)
+        by (rule MGFn_Init)
     next
       case Nil_expr
       show "G,A\<turnstile>{=:n} \<langle>[]\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Nil [THEN conseq1])
-	apply (fastsimp intro: eval.Nil)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Nil [THEN conseq1])
+        apply (fastsimp intro: eval.Nil)
+        done
     next
       case (Cons_expr e es)
       thus "G,A\<turnstile>{=:n} \<langle>e# es\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Cons)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (erule MGFnD'[THEN conseq12])
-	apply (fastsimp intro: eval.Cons)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Cons)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (erule MGFnD'[THEN conseq12])
+        apply (fastsimp intro: eval.Cons)
+        done
     qed
   }
   thus ?thesis
@@ -1519,15 +1519,15 @@
     proof -
       from eval_t type_ok wf 
       obtain n where evaln: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s')"
-	by (rule eval_to_evaln [elim_format]) iprover
+        by (rule eval_to_evaln [elim_format]) iprover
       from valid have 
-	valid_expanded:
-	"\<forall>n Y s Z. P Y s Z \<longrightarrow> type_ok G t s 
+        valid_expanded:
+        "\<forall>n Y s Z. P Y s Z \<longrightarrow> type_ok G t s 
                    \<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s') \<longrightarrow> Q Y' s' Z)"
-	by (simp add: ax_valids_def triple_valid_def)
+        by (simp add: ax_valids_def triple_valid_def)
       from P type_ok evaln
       show "Q Y' s' Z"
-	by (rule valid_expanded [rule_format])
+        by (rule valid_expanded [rule_format])
     qed
   qed 
 qed
--- a/src/HOL/Bali/AxSem.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/AxSem.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -519,7 +519,7 @@
                                  G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}"
 
 | NewA: "\<lbrakk>G,A\<turnstile>{Normal P} .init_comp_ty T. {Q};  G,A\<turnstile>{Q} e-\<succ>
-	  {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\<rbrakk> \<Longrightarrow>
+          {\<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..
--- a/src/HOL/Bali/AxSound.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/AxSound.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -118,14 +118,14 @@
       case 0
       show "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}.  G\<Turnstile>0\<Colon>t"
       proof -
-	{
-	  fix C sig
-	  assume "(C,sig) \<in> ms" 
-	  have "G\<Turnstile>0\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
-	    by (rule Methd_triple_valid2_0)
-	}
-	thus ?thesis
-	  by (simp add: mtriples_def split_def)
+        {
+          fix C sig
+          assume "(C,sig) \<in> ms" 
+          have "G\<Turnstile>0\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
+            by (rule Methd_triple_valid2_0)
+        }
+        thus ?thesis
+          by (simp add: mtriples_def split_def)
       qed
     next
       case (Suc m)
@@ -133,28 +133,28 @@
       note prem = `\<forall>t\<in>A. G\<Turnstile>Suc m\<Colon>t`
       show "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}.  G\<Turnstile>Suc m\<Colon>t"
       proof -
-	{
-	  fix C sig
-	  assume m: "(C,sig) \<in> ms" 
-	  have "G\<Turnstile>Suc m\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
-	  proof -
-	    from prem have prem_m: "\<forall>t\<in>A. G\<Turnstile>m\<Colon>t"
-	      by (rule triples_valid2_Suc)
-	    hence "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}.  G\<Turnstile>m\<Colon>t"
-	      by (rule hyp)
-	    with prem_m
-	    have "\<forall>t\<in>(A \<union> {{P} Methd-\<succ> {Q} | ms}). G\<Turnstile>m\<Colon>t"
-	      by (simp add: ball_Un)
-	    hence "\<forall>t\<in>{{P} body G-\<succ> {Q} | ms}.  G\<Turnstile>m\<Colon>t"
-	      by (rule recursive)
-	    with m have "G\<Turnstile>m\<Colon>{Normal (P C sig)} body G C sig-\<succ> {Q C sig}"
-	      by (auto simp add: mtriples_def split_def)
-	    thus ?thesis
-	      by (rule Methd_triple_valid2_SucI)
-	  qed
-	}
-	thus ?thesis
-	  by (simp add: mtriples_def split_def)
+        {
+          fix C sig
+          assume m: "(C,sig) \<in> ms" 
+          have "G\<Turnstile>Suc m\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
+          proof -
+            from prem have prem_m: "\<forall>t\<in>A. G\<Turnstile>m\<Colon>t"
+              by (rule triples_valid2_Suc)
+            hence "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}.  G\<Turnstile>m\<Colon>t"
+              by (rule hyp)
+            with prem_m
+            have "\<forall>t\<in>(A \<union> {{P} Methd-\<succ> {Q} | ms}). G\<Turnstile>m\<Colon>t"
+              by (simp add: ball_Un)
+            hence "\<forall>t\<in>{{P} body G-\<succ> {Q} | ms}.  G\<Turnstile>m\<Colon>t"
+              by (rule recursive)
+            with m have "G\<Turnstile>m\<Colon>{Normal (P C sig)} body G C sig-\<succ> {Q C sig}"
+              by (auto simp add: mtriples_def split_def)
+            thus ?thesis
+              by (rule Methd_triple_valid2_SucI)
+          qed
+        }
+        thus ?thesis
+          by (simp add: mtriples_def split_def)
       qed
     qed
   }
@@ -361,10 +361,10 @@
     have "G\<Turnstile>n\<Colon>t" and "\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t"
     proof -
       from valid_A valid_t show "G\<Turnstile>n\<Colon>t"
-	by (simp add: ax_valids2_def)
+        by (simp add: ax_valids2_def)
     next
       from valid_A valid_ts show "\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t"
-	by (unfold ax_valids2_def) blast
+        by (unfold ax_valids2_def) blast
     qed
     hence "\<forall>t'\<in>insert t ts. G\<Turnstile>n\<Colon>t'"
       by simp
@@ -402,27 +402,27 @@
     proof -
       from valid_A conf wt da eval P con
       have "Q v s1 Z"
-	apply (simp add: ax_valids2_def triple_valid2_def2)
-	apply (tactic "smp_tac 3 1")
-	apply clarify
-	apply (tactic "smp_tac 1 1")
-	apply (erule allE,erule allE, erule mp)
-	apply (intro strip)
-	apply (tactic "smp_tac 3 1")
-	apply (tactic "smp_tac 2 1")
-	apply (tactic "smp_tac 1 1")
-	by blast
+        apply (simp add: ax_valids2_def triple_valid2_def2)
+        apply (tactic "smp_tac 3 1")
+        apply clarify
+        apply (tactic "smp_tac 1 1")
+        apply (erule allE,erule allE, erule mp)
+        apply (intro strip)
+        apply (tactic "smp_tac 3 1")
+        apply (tactic "smp_tac 2 1")
+        apply (tactic "smp_tac 1 1")
+        by blast
       moreover have "s1\<Colon>\<preceq>(G, L)"
       proof (cases "normal s0")
-	case True
-	from eval wt [OF True] da [OF True] conf wf 
-	show ?thesis
-	  by (rule evaln_type_sound [elim_format]) simp
+        case True
+        from eval wt [OF True] da [OF True] conf wf 
+        show ?thesis
+          by (rule evaln_type_sound [elim_format]) simp
       next
-	case False
-	with eval have "s1=s0"
-	  by auto
-	with conf show ?thesis by simp
+        case False
+        with eval have "s1=s0"
+          by auto
+        with conf show ?thesis by simp
       qed
       ultimately show ?thesis ..
     qed
@@ -461,13 +461,13 @@
     show "P (In2 vf) s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
     proof 
       from eval normal_s0 obtain "s1=s0" "vf=lvar vn (store s0)"
-	by (fastsimp elim: evaln_elim_cases)
+        by (fastsimp elim: evaln_elim_cases)
       with P show "P (In2 vf) s1 Z"
-	by simp
+        by simp
     next
       from eval wt da conf_s0 wf
       show "s1\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
     qed
   qed
 next
@@ -490,78 +490,78 @@
       from wt obtain statC f where
         wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-Class statC" and
         accfield: "accfield G accC statC fn = Some (statDeclC,f)" and
-	eq_accC: "accC=accC'" and
+        eq_accC: "accC=accC'" and
         stat: "stat=is_static f" and
-	T: "T=(type f)"
-	by (cases) (auto simp add: member_is_static_simp)
+        T: "T=(type f)"
+        by (cases) (auto simp add: member_is_static_simp)
       from da eq_accC
       have da_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V"
-	by cases simp
+        by cases simp
       from eval obtain a s1 s2 s2' where
-	eval_init: "G\<turnstile>s0 \<midarrow>Init statDeclC\<midarrow>n\<rightarrow> s1" and 
+        eval_init: "G\<turnstile>s0 \<midarrow>Init statDeclC\<midarrow>n\<rightarrow> s1" and 
         eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s2" and 
-	fvar: "(vf,s2')=fvar statDeclC stat fn a s2" and
-	s3: "s3 = check_field_access G accC statDeclC fn stat a s2'"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases) 
+        fvar: "(vf,s2')=fvar statDeclC stat fn a s2" and
+        s3: "s3 = check_field_access G accC statDeclC fn stat a s2'"
+        using normal_s0 by (fastsimp elim: evaln_elim_cases) 
       have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
       proof -
-	from wf wt_e 
-	have iscls_statC: "is_class G statC"
-	  by (auto dest: ty_expr_is_type type_is_class)
-	with wf accfield 
-	have iscls_statDeclC: "is_class G statDeclC"
-	  by (auto dest!: accfield_fields dest: fields_declC)
-	thus ?thesis by simp
+        from wf wt_e 
+        have iscls_statC: "is_class G statC"
+          by (auto dest: ty_expr_is_type type_is_class)
+        with wf accfield 
+        have iscls_statDeclC: "is_class G statDeclC"
+          by (auto dest!: accfield_fields dest: fields_declC)
+        thus ?thesis by simp
       qed
       obtain I where 
-	da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                     \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init statDeclC\<rangle>\<^sub>s\<guillemotright> I"
-	by (auto intro: da_Init [simplified] assigned.select_convs)
+        by (auto intro: da_Init [simplified] assigned.select_convs)
       from valid_init P valid_A conf_s0 eval_init wt_init da_init
       obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
-	by (rule validE)
+        by (rule validE)
       obtain 
-	R: "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and 
+        R: "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and 
         conf_s2: "s2\<Colon>\<preceq>(G, L)" and
-	conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+        conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
       proof (cases "normal s1")
-	case True
-	obtain V' where 
-	  da_e':
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V'"
-	proof -
-	  from eval_init 
-	  have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
-	    by (rule dom_locals_evaln_mono_elim)
-	  with da_e show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	with valid_e Q valid_A conf_s1 eval_e wt_e
-	obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
-	  by (rule validE) (simp add: fvar [symmetric])
-	moreover
-	from eval_e wt_e da_e' conf_s1 wf
-	have "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
-	  by (rule evaln_type_sound [elim_format]) simp
-	ultimately show ?thesis ..
+        case True
+        obtain V' where 
+          da_e':
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V'"
+        proof -
+          from eval_init 
+          have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
+            by (rule dom_locals_evaln_mono_elim)
+          with da_e show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        with valid_e Q valid_A conf_s1 eval_e wt_e
+        obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
+          by (rule validE) (simp add: fvar [symmetric])
+        moreover
+        from eval_e wt_e da_e' conf_s1 wf
+        have "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+          by (rule evaln_type_sound [elim_format]) simp
+        ultimately show ?thesis ..
       next
-	case False
-	with valid_e Q valid_A conf_s1 eval_e
-	obtain  "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
-	  by (cases rule: validE) (simp add: fvar [symmetric])+
-	moreover from False eval_e have "\<not> normal s2"
-	  by auto
-	hence "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
-	  by auto
-	ultimately show ?thesis ..
+        case False
+        with valid_e Q valid_A conf_s1 eval_e
+        obtain  "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
+          by (cases rule: validE) (simp add: fvar [symmetric])+
+        moreover from False eval_e have "\<not> normal s2"
+          by auto
+        hence "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+          by auto
+        ultimately show ?thesis ..
       qed
       from accfield wt_e eval_init eval_e conf_s2 conf_a fvar stat s3 wf
       have eq_s3_s2': "s3=s2'"  
-	using normal_s0 by (auto dest!: error_free_field_access evaln_eval)
+        using normal_s0 by (auto dest!: error_free_field_access evaln_eval)
       moreover
       from eval wt da conf_s0 wf
       have "s3\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis using Q R by simp
     qed
   qed
@@ -584,48 +584,48 @@
     show "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z \<and> s2'\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain 
-	wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T.[]" and
+        wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T.[]" and
         wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-PrimT Integer" 
-	by (rule wt_elim_cases) simp
+        by (rule wt_elim_cases) simp
       from da obtain E1 where
-	da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
-	da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V"
-	by (rule da_elim_cases) simp
+        da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
+        da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V"
+        by (rule da_elim_cases) simp
       from eval obtain s1 a i s2 where
-	eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1" and
-	eval_e2: "G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2" and
-	avar: "avar G i a s2 =(vf, s2')"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1" and
+        eval_e2: "G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2" and
+        avar: "avar G i a s2 =(vf, s2')"
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_e1 P valid_A conf_s0 eval_e1 wt_e1 da_e1
       obtain Q: "Q \<lfloor>a\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
-	by (rule validE)
+        by (rule validE)
       from Q have Q': "\<And> v. (Q\<leftarrow>In1 a) v s1 Z"
-	by simp
+        by simp
       have "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z"
       proof (cases "normal s1")
-	case True
-	obtain V' where 
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V'"
-	proof -
-	  from eval_e1  wt_e1 da_e1 wf True
-	  have "nrm E1 \<subseteq> dom (locals (store s1))"
-	    by (cases rule: da_good_approx_evalnE) iprover
-	  with da_e2 show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	with valid_e2 Q' valid_A conf_s1 eval_e2 wt_e2 
-	show ?thesis
-	  by (rule validE) (simp add: avar)
+        case True
+        obtain V' where 
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V'"
+        proof -
+          from eval_e1  wt_e1 da_e1 wf True
+          have "nrm E1 \<subseteq> dom (locals (store s1))"
+            by (cases rule: da_good_approx_evalnE) iprover
+          with da_e2 show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        with valid_e2 Q' valid_A conf_s1 eval_e2 wt_e2 
+        show ?thesis
+          by (rule validE) (simp add: avar)
       next
-	case False
-	with valid_e2 Q' valid_A conf_s1 eval_e2
-	show ?thesis
-	  by (cases rule: validE) (simp add: avar)+
+        case False
+        with valid_e2 Q' valid_A conf_s1 eval_e2
+        show ?thesis
+          by (cases rule: validE) (simp add: avar)+
       qed
       moreover
       from eval wt da conf_s0 wf
       have "s2'\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -646,26 +646,26 @@
     show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain is_cls_C: "is_class G C" 
-	by (rule wt_elim_cases) (auto dest: is_acc_classD)
+        by (rule wt_elim_cases) (auto dest: is_acc_classD)
       hence wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>Init C\<Colon>\<surd>" 
-	by auto
+        by auto
       obtain I where 
-	da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init C\<rangle>\<^sub>s\<guillemotright> I"
-	by (auto intro: da_Init [simplified] assigned.select_convs)
+        da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init C\<rangle>\<^sub>s\<guillemotright> I"
+        by (auto intro: da_Init [simplified] assigned.select_convs)
       from eval obtain s1 a where
-	eval_init: "G\<turnstile>s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1" and 
+        eval_init: "G\<turnstile>s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1" and 
         alloc: "G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2" and
-	v: "v=Addr a"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        v: "v=Addr a"
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_init P valid_A conf_s0 eval_init wt_init da_init
       obtain "(Alloc G (CInst C) Q) \<diamondsuit> s1 Z" 
-	by (rule validE)
+        by (rule validE)
       with alloc v have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s2\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -687,58 +687,58 @@
     show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z \<and> s3\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain
-	wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and 
-	wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Integer" 
-	by (rule wt_elim_cases) (auto intro: wt_init_comp_ty )
+        wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and 
+        wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Integer" 
+        by (rule wt_elim_cases) (auto intro: wt_init_comp_ty )
       from da obtain
-	da_e:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
-	by cases simp
+        da_e:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+        by cases simp
       from eval obtain s1 i s2 a where
-	eval_init: "G\<turnstile>s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1" and 
+        eval_init: "G\<turnstile>s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1" and 
         eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>i\<midarrow>n\<rightarrow> s2" and
         alloc: "G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3" and
         v: "v=Addr a"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       obtain I where
-	da_init:
-	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>init_comp_ty T\<rangle>\<^sub>s\<guillemotright> I"
+        da_init:
+        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>init_comp_ty T\<rangle>\<^sub>s\<guillemotright> I"
       proof (cases "\<exists>C. T = Class C")
-	case True
-	thus ?thesis
-	  by - (rule that, (auto intro: da_Init [simplified] 
+        case True
+        thus ?thesis
+          by - (rule that, (auto intro: da_Init [simplified] 
                                         assigned.select_convs
                               simp add: init_comp_ty_def))
-	 (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
+         (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
       next
-	case False
-	thus ?thesis
-	  by - (rule that, (auto intro: da_Skip [simplified] 
+        case False
+        thus ?thesis
+          by - (rule that, (auto intro: da_Skip [simplified] 
                                       assigned.select_convs
                            simp add: init_comp_ty_def))
          (* simplified: to rewrite \<langle>Skip\<rangle> to In1r (Skip) *)
       qed
       with valid_init P valid_A conf_s0 eval_init wt_init 
       obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
-	by (rule validE)
+        by (rule validE)
       obtain E' where
        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
       proof -
-	from eval_init 
-	have "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
-	  by (rule dom_locals_evaln_mono_elim)
-	with da_e show thesis
-	  by (rule da_weakenE) (rule that)
+        from eval_init 
+        have "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+          by (rule dom_locals_evaln_mono_elim)
+        with da_e show thesis
+          by (rule da_weakenE) (rule that)
       qed
       with valid_e Q valid_A conf_s1 eval_e wt_e
       have "(\<lambda>Val:i:. abupd (check_neg i) .; 
                       Alloc G (Arr T (the_Intg i)) R) \<lfloor>i\<rfloor>\<^sub>e s2 Z"
-	by (rule validE)
+        by (rule validE)
       with alloc v have "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
-	by simp
+        by simp
       moreover 
       from eval wt da conf_s0 wf
       have "s3\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -760,25 +760,25 @@
     show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain eT where 
-	wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
-	by cases simp
+        wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
+        by cases simp
       from da obtain
-	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
-	by cases simp
+        da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+        by cases simp
       from eval obtain s1 where
-	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1" and
+        eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1" and
         s2: "s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits T) ClassCast) s1"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_e P valid_A conf_s0 eval_e wt_e da_e
       have "(\<lambda>Val:v:. \<lambda>s.. abupd (raise_if (\<not> G,s\<turnstile>v fits T) ClassCast) .;
                   Q\<leftarrow>In1 v) \<lfloor>v\<rfloor>\<^sub>e s1 Z"
-	by (rule validE)
+        by (rule validE)
       with s2 have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s2\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -799,25 +799,25 @@
     show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain eT where 
-	wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
-	by cases simp
+        wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
+        by cases simp
       from da obtain
-	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
-	by cases simp
+        da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+        by cases simp
       from eval obtain a where
-	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
+        eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
         v: "v = Bool (a \<noteq> Null \<and> G,store s1\<turnstile>a fits RefT T)"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_e P valid_A conf_s0 eval_e wt_e da_e
       have "(\<lambda>Val:v:. \<lambda>s.. Q\<leftarrow>In1 (Bool (v \<noteq> Null \<and> G,s\<turnstile>v fits RefT T))) 
               \<lfloor>a\<rfloor>\<^sub>e s1 Z"
-	by (rule validE)
+        by (rule validE)
       with v have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s1\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -833,7 +833,7 @@
     show "P \<lfloor>v'\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
     proof -
       from eval have "s1=s0" and  "v'=v"
-	using normal_s0 by (auto elim: evaln_elim_cases)
+        using normal_s0 by (auto elim: evaln_elim_cases)
       with P conf_s0 show ?thesis by simp
     qed
   qed
@@ -853,24 +853,24 @@
     show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain eT where 
-	wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
-	by cases simp
+        wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" 
+        by cases simp
       from da obtain
-	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
-	by cases simp
+        da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+        by cases simp
       from eval obtain ve where
-	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
+        eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
         v: "v = eval_unop unop ve"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_e P valid_A conf_s0 eval_e wt_e da_e
       have "(\<lambda>Val:v:. Q\<leftarrow>In1 (eval_unop unop v)) \<lfloor>ve\<rfloor>\<^sub>e s1 Z"
-	by (rule validE)
+        by (rule validE)
       with v have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s1\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -897,63 +897,63 @@
       from wt obtain e1T e2T where
         wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-e1T" and
         wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-e2T" and
-	wt_binop: "wt_binop G binop e1T e2T" 
-	by cases simp
+        wt_binop: "wt_binop G binop e1T e2T" 
+        by cases simp
       have wt_Skip: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>Skip\<Colon>\<surd>"
-	by simp
+        by simp
       (*
       obtain S where
-	daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                    \<turnstile> dom (locals (store s1)) \<guillemotright>In1r Skip\<guillemotright> S"
-	by (auto intro: da_Skip [simplified] assigned.select_convs) *)
+        by (auto intro: da_Skip [simplified] assigned.select_convs) *)
       from da obtain E1 where
-	da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1"
-	by cases simp+
+        da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1"
+        by cases simp+
       from eval obtain v1 s1 v2 where
-	eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<midarrow>n\<rightarrow> s1" and
-	eval_e2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)
+        eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<midarrow>n\<rightarrow> s1" and
+        eval_e2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)
                         \<succ>\<midarrow>n\<rightarrow> (\<lfloor>v2\<rfloor>\<^sub>e, s2)" and
         v: "v=eval_binop binop v1 v2"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_e1 P valid_A conf_s0 eval_e1 wt_e1 da_e1
       obtain Q: "Q \<lfloor>v1\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
-	by (rule validE)
+        by (rule validE)
       from Q have Q': "\<And> v. (Q\<leftarrow>In1 v1) v s1 Z"
-	by simp
+        by simp
       have "(\<lambda>Val:v2:. R\<leftarrow>In1 (eval_binop binop v1 v2)) \<lfloor>v2\<rfloor>\<^sub>e s2 Z"
       proof (cases "normal s1")
-	case True
-	from eval_e1 wt_e1 da_e1 conf_s0 wf
-	have conf_v1: "G,store s1\<turnstile>v1\<Colon>\<preceq>e1T" 
-	  by (rule evaln_type_sound [elim_format]) (insert True,simp)
-	from eval_e1 
-	have "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
-	  by (rule evaln_eval)
-	from da wt_e1 wt_e2 wt_binop conf_s0 True this conf_v1 wf
-	obtain E2 where
-	  da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) 
+        case True
+        from eval_e1 wt_e1 da_e1 conf_s0 wf
+        have conf_v1: "G,store s1\<turnstile>v1\<Colon>\<preceq>e1T" 
+          by (rule evaln_type_sound [elim_format]) (insert True,simp)
+        from eval_e1 
+        have "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
+          by (rule evaln_eval)
+        from da wt_e1 wt_e2 wt_binop conf_s0 True this conf_v1 wf
+        obtain E2 where
+          da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) 
                    \<guillemotright>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<guillemotright> E2"
-	  by (rule da_e2_BinOp [elim_format]) iprover
-	from wt_e2 wt_Skip obtain T2 
-	  where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+          by (rule da_e2_BinOp [elim_format]) iprover
+        from wt_e2 wt_Skip obtain T2 
+          where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                   \<turnstile>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<Colon>T2"
-	  by (cases "need_second_arg binop v1") auto
-	note ve=validE [OF valid_e2,OF  Q' valid_A conf_s1 eval_e2 this da_e2]
-	(* chaining Q', without extra OF causes unification error *)
-	thus ?thesis
-	  by (rule ve)
+          by (cases "need_second_arg binop v1") auto
+        note ve=validE [OF valid_e2,OF  Q' valid_A conf_s1 eval_e2 this da_e2]
+        (* chaining Q', without extra OF causes unification error *)
+        thus ?thesis
+          by (rule ve)
       next
-	case False
-	note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2]
-	with False show ?thesis
-	  by iprover
+        case False
+        note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2]
+        with False show ?thesis
+          by iprover
       qed
       with v have "R \<lfloor>v\<rfloor>\<^sub>e s2 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s2\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -969,7 +969,7 @@
     show "P \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
     proof -
       from eval have "s1=s0" and  "v=val_this (store s0)"
-	using normal_s0 by (auto elim: evaln_elim_cases)
+        using normal_s0 by (auto elim: evaln_elim_cases)
       with P conf_s0 show ?thesis by simp
     qed
   qed
@@ -989,23 +989,23 @@
     show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain 
-	wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=T" 
-	by cases simp
+        wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=T" 
+        by cases simp
       from da obtain V where 
-	da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
-	by (cases "\<exists> n. var=LVar n") (insert da.LVar,auto elim!: da_elim_cases)
+        da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
+        by (cases "\<exists> n. var=LVar n") (insert da.LVar,auto elim!: da_elim_cases)
       from eval obtain w upd where
-	eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(v, upd)\<midarrow>n\<rightarrow> s1"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(v, upd)\<midarrow>n\<rightarrow> s1"
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_var P valid_A conf_s0 eval_var wt_var da_var
       have "(\<lambda>Var:(v, f):. Q\<leftarrow>In1 v) \<lfloor>(v, upd)\<rfloor>\<^sub>v s1 Z"
-	by (rule validE)
+        by (rule validE)
       then have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s1\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -1028,96 +1028,96 @@
     show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z \<and> s3\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain varT  where
-	wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=varT" and
-	wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T" 
-	by cases simp
+        wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=varT" and
+        wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T" 
+        by cases simp
       from eval obtain w upd s1 s2 where
-	eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(w, upd)\<midarrow>n\<rightarrow> s1" and
+        eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(w, upd)\<midarrow>n\<rightarrow> s1" and
         eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s2" and
-	s3: "s3=assign upd v s2"
-	using normal_s0 by (auto elim: evaln_elim_cases)
+        s3: "s3=assign upd v s2"
+        using normal_s0 by (auto elim: evaln_elim_cases)
       have "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
       proof (cases "\<exists> vn. var = LVar vn")
-	case False
-	with da obtain V where
-	  da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        case False
+        with da obtain V where
+          da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                       \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V" and
-	  da_e:   "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
-	  by cases simp+
-	from valid_var P valid_A conf_s0 eval_var wt_var da_var
-	obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
-	  by (rule validE) 
-	hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
-	  by simp
-	have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
-	proof (cases "normal s1")
-	  case True
-	  obtain E' where 
-	    da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
-	  proof -
-	    from eval_var wt_var da_var wf True
-	    have "nrm V \<subseteq>  dom (locals (store s1))"
-	      by (cases rule: da_good_approx_evalnE) iprover
-	    with da_e show thesis
-	      by (rule da_weakenE) (rule that)
-	  qed
-	  note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
-	  show ?thesis
-	    by (rule ve)
-	next
-	  case False
-	  note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
-	  with False show ?thesis
-	    by iprover
-	qed
-	with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
-	  by simp
+          da_e:   "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+          by cases simp+
+        from valid_var P valid_A conf_s0 eval_var wt_var da_var
+        obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
+          by (rule validE) 
+        hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
+          by simp
+        have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+        proof (cases "normal s1")
+          case True
+          obtain E' where 
+            da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
+          proof -
+            from eval_var wt_var da_var wf True
+            have "nrm V \<subseteq>  dom (locals (store s1))"
+              by (cases rule: da_good_approx_evalnE) iprover
+            with da_e show thesis
+              by (rule da_weakenE) (rule that)
+          qed
+          note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
+          show ?thesis
+            by (rule ve)
+        next
+          case False
+          note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
+          with False show ?thesis
+            by iprover
+        qed
+        with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
+          by simp
       next
-	case True
-	then obtain vn where 
-	  vn: "var = LVar vn" 
-	  by auto
-	with da obtain E where
-	    da_e:   "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
-	  by cases simp+
-	from da.LVar vn obtain  V where
-	  da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        case True
+        then obtain vn where 
+          vn: "var = LVar vn" 
+          by auto
+        with da obtain E where
+            da_e:   "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+          by cases simp+
+        from da.LVar vn obtain  V where
+          da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                       \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
-	  by auto
-	from valid_var P valid_A conf_s0 eval_var wt_var da_var
-	obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
-	  by (rule validE) 
-	hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
-	  by simp
-	have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
-	proof (cases "normal s1")
-	  case True
-	  obtain E' where
-	    da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+          by auto
+        from valid_var P valid_A conf_s0 eval_var wt_var da_var
+        obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
+          by (rule validE) 
+        hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
+          by simp
+        have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+        proof (cases "normal s1")
+          case True
+          obtain E' where
+            da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                        \<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
-	  proof -
-	    from eval_var
-	    have "dom (locals (store s0)) \<subseteq> dom (locals (store (s1)))"
-	      by (rule dom_locals_evaln_mono_elim)
-	    with da_e show thesis
-	      by (rule da_weakenE) (rule that)
-	  qed
-	  note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
-	  show ?thesis
-	    by (rule ve)
-	next
-	  case False
-	  note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
-	  with False show ?thesis
-	    by iprover
-	qed
-	with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
-	  by simp
+          proof -
+            from eval_var
+            have "dom (locals (store s0)) \<subseteq> dom (locals (store (s1)))"
+              by (rule dom_locals_evaln_mono_elim)
+            with da_e show thesis
+              by (rule da_weakenE) (rule that)
+          qed
+          note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
+          show ?thesis
+            by (rule ve)
+        next
+          case False
+          note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
+          with False show ?thesis
+            by iprover
+        qed
+        with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
+          by simp
       qed
       moreover
       from eval wt da conf_s0 wf
       have "s3\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -1139,75 +1139,75 @@
     show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain T1 T2 where
-	wt_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
-	wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T1" and
-	wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-T2" 
-	by cases simp
+        wt_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
+        wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T1" and
+        wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-T2" 
+        by cases simp
       from da obtain E0 E1 E2 where
         da_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e0\<rangle>\<^sub>e\<guillemotright> E0" and
         da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                  \<turnstile>(dom (locals (store s0)) \<union> assigns_if True e0)\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
         da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                  \<turnstile>(dom (locals (store s0)) \<union> assigns_if False e0)\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> E2"
-	by cases simp+
+        by cases simp+
       from eval obtain b s1 where
-	eval_e0: "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1" and
+        eval_e0: "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1" and
         eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<midarrow>n\<rightarrow> s2"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_e0 P valid_A conf_s0 eval_e0 wt_e0 da_e0
       obtain "P' \<lfloor>b\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
-	by (rule validE)
+        by (rule validE)
       hence P': "\<And> v. (P'\<leftarrow>=(the_Bool b)) v s1 Z"
-	by (cases "normal s1") auto
+        by (cases "normal s1") auto
       have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
       proof (cases "normal s1")
-	case True
-	note normal_s1=this
-	from wt_e1 wt_e2 obtain T' where
-	  wt_then_else: 
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then e1 else e2)\<Colon>-T'"
-	  by (cases "the_Bool b") simp+
-	have s0_s1: "dom (locals (store s0)) 
+        case True
+        note normal_s1=this
+        from wt_e1 wt_e2 obtain T' where
+          wt_then_else: 
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then e1 else e2)\<Colon>-T'"
+          by (cases "the_Bool b") simp+
+        have s0_s1: "dom (locals (store s0)) 
                       \<union> assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
-	proof -
-	  from eval_e0 
-	  have eval_e0': "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<rightarrow> s1"
-	    by (rule evaln_eval)
-	  hence
-	    "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
-	    by (rule dom_locals_eval_mono_elim)
+        proof -
+          from eval_e0 
+          have eval_e0': "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<rightarrow> s1"
+            by (rule evaln_eval)
+          hence
+            "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+            by (rule dom_locals_eval_mono_elim)
           moreover
-	  from eval_e0' True wt_e0 
-	  have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
-	    by (rule assigns_if_good_approx') 
-	  ultimately show ?thesis by (rule Un_least)
-	qed
-	obtain E' where
-	  da_then_else:
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+          from eval_e0' True wt_e0 
+          have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
+            by (rule assigns_if_good_approx') 
+          ultimately show ?thesis by (rule Un_least)
+        qed
+        obtain E' where
+          da_then_else:
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
               \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>if the_Bool b then e1 else e2\<rangle>\<^sub>e\<guillemotright> E'"
-	proof (cases "the_Bool b")
-	  case True
-	  with that da_e1 s0_s1 show ?thesis
-	    by simp (erule da_weakenE,auto)
-	next
-	  case False
-	  with that da_e2 s0_s1 show ?thesis
-	    by simp (erule da_weakenE,auto)
-	qed
-	with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
-	show ?thesis
-	  by (rule validE)
+        proof (cases "the_Bool b")
+          case True
+          with that da_e1 s0_s1 show ?thesis
+            by simp (erule da_weakenE,auto)
+        next
+          case False
+          with that da_e2 s0_s1 show ?thesis
+            by simp (erule da_weakenE,auto)
+        qed
+        with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
+        show ?thesis
+          by (rule validE)
       next
-	case False
-	with valid_then_else P' valid_A conf_s1 eval_then_else
-	show ?thesis
-	  by (cases rule: validE) iprover+
+        case False
+        with valid_then_else P' valid_A conf_s1 eval_then_else
+        show ?thesis
+          by (cases rule: validE) iprover+
       qed
       moreover
       from eval wt da conf_s0 wf
       have "s2\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -1248,93 +1248,93 @@
                  mode: "mode = invmode statM e" and
                     T: "T =(resTy statM)" and
         eq_accC_accC': "accC=accC'"
-	by cases fastsimp+
+        by cases fastsimp+
       from da obtain C where
-	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
-	da_args: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E" 
-	by cases simp
+        da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
+        da_args: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E" 
+        by cases simp
       from eval eq_accC_accC' obtain a s1 vs s2 s3 s3' s4 invDeclC where
-	evaln_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
+        evaln_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
         evaln_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
-	invDeclC: "invDeclC = invocation_declclass 
+        invDeclC: "invDeclC = invocation_declclass 
                 G mode (store s2) a statT \<lparr>name=mn,parTs=pTs'\<rparr>" and
         s3: "s3 = init_lvars G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> mode a vs s2" and
         check: "s3' = check_method_access G 
                            accC' statT mode \<lparr>name = mn, parTs = pTs'\<rparr> a s3" and
-	evaln_methd:
+        evaln_methd:
            "G\<turnstile>s3' \<midarrow>Methd invDeclC  \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<midarrow>n\<rightarrow> s4" and
         s5: "s5=(set_lvars (locals (store s2))) s4"
-	using normal_s0 by (auto elim: evaln_elim_cases)
+        using normal_s0 by (auto elim: evaln_elim_cases)
 
       from evaln_e
       have eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
-	by (rule evaln_eval)
+        by (rule evaln_eval)
       
       from eval_e _ wt_e wf
       have s1_no_return: "abrupt s1 \<noteq> Some (Jump Ret)"
-	by (rule eval_expression_no_jump 
+        by (rule eval_expression_no_jump 
                  [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
-	   (insert normal_s0,auto)
+           (insert normal_s0,auto)
 
       from valid_e P valid_A conf_s0 evaln_e wt_e da_e
       obtain "Q \<lfloor>a\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
-	by (rule validE)
+        by (rule validE)
       hence Q: "\<And> v. (Q\<leftarrow>In1 a) v s1 Z"
-	by simp
+        by simp
       obtain 
-	R: "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" and 
-	conf_s2: "s2\<Colon>\<preceq>(G,L)" and 
-	s2_no_return: "abrupt s2 \<noteq> Some (Jump Ret)"
+        R: "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" and 
+        conf_s2: "s2\<Colon>\<preceq>(G,L)" and 
+        s2_no_return: "abrupt s2 \<noteq> Some (Jump Ret)"
       proof (cases "normal s1")
-	case True
-	obtain E' where 
-	  da_args':
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
-	proof -
-	  from evaln_e wt_e da_e wf True
-	  have "nrm C \<subseteq>  dom (locals (store s1))"
-	    by (cases rule: da_good_approx_evalnE) iprover
-	  with da_args show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	with valid_args Q valid_A conf_s1 evaln_args wt_args 
-	obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)" 
-	  by (rule validE)
-	moreover
-	from evaln_args
-	have e: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
-	  by (rule evaln_eval)
-	from this s1_no_return wt_args wf
-	have "abrupt s2 \<noteq> Some (Jump Ret)"
-	  by (rule eval_expression_list_no_jump 
+        case True
+        obtain E' where 
+          da_args':
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
+        proof -
+          from evaln_e wt_e da_e wf True
+          have "nrm C \<subseteq>  dom (locals (store s1))"
+            by (cases rule: da_good_approx_evalnE) iprover
+          with da_args show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        with valid_args Q valid_A conf_s1 evaln_args wt_args 
+        obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)" 
+          by (rule validE)
+        moreover
+        from evaln_args
+        have e: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
+          by (rule evaln_eval)
+        from this s1_no_return wt_args wf
+        have "abrupt s2 \<noteq> Some (Jump Ret)"
+          by (rule eval_expression_list_no_jump 
                  [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
-	ultimately show ?thesis ..
+        ultimately show ?thesis ..
       next
-	case False
-	with valid_args Q valid_A conf_s1 evaln_args
-	obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)" 
-	  by (cases rule: validE) iprover+
-	moreover
-	from False evaln_args have "s2=s1"
-	  by auto
-	with s1_no_return have "abrupt s2 \<noteq> Some (Jump Ret)"
-	  by simp
-	ultimately show ?thesis ..
+        case False
+        with valid_args Q valid_A conf_s1 evaln_args
+        obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)" 
+          by (cases rule: validE) iprover+
+        moreover
+        from False evaln_args have "s2=s1"
+          by auto
+        with s1_no_return have "abrupt s2 \<noteq> Some (Jump Ret)"
+          by simp
+        ultimately show ?thesis ..
       qed
 
       obtain invC where
-	invC: "invC = invocation_class mode (store s2) a statT"
-	by simp
+        invC: "invC = invocation_class mode (store s2) a statT"
+        by simp
       with s3
       have invC': "invC = (invocation_class mode (store s3) a statT)"
-	by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
+        by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
       obtain l where
-	l: "l = locals (store s2)"
-	by simp
+        l: "l = locals (store s2)"
+        by simp
 
       from eval wt da conf_s0 wf
       have conf_s5: "s5\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       let "PROP ?R" = "\<And> v.
              (R a\<leftarrow>In3 vs \<and>.
                  (\<lambda>s. invDeclC = invocation_declclass G mode (store s) a statT
@@ -1345,128 +1345,128 @@
                   (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)
                ) v s3' Z"
       {
-	assume abrupt_s3: "\<not> normal s3"
-	have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
-	proof -
-	  from abrupt_s3 check have eq_s3'_s3: "s3'=s3"
-	    by (auto simp add: check_method_access_def Let_def)
-	  with R s3 invDeclC invC l abrupt_s3
-	  have R': "PROP ?R"
-	    by auto
-	  have conf_s3': "s3'\<Colon>\<preceq>(G, empty)"
-	   (* we need an arbirary environment (here empty) that s2' conforms to
+        assume abrupt_s3: "\<not> normal s3"
+        have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
+        proof -
+          from abrupt_s3 check have eq_s3'_s3: "s3'=s3"
+            by (auto simp add: check_method_access_def Let_def)
+          with R s3 invDeclC invC l abrupt_s3
+          have R': "PROP ?R"
+            by auto
+          have conf_s3': "s3'\<Colon>\<preceq>(G, empty)"
+           (* we need an arbirary environment (here empty) that s2' conforms to
               to apply validE *)
-	  proof -
-	    from s2_no_return s3
-	    have "abrupt s3 \<noteq> Some (Jump Ret)"
-	      by (cases s2) (auto simp add: init_lvars_def2 split: split_if_asm)
-	    moreover
-	    obtain abr2 str2 where s2: "s2=(abr2,str2)"
-	      by (cases s2)
-	    from s3 s2 conf_s2 have "(abrupt s3,str2)\<Colon>\<preceq>(G, L)"
-	      by (auto simp add: init_lvars_def2 split: split_if_asm)
-	    ultimately show ?thesis
-	      using s3 s2 eq_s3'_s3
-	      apply (simp add: init_lvars_def2)
-	      apply (rule conforms_set_locals [OF _ wlconf_empty])
-	      by auto
-	  qed
-	  from valid_methd R' valid_A conf_s3' evaln_methd abrupt_s3 eq_s3'_s3
-	  have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
-	    by (cases rule: validE) simp+
-	  with s5 l show ?thesis
-	    by simp
-	qed
+          proof -
+            from s2_no_return s3
+            have "abrupt s3 \<noteq> Some (Jump Ret)"
+              by (cases s2) (auto simp add: init_lvars_def2 split: split_if_asm)
+            moreover
+            obtain abr2 str2 where s2: "s2=(abr2,str2)"
+              by (cases s2)
+            from s3 s2 conf_s2 have "(abrupt s3,str2)\<Colon>\<preceq>(G, L)"
+              by (auto simp add: init_lvars_def2 split: split_if_asm)
+            ultimately show ?thesis
+              using s3 s2 eq_s3'_s3
+              apply (simp add: init_lvars_def2)
+              apply (rule conforms_set_locals [OF _ wlconf_empty])
+              by auto
+          qed
+          from valid_methd R' valid_A conf_s3' evaln_methd abrupt_s3 eq_s3'_s3
+          have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
+            by (cases rule: validE) simp+
+          with s5 l show ?thesis
+            by simp
+        qed
       } note abrupt_s3_lemma = this
 
       have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
       proof (cases "normal s2")
-	case False
-	with s3 have abrupt_s3: "\<not> normal s3"
-	  by (cases s2) (simp add: init_lvars_def2)
-	thus ?thesis
-	  by (rule abrupt_s3_lemma)
+        case False
+        with s3 have abrupt_s3: "\<not> normal s3"
+          by (cases s2) (simp add: init_lvars_def2)
+        thus ?thesis
+          by (rule abrupt_s3_lemma)
       next
-	case True
-	note normal_s2 = this
-	with evaln_args 
-	have normal_s1: "normal s1"
-	  by (rule evaln_no_abrupt)
-	obtain E' where 
-	  da_args':
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
-	proof -
-	  from evaln_e wt_e da_e wf normal_s1
-	  have "nrm C \<subseteq>  dom (locals (store s1))"
-	    by (cases rule: da_good_approx_evalnE) iprover
-	  with da_args show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	from evaln_args
-	have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
-	  by (rule evaln_eval)
-	from evaln_e wt_e da_e conf_s0 wf
-	have conf_a: "G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
-	  by (rule evaln_type_sound [elim_format]) (insert normal_s1,simp)
-	with normal_s1 normal_s2 eval_args 
-	have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
-	  by (auto dest: eval_gext intro: conf_gext)
-	from evaln_args wt_args da_args' conf_s1 wf
-	have conf_args: "list_all2 (conf G (store s2)) vs pTs"
-	  by (rule evaln_type_sound [elim_format]) (insert normal_s2,simp)
-	from statM 
-	obtain
-	  statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>" 
-	  and
-	  pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
-	  by (blast dest: max_spec2mheads)
-	show ?thesis
-	proof (cases "normal s3")
-	  case False
-	  thus ?thesis
-	    by (rule abrupt_s3_lemma)
-	next
-	  case True
-	  note normal_s3 = this
-	  with s3 have notNull: "mode = IntVir \<longrightarrow> a \<noteq> Null"
-	    by (cases s2) (auto simp add: init_lvars_def2)
-	  from conf_s2 conf_a_s2 wf notNull invC
-	  have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
-	    by (cases s2) (auto intro: DynT_propI)
+        case True
+        note normal_s2 = this
+        with evaln_args 
+        have normal_s1: "normal s1"
+          by (rule evaln_no_abrupt)
+        obtain E' where 
+          da_args':
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
+        proof -
+          from evaln_e wt_e da_e wf normal_s1
+          have "nrm C \<subseteq>  dom (locals (store s1))"
+            by (cases rule: da_good_approx_evalnE) iprover
+          with da_args show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        from evaln_args
+        have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
+          by (rule evaln_eval)
+        from evaln_e wt_e da_e conf_s0 wf
+        have conf_a: "G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+          by (rule evaln_type_sound [elim_format]) (insert normal_s1,simp)
+        with normal_s1 normal_s2 eval_args 
+        have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
+          by (auto dest: eval_gext intro: conf_gext)
+        from evaln_args wt_args da_args' conf_s1 wf
+        have conf_args: "list_all2 (conf G (store s2)) vs pTs"
+          by (rule evaln_type_sound [elim_format]) (insert normal_s2,simp)
+        from statM 
+        obtain
+          statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>" 
+          and
+          pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
+          by (blast dest: max_spec2mheads)
+        show ?thesis
+        proof (cases "normal s3")
+          case False
+          thus ?thesis
+            by (rule abrupt_s3_lemma)
+        next
+          case True
+          note normal_s3 = this
+          with s3 have notNull: "mode = IntVir \<longrightarrow> a \<noteq> Null"
+            by (cases s2) (auto simp add: init_lvars_def2)
+          from conf_s2 conf_a_s2 wf notNull invC
+          have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
+            by (cases s2) (auto intro: DynT_propI)
 
-	  with wt_e statM' invC mode wf 
-	  obtain dynM where 
+          with wt_e statM' invC mode wf 
+          obtain dynM where 
             dynM: "dynlookup G statT invC  \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
             acc_dynM: "G \<turnstile>Methd  \<lparr>name=mn,parTs=pTs'\<rparr> dynM 
                             in invC dyn_accessible_from accC"
-	    by (force dest!: call_access_ok)
-	  with invC' check eq_accC_accC'
-	  have eq_s3'_s3: "s3'=s3"
-	    by (auto simp add: check_method_access_def Let_def)
-	  
-	  with dynT_prop R s3 invDeclC invC l 
-	  have R': "PROP ?R"
-	    by auto
+            by (force dest!: call_access_ok)
+          with invC' check eq_accC_accC'
+          have eq_s3'_s3: "s3'=s3"
+            by (auto simp add: check_method_access_def Let_def)
+          
+          with dynT_prop R s3 invDeclC invC l 
+          have R': "PROP ?R"
+            by auto
 
-	  from dynT_prop wf wt_e statM' mode invC invDeclC dynM
-	  obtain 
+          from dynT_prop wf wt_e statM' mode invC invDeclC dynM
+          obtain 
             dynM: "dynlookup G statT invC  \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
-	    wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
-	      dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+            wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
+              dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
             iscls_invDeclC: "is_class G invDeclC" and
-	         invDeclC': "invDeclC = declclass dynM" and
-	      invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
-	     resTy_widen: "G\<turnstile>resTy dynM\<preceq>resTy statM" and
-	    is_static_eq: "is_static dynM = is_static statM" and
-	    involved_classes_prop:
+                 invDeclC': "invDeclC = declclass dynM" and
+              invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
+             resTy_widen: "G\<turnstile>resTy dynM\<preceq>resTy statM" and
+            is_static_eq: "is_static dynM = is_static statM" and
+            involved_classes_prop:
              "(if invmode statM e = IntVir
                then \<forall>statC. statT = ClassT statC \<longrightarrow> G\<turnstile>invC\<preceq>\<^sub>C statC
                else ((\<exists>statC. statT = ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C invDeclC) \<or>
                      (\<forall>statC. statT \<noteq> ClassT statC \<and> invDeclC = Object)) \<and>
                       statDeclT = ClassT invDeclC)"
-	    by (cases rule: DynT_mheadsE) simp
-	  obtain L' where 
-	    L':"L'=(\<lambda> k. 
+            by (cases rule: DynT_mheadsE) simp
+          obtain L' where 
+            L':"L'=(\<lambda> k. 
                     (case k of
                        EName e
                        \<Rightarrow> (case e of 
@@ -1476,121 +1476,121 @@
                            | Res \<Rightarrow> Some (resTy dynM))
                      | This \<Rightarrow> if is_static statM 
                                then None else Some (Class invDeclC)))"
-	    by simp
-	  from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
+            by simp
+          from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
             wf eval_args conf_a mode notNull wf_dynM involved_classes_prop
-	  have conf_s3: "s3\<Colon>\<preceq>(G,L')"
-	    apply - 
+          have conf_s3: "s3\<Colon>\<preceq>(G,L')"
+            apply - 
                (* FIXME confomrs_init_lvars should be 
                   adjusted to be more directy applicable *)
-	    apply (drule conforms_init_lvars [of G invDeclC 
+            apply (drule conforms_init_lvars [of G invDeclC 
                     "\<lparr>name=mn,parTs=pTs'\<rparr>" dynM "store s2" vs pTs "abrupt s2" 
                     L statT invC a "(statDeclT,statM)" e])
-	    apply (rule wf)
-	    apply (rule conf_args)
-	    apply (simp add: pTs_widen)
-	    apply (cases s2,simp)
-	    apply (rule dynM')
-	    apply (force dest: ty_expr_is_type)
-	    apply (rule invC_widen)
-	    apply (force intro: conf_gext dest: eval_gext)
-	    apply simp
-	    apply simp
-	    apply (simp add: invC)
-	    apply (simp add: invDeclC)
-	    apply (simp add: normal_s2)
-	    apply (cases s2, simp add: L' init_lvars_def2 s3
-	                     cong add: lname.case_cong ename.case_cong)
-	    done
-	  with eq_s3'_s3 have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
-	  from is_static_eq wf_dynM L'
-	  obtain mthdT where
-	    "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+            apply (rule wf)
+            apply (rule conf_args)
+            apply (simp add: pTs_widen)
+            apply (cases s2,simp)
+            apply (rule dynM')
+            apply (force dest: ty_expr_is_type)
+            apply (rule invC_widen)
+            apply (force intro: conf_gext dest: eval_gext)
+            apply simp
+            apply simp
+            apply (simp add: invC)
+            apply (simp add: invDeclC)
+            apply (simp add: normal_s2)
+            apply (cases s2, simp add: L' init_lvars_def2 s3
+                             cong add: lname.case_cong ename.case_cong)
+            done
+          with eq_s3'_s3 have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
+          from is_static_eq wf_dynM L'
+          obtain mthdT where
+            "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
                \<turnstile>Body invDeclC (stmt (mbody (mthd dynM)))\<Colon>-mthdT" and
-	    mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
-	    by - (drule wf_mdecl_bodyD,
+            mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
+            by - (drule wf_mdecl_bodyD,
                   auto simp add: callee_lcl_def  
                        cong add: lname.case_cong ename.case_cong)
-	  with dynM' iscls_invDeclC invDeclC'
-	  have
-	    wt_methd:
-	    "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+          with dynM' iscls_invDeclC invDeclC'
+          have
+            wt_methd:
+            "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
                \<turnstile>(Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<Colon>-mthdT"
-	    by (auto intro: wt.Methd)
-	  obtain M where 
-	    da_methd:
-	    "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr> 
-	       \<turnstile> dom (locals (store s3')) 
+            by (auto intro: wt.Methd)
+          obtain M where 
+            da_methd:
+            "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr> 
+               \<turnstile> dom (locals (store s3')) 
                    \<guillemotright>\<langle>Methd invDeclC \<lparr>name=mn,parTs=pTs'\<rparr>\<rangle>\<^sub>e\<guillemotright> M"
-	  proof -
-	    from wf_dynM
-	    obtain M' where
-	      da_body: 
-	      "\<lparr>prg=G, cls=invDeclC
+          proof -
+            from wf_dynM
+            obtain M' where
+              da_body: 
+              "\<lparr>prg=G, cls=invDeclC
                ,lcl=callee_lcl invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> (mthd dynM)
                \<rparr> \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'" and
               res: "Result \<in> nrm M'"
-	      by (rule wf_mdeclE) iprover
-	    from da_body is_static_eq L' have
-	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
+              by (rule wf_mdeclE) iprover
+            from da_body is_static_eq L' have
+              "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
                  \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'"
-	      by (simp add: callee_lcl_def  
+              by (simp add: callee_lcl_def  
                   cong add: lname.case_cong ename.case_cong)
-	    moreover have "parameters (mthd dynM) \<subseteq>  dom (locals (store s3'))"
-	    proof -
-	      from is_static_eq 
-	      have "(invmode (mthd dynM) e) = (invmode statM e)"
-		by (simp add: invmode_def)
-	      moreover
-	      have "length (pars (mthd dynM)) = length vs" 
-	      proof -
-		from normal_s2 conf_args
-		have "length vs = length pTs"
-		  by (simp add: list_all2_def)
-		also from pTs_widen
-		have "\<dots> = length pTs'"
-		  by (simp add: widens_def list_all2_def)
-		also from wf_dynM
-		have "\<dots> = length (pars (mthd dynM))"
-		  by (simp add: wf_mdecl_def wf_mhead_def)
-		finally show ?thesis ..
-	      qed
-	      moreover note s3 dynM' is_static_eq normal_s2 mode 
-	      ultimately
-	      have "parameters (mthd dynM) = dom (locals (store s3))"
-		using dom_locals_init_lvars 
+            moreover have "parameters (mthd dynM) \<subseteq>  dom (locals (store s3'))"
+            proof -
+              from is_static_eq 
+              have "(invmode (mthd dynM) e) = (invmode statM e)"
+                by (simp add: invmode_def)
+              moreover
+              have "length (pars (mthd dynM)) = length vs" 
+              proof -
+                from normal_s2 conf_args
+                have "length vs = length pTs"
+                  by (simp add: list_all2_def)
+                also from pTs_widen
+                have "\<dots> = length pTs'"
+                  by (simp add: widens_def list_all2_def)
+                also from wf_dynM
+                have "\<dots> = length (pars (mthd dynM))"
+                  by (simp add: wf_mdecl_def wf_mhead_def)
+                finally show ?thesis ..
+              qed
+              moreover note s3 dynM' is_static_eq normal_s2 mode 
+              ultimately
+              have "parameters (mthd dynM) = dom (locals (store s3))"
+                using dom_locals_init_lvars 
                   [of "mthd dynM" G invDeclC "\<lparr>name=mn,parTs=pTs'\<rparr>" vs e a s2]
-		by simp
-	      thus ?thesis using eq_s3'_s3 by simp
-	    qed
-	    ultimately obtain M2 where
-	      da:
-	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
+                by simp
+              thus ?thesis using eq_s3'_s3 by simp
+            qed
+            ultimately obtain M2 where
+              da:
+              "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
                 \<turnstile> dom (locals (store s3')) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M2" and
               M2: "nrm M' \<subseteq> nrm M2"
-	      by (rule da_weakenE)
-	    from res M2 have "Result \<in> nrm M2"
-	      by blast
-	    moreover from wf_dynM
-	    have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
-	      by (rule wf_mdeclE)
-	    ultimately
-	    obtain M3 where
-	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3')) 
+              by (rule da_weakenE)
+            from res M2 have "Result \<in> nrm M2"
+              by blast
+            moreover from wf_dynM
+            have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
+              by (rule wf_mdeclE)
+            ultimately
+            obtain M3 where
+              "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3')) 
                      \<guillemotright>\<langle>Body (declclass dynM) (stmt (mbody (mthd dynM)))\<rangle>\<guillemotright> M3"
-	      using da
-	      by (iprover intro: da.Body assigned.select_convs)
-	    from _ this [simplified]
-	    show thesis
-	      by (rule da.Methd [simplified,elim_format])
-	         (auto intro: dynM' that)
-	  qed
-	  from valid_methd R' valid_A conf_s3' evaln_methd wt_methd da_methd
-	  have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
-	    by (cases rule: validE) iprover+
-	  with s5 l show ?thesis
-	    by simp
-	qed
+              using da
+              by (iprover intro: da.Body assigned.select_convs)
+            from _ this [simplified]
+            show thesis
+              by (rule da.Methd [simplified,elim_format])
+                 (auto intro: dynM' that)
+          qed
+          from valid_methd R' valid_A conf_s3' evaln_methd wt_methd da_methd
+          have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
+            by (cases rule: validE) iprover+
+          with s5 l show ?thesis
+            by simp
+        qed
       qed
       with conf_s5 show ?thesis by iprover
     qed
@@ -1618,61 +1618,61 @@
     show "R \<lfloor>v\<rfloor>\<^sub>e s4 Z \<and> s4\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain 
-	iscls_D: "is_class G D" and
+        iscls_D: "is_class G D" and
         wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Init D\<Colon>\<surd>" and
         wt_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>" 
-	by cases auto
+        by cases auto
       obtain I where 
-	da_init:"\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init D\<rangle>\<^sub>s\<guillemotright> I"
-	by (auto intro: da_Init [simplified] assigned.select_convs)
+        da_init:"\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init D\<rangle>\<^sub>s\<guillemotright> I"
+        by (auto intro: da_Init [simplified] assigned.select_convs)
       from da obtain C where
-	da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C" and
-	jmpOk: "jumpNestingOkS {Ret} c" 
-	by cases simp
+        da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C" and
+        jmpOk: "jumpNestingOkS {Ret} c" 
+        by cases simp
       from eval obtain s1 s2 s3 where
-	eval_init: "G\<turnstile>s0 \<midarrow>Init D\<midarrow>n\<rightarrow> s1" and
+        eval_init: "G\<turnstile>s0 \<midarrow>Init D\<midarrow>n\<rightarrow> s1" and
         eval_c: "G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2" and
-	v: "v = the (locals (store s2) Result)" and
+        v: "v = the (locals (store s2) Result)" and
         s3: "s3 =(if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or> 
                          abrupt s2 = Some (Jump (Cont l))
                   then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)"and
         s4: "s4 = abupd (absorb Ret) s3"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       obtain C' where 
-	da_c': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C'"
+        da_c': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C'"
       proof -
-	from eval_init 
-	have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
-	  by (rule dom_locals_evaln_mono_elim)
-	with da_c show thesis by (rule da_weakenE) (rule that)
+        from eval_init 
+        have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
+          by (rule dom_locals_evaln_mono_elim)
+        with da_c show thesis by (rule da_weakenE) (rule that)
       qed
       from valid_init P valid_A conf_s0 eval_init wt_init da_init
       obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
-	by (rule validE)
+        by (rule validE)
       from valid_c Q valid_A conf_s1 eval_c wt_c da_c' 
       have R: "(\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>In1 (the (locals s Result))) 
                 \<diamondsuit> s2 Z"
-	by (rule validE)
+        by (rule validE)
       have "s3=s2"
       proof -
-	from eval_init [THEN evaln_eval] wf
-	have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
-	  by - (rule eval_statement_no_jump [OF _ _ _ wt_init],
+        from eval_init [THEN evaln_eval] wf
+        have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+          by - (rule eval_statement_no_jump [OF _ _ _ wt_init],
                 insert normal_s0,auto)
-	from eval_c [THEN evaln_eval] _ wt_c wf
-	have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
-	  by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
-	moreover note s3
-	ultimately show ?thesis 
-	  by (force split: split_if)
+        from eval_c [THEN evaln_eval] _ wt_c wf
+        have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
+          by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
+        moreover note s3
+        ultimately show ?thesis 
+          by (force split: split_if)
       qed
       with R v s4 
       have "R \<lfloor>v\<rfloor>\<^sub>e s4 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s4\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -1688,9 +1688,9 @@
     show "P \<lfloor>vs\<rfloor>\<^sub>l s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
     proof -
       from eval obtain "vs=[]" "s1=s0"
-	using normal_s0 by (auto elim: evaln_elim_cases)
+        using normal_s0 by (auto elim: evaln_elim_cases)
       with P conf_s0 show ?thesis
-	by simp
+        by simp
     qed
   qed
 next
@@ -1711,50 +1711,50 @@
     show "R \<lfloor>v\<rfloor>\<^sub>l s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain eT esT where
-	wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-eT" and
-	wt_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>es\<Colon>\<doteq>esT"
-	by cases simp
+        wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-eT" and
+        wt_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>es\<Colon>\<doteq>esT"
+        by cases simp
       from da obtain E1 where
-	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E1" and
-	da_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E" 
-	by cases simp
+        da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E1" and
+        da_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E" 
+        by cases simp
       from eval obtain s1 ve vs where
-	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
-	eval_es: "G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
-	v: "v=ve#vs"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
+        eval_es: "G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
+        v: "v=ve#vs"
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_e P valid_A conf_s0 eval_e wt_e da_e 
       obtain Q: "Q \<lfloor>ve\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
-	by (rule validE)
+        by (rule validE)
       from Q have Q': "\<And> v. (Q\<leftarrow>\<lfloor>ve\<rfloor>\<^sub>e) v s1 Z"
-	by simp
+        by simp
       have "(\<lambda>Vals:vs:. R\<leftarrow>\<lfloor>(ve # vs)\<rfloor>\<^sub>l) \<lfloor>vs\<rfloor>\<^sub>l s2 Z"
       proof (cases "normal s1")
-	case True
-	obtain E' where 
-	  da_es': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E'"
-	proof -
-	  from eval_e wt_e da_e wf True
-	  have "nrm E1 \<subseteq> dom (locals (store s1))"
-	    by (cases rule: da_good_approx_evalnE) iprover
-	  with da_es show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	from valid_es Q' valid_A conf_s1 eval_es wt_es da_es'
-	show ?thesis
-	  by (rule validE)
+        case True
+        obtain E' where 
+          da_es': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E'"
+        proof -
+          from eval_e wt_e da_e wf True
+          have "nrm E1 \<subseteq> dom (locals (store s1))"
+            by (cases rule: da_good_approx_evalnE) iprover
+          with da_es show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        from valid_es Q' valid_A conf_s1 eval_es wt_es da_es'
+        show ?thesis
+          by (rule validE)
       next
-	case False
-	with valid_es Q' valid_A conf_s1 eval_es 
-	show ?thesis
-	  by (cases rule: validE) iprover+
+        case False
+        with valid_es Q' valid_A conf_s1 eval_es 
+        show ?thesis
+          by (cases rule: validE) iprover+
       qed
       with v have "R \<lfloor>v\<rfloor>\<^sub>l s2 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s2\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -1770,9 +1770,9 @@
     show "P \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
     proof -
       from eval obtain "s1=s0"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       with P conf_s0 show ?thesis
-	by simp
+        by simp
     qed
   qed
 next
@@ -1791,17 +1791,17 @@
     show "Q \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain eT where 
-	wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
-	by cases simp
+        wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
+        by cases simp
       from da obtain E where
-	da_e: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright>E"
-	by cases simp
+        da_e: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright>E"
+        by cases simp
       from eval obtain v where
-	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_e P valid_A conf_s0 eval_e wt_e da_e
       obtain Q: "(Q\<leftarrow>\<diamondsuit>) \<lfloor>v\<rfloor>\<^sub>e s1 Z" and "s1\<Colon>\<preceq>(G,L)"
-	by (rule validE)
+        by (rule validE)
       thus ?thesis by simp
     qed
   qed
@@ -1821,24 +1821,24 @@
     show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
     proof -
       from wt obtain 
-	wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
-	by cases simp
+        wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
+        by cases simp
       from da obtain E where
-	da_c: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>E"
-	by cases simp
+        da_c: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>E"
+        by cases simp
       from eval obtain s1 where
-	eval_c: "G\<turnstile>s0 \<midarrow>c\<midarrow>n\<rightarrow> s1" and
-	s2: "s2 = abupd (absorb l) s1"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        eval_c: "G\<turnstile>s0 \<midarrow>c\<midarrow>n\<rightarrow> s1" and
+        s2: "s2 = abupd (absorb l) s1"
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from valid_c P valid_A conf_s0 eval_c wt_c da_c
       obtain Q: "(abupd (absorb l) .; Q) \<diamondsuit> s1 Z" 
-	by (rule validE)
+        by (rule validE)
       with s2 have "Q \<diamondsuit> s2 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s2\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -1859,45 +1859,45 @@
     show "R \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
     proof -
       from eval  obtain s1 where
-	eval_c1: "G\<turnstile>s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1" and
-	eval_c2: "G\<turnstile>s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        eval_c1: "G\<turnstile>s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1" and
+        eval_c2: "G\<turnstile>s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2"
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from wt obtain 
-	wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+        wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
         wt_c2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c2\<Colon>\<surd>"
-	by cases simp
+        by cases simp
       from da obtain C1 C2 where 
-	da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and 
-	da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>nrm C1 \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2" 
-	by cases simp
+        da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and 
+        da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>nrm C1 \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2" 
+        by cases simp
       from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1  
       obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"  
-	by (rule validE) 
+        by (rule validE) 
       have "R \<diamondsuit> s2 Z"
       proof (cases "normal s1")
-	case True
-	obtain C2' where 
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
-	proof -
-	  from eval_c1 wt_c1 da_c1 wf True
-	  have "nrm C1 \<subseteq> dom (locals (store s1))"
-	    by (cases rule: da_good_approx_evalnE) iprover
-	  with da_c2 show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	with valid_c2 Q valid_A conf_s1 eval_c2 wt_c2 
-	show ?thesis
-	  by (rule validE)
+        case True
+        obtain C2' where 
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
+        proof -
+          from eval_c1 wt_c1 da_c1 wf True
+          have "nrm C1 \<subseteq> dom (locals (store s1))"
+            by (cases rule: da_good_approx_evalnE) iprover
+          with da_c2 show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        with valid_c2 Q valid_A conf_s1 eval_c2 wt_c2 
+        show ?thesis
+          by (rule validE)
       next
-	case False
-	from valid_c2 Q valid_A conf_s1 eval_c2 False
-	show ?thesis
-	  by (cases rule: validE) iprover+
+        case False
+        from valid_c2 Q valid_A conf_s1 eval_c2 False
+        show ?thesis
+          by (cases rule: validE) iprover+
       qed
       moreover
       from eval wt da conf_s0 wf
       have "s2\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -1920,61 +1920,61 @@
     show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
     proof -
       from eval obtain b s1 where
-	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1" and
-	eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2"
-	using normal_s0 by (auto elim: evaln_elim_cases)
+        eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1" and
+        eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2"
+        using normal_s0 by (auto elim: evaln_elim_cases)
       from wt obtain  
-	wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
-	wt_then_else: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
-	by cases (simp split: split_if)
+        wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
+        wt_then_else: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
+        by cases (simp split: split_if)
       from da obtain E S where
-	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
-	da_then_else: 
-	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+        da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
+        da_then_else: 
+        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
              (dom (locals (store s0)) \<union> assigns_if (the_Bool b) e)
                \<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> S"
-	by cases (cases "the_Bool b",auto)
+        by cases (cases "the_Bool b",auto)
       from valid_e P valid_A conf_s0 eval_e wt_e da_e
       obtain "P' \<lfloor>b\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
-	by (rule validE)
+        by (rule validE)
       hence P': "\<And>v. (P'\<leftarrow>=the_Bool b) v s1 Z"
-	by (cases "normal s1") auto
+        by (cases "normal s1") auto
       have "Q \<diamondsuit> s2 Z"
       proof (cases "normal s1")
-	case True
-	have s0_s1: "dom (locals (store s0)) 
+        case True
+        have s0_s1: "dom (locals (store s0)) 
                       \<union> assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
-	proof -
-	  from eval_e 
-	  have eval_e': "G\<turnstile>s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
-	    by (rule evaln_eval)
-	  hence
-	    "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
-	    by (rule dom_locals_eval_mono_elim)
+        proof -
+          from eval_e 
+          have eval_e': "G\<turnstile>s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
+            by (rule evaln_eval)
+          hence
+            "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+            by (rule dom_locals_eval_mono_elim)
           moreover
-	  from eval_e' True wt_e
-	  have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
-	    by (rule assigns_if_good_approx') 
-	  ultimately show ?thesis by (rule Un_least)
-	qed
-	with da_then_else
-	obtain S' where
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+          from eval_e' True wt_e
+          have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+            by (rule assigns_if_good_approx') 
+          ultimately show ?thesis by (rule Un_least)
+        qed
+        with da_then_else
+        obtain S' where
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
               \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> S'"
-	  by (rule da_weakenE)
-	with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
-	show ?thesis
-	  by (rule validE)
+          by (rule da_weakenE)
+        with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
+        show ?thesis
+          by (rule validE)
       next
-	case False
-	with valid_then_else P' valid_A conf_s1 eval_then_else
-	show ?thesis
-	  by (cases rule: validE) iprover+
+        case False
+        with valid_then_else P' valid_A conf_s1 eval_then_else
+        show ?thesis
+          by (cases rule: validE) iprover+
       qed
       moreover
       from eval wt da conf_s0 wf
       have "s2\<Colon>\<preceq>(G, L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -1998,7 +1998,7 @@
     proof -
         --{* From the given hypothesises @{text valid_e} and @{text valid_c} 
            we can only reach the state after unfolding the loop once, i.e. 
-	   @{term "P \<diamondsuit> s2 Z"}, where @{term s2} is the state after executing
+           @{term "P \<diamondsuit> s2 Z"}, where @{term s2} is the state after executing
            @{term c}. To gain validity of the further execution of while, to
            finally get @{term "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3 Z"} we have to get 
            a hypothesis about the subsequent unfoldings (the whole loop again),
@@ -2008,202 +2008,202 @@
            @{text valid_c} in the goal.
         *}
       {
-	fix t s s' v 
-	assume "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
-	hence "\<And> Y' T E. 
+        fix t s s' v 
+        assume "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
+        hence "\<And> Y' T E. 
                 \<lbrakk>t =  \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s; \<forall>t\<in>A. G\<Turnstile>n\<Colon>t; P Y' s Z; s\<Colon>\<preceq>(G, L);
                  normal s \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T; 
                  normal s \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>E
                 \<rbrakk>\<Longrightarrow> (P'\<leftarrow>=False\<down>=\<diamondsuit>) v s' Z"
-	  (is "PROP ?Hyp n t s v s'")
-	proof (induct)
-	  case (Loop s0' e' b n' s1' c' s2' l' s3' Y' T E)
-	  note while = `(\<langle>l'\<bullet> While(e') c'\<rangle>\<^sub>s::term) = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
+          (is "PROP ?Hyp n t s v s'")
+        proof (induct)
+          case (Loop s0' e' b n' s1' c' s2' l' s3' Y' T E)
+          note 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
-	  note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
-	  note P = `P Y' (Norm s0') Z`
-	  note conf_s0' = `Norm s0'\<Colon>\<preceq>(G, L)`
+          note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
+          note P = `P Y' (Norm s0') Z`
+          note conf_s0' = `Norm s0'\<Colon>\<preceq>(G, L)`
           have wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<Colon>T"
-	    using Loop.prems eqs by simp
-	  have da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
+            using Loop.prems eqs by simp
+          have da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
                     dom (locals (store ((Norm s0')::state)))\<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<guillemotright>E"
-	    using Loop.prems eqs by simp
-	  have evaln_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<midarrow>n'\<rightarrow> s1'" 
-	    using Loop.hyps eqs by simp
-	  show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
-	  proof -
-	    from wt  obtain 
-	      wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
+            using Loop.prems eqs by simp
+          have evaln_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<midarrow>n'\<rightarrow> s1'" 
+            using Loop.hyps eqs by simp
+          show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
+          proof -
+            from wt  obtain 
+              wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
               wt_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>"
-	      by cases (simp add: eqs)
-	    from da obtain E S where
-	      da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+              by cases (simp add: eqs)
+            from da obtain E S where
+              da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                      \<turnstile> dom (locals (store ((Norm s0')::state))) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
-	      da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+              da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                      \<turnstile> (dom (locals (store ((Norm s0')::state))) 
                             \<union> assigns_if True e) \<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> S"
-	      by cases (simp add: eqs)
-	    from evaln_e 
-	    have eval_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<rightarrow> s1'"
-	      by (rule evaln_eval)
-	    from valid_e P valid_A conf_s0' evaln_e wt_e da_e
-	    obtain P': "P' \<lfloor>b\<rfloor>\<^sub>e s1' Z" and conf_s1': "s1'\<Colon>\<preceq>(G,L)"
-	      by (rule validE)
-	    show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
-	    proof (cases "normal s1'")
-	      case True
-	      note normal_s1'=this
-	      show ?thesis
-	      proof (cases "the_Bool b")
-		case True
-		with P' normal_s1' have P'': "(Normal (P'\<leftarrow>=True)) \<lfloor>b\<rfloor>\<^sub>e s1' Z"
-		  by auto
-		from True Loop.hyps obtain
-		  eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and 
-		  eval_while:  
-		     "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
-		  by (simp add: eqs)
-		from True Loop.hyps have
-		  hyp: "PROP ?Hyp n' \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s 
+              by cases (simp add: eqs)
+            from evaln_e 
+            have eval_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<rightarrow> s1'"
+              by (rule evaln_eval)
+            from valid_e P valid_A conf_s0' evaln_e wt_e da_e
+            obtain P': "P' \<lfloor>b\<rfloor>\<^sub>e s1' Z" and conf_s1': "s1'\<Colon>\<preceq>(G,L)"
+              by (rule validE)
+            show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
+            proof (cases "normal s1'")
+              case True
+              note normal_s1'=this
+              show ?thesis
+              proof (cases "the_Bool b")
+                case True
+                with P' normal_s1' have P'': "(Normal (P'\<leftarrow>=True)) \<lfloor>b\<rfloor>\<^sub>e s1' Z"
+                  by auto
+                from True Loop.hyps obtain
+                  eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and 
+                  eval_while:  
+                     "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
+                  by (simp add: eqs)
+                from True Loop.hyps have
+                  hyp: "PROP ?Hyp n' \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s 
                           (abupd (absorb (Cont l')) s2') \<diamondsuit> s3'"
-		  apply (simp only: True if_True eqs)
-		  apply (elim conjE)
-		  apply (tactic "smp_tac 3 1")
-		  apply fast
-		  done
-		from eval_e
-		have s0'_s1': "dom (locals (store ((Norm s0')::state))) 
+                  apply (simp only: True if_True eqs)
+                  apply (elim conjE)
+                  apply (tactic "smp_tac 3 1")
+                  apply fast
+                  done
+                from eval_e
+                have s0'_s1': "dom (locals (store ((Norm s0')::state))) 
                                   \<subseteq> dom (locals (store s1'))"
-		  by (rule dom_locals_eval_mono_elim)
-		obtain S' where
-		  da_c':
-		   "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(dom (locals (store s1')))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> S'" 
-		proof -
-		  note s0'_s1'
-		  moreover
-		  from eval_e normal_s1' wt_e 
-		  have "assigns_if True e \<subseteq> dom (locals (store s1'))"
-		    by (rule assigns_if_good_approx' [elim_format]) 
+                  by (rule dom_locals_eval_mono_elim)
+                obtain S' where
+                  da_c':
+                   "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(dom (locals (store s1')))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> S'" 
+                proof -
+                  note s0'_s1'
+                  moreover
+                  from eval_e normal_s1' wt_e 
+                  have "assigns_if True e \<subseteq> dom (locals (store s1'))"
+                    by (rule assigns_if_good_approx' [elim_format]) 
                        (simp add: True)
-		  ultimately 
-		  have "dom (locals (store ((Norm s0')::state)))
+                  ultimately 
+                  have "dom (locals (store ((Norm s0')::state)))
                            \<union> assigns_if True e \<subseteq> dom (locals (store s1'))"
-		    by (rule Un_least)
-		  with da_c show thesis
-		    by (rule da_weakenE) (rule that)
-		qed
-		with valid_c P'' valid_A conf_s1' eval_c wt_c
-		obtain "(abupd (absorb (Cont l)) .; P) \<diamondsuit> s2' Z" and 
+                    by (rule Un_least)
+                  with da_c show thesis
+                    by (rule da_weakenE) (rule that)
+                qed
+                with valid_c P'' valid_A conf_s1' eval_c wt_c
+                obtain "(abupd (absorb (Cont l)) .; P) \<diamondsuit> s2' Z" and 
                   conf_s2': "s2'\<Colon>\<preceq>(G,L)"
-		  by (rule validE)
-		hence P_s2': "P \<diamondsuit> (abupd (absorb (Cont l)) s2') Z"
-		  by simp
-		from conf_s2'
-		have conf_absorb: "abupd (absorb (Cont l)) s2' \<Colon>\<preceq>(G, L)"
-		  by (cases s2') (auto intro: conforms_absorb)
-		moreover
-		obtain E' where 
-		  da_while':
-		   "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
-		     dom (locals(store (abupd (absorb (Cont l)) s2')))
+                  by (rule validE)
+                hence P_s2': "P \<diamondsuit> (abupd (absorb (Cont l)) s2') Z"
+                  by simp
+                from conf_s2'
+                have conf_absorb: "abupd (absorb (Cont l)) s2' \<Colon>\<preceq>(G, L)"
+                  by (cases s2') (auto intro: conforms_absorb)
+                moreover
+                obtain E' where 
+                  da_while':
+                   "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+                     dom (locals(store (abupd (absorb (Cont l)) s2')))
                       \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<guillemotright> E'"
-		proof -
-		  note s0'_s1'
-		  also 
-		  from eval_c 
-		  have "G\<turnstile>s1' \<midarrow>c\<rightarrow> s2'"
-		    by (rule evaln_eval)
-		  hence "dom (locals (store s1')) \<subseteq> dom (locals (store s2'))"
-		    by (rule dom_locals_eval_mono_elim)
-		  also 
-		  have "\<dots>\<subseteq>dom (locals (store (abupd (absorb (Cont l)) s2')))"
-		    by simp
-		  finally
-		  have "dom (locals (store ((Norm s0')::state))) \<subseteq> \<dots>" .
-		  with da show thesis
-		    by (rule da_weakenE) (rule that)
-		qed
-		from valid_A P_s2' conf_absorb wt da_while'
-		show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z" 
-		  using hyp by (simp add: eqs)
-	      next
-		case False
-		with Loop.hyps obtain "s3'=s1'"
-		  by simp
-		with P' False show ?thesis
-		  by auto
-	      qed 
-	    next
-	      case False
-	      note abnormal_s1'=this
-	      have "s3'=s1'"
-	      proof -
-		from False obtain abr where abr: "abrupt s1' = Some abr"
-		  by (cases s1') auto
-		from eval_e _ wt_e wf
-		have no_jmp: "\<And> j. abrupt s1' \<noteq> Some (Jump j)"
-		  by (rule eval_expression_no_jump 
+                proof -
+                  note s0'_s1'
+                  also 
+                  from eval_c 
+                  have "G\<turnstile>s1' \<midarrow>c\<rightarrow> s2'"
+                    by (rule evaln_eval)
+                  hence "dom (locals (store s1')) \<subseteq> dom (locals (store s2'))"
+                    by (rule dom_locals_eval_mono_elim)
+                  also 
+                  have "\<dots>\<subseteq>dom (locals (store (abupd (absorb (Cont l)) s2')))"
+                    by simp
+                  finally
+                  have "dom (locals (store ((Norm s0')::state))) \<subseteq> \<dots>" .
+                  with da show thesis
+                    by (rule da_weakenE) (rule that)
+                qed
+                from valid_A P_s2' conf_absorb wt da_while'
+                show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z" 
+                  using hyp by (simp add: eqs)
+              next
+                case False
+                with Loop.hyps obtain "s3'=s1'"
+                  by simp
+                with P' False show ?thesis
+                  by auto
+              qed 
+            next
+              case False
+              note abnormal_s1'=this
+              have "s3'=s1'"
+              proof -
+                from False obtain abr where abr: "abrupt s1' = Some abr"
+                  by (cases s1') auto
+                from eval_e _ wt_e wf
+                have no_jmp: "\<And> j. abrupt s1' \<noteq> Some (Jump j)"
+                  by (rule eval_expression_no_jump 
                        [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
-		     simp
-		show ?thesis
-		proof (cases "the_Bool b")
-		  case True  
-		  with Loop.hyps obtain
-		    eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and 
-		    eval_while:  
-		     "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
-		    by (simp add: eqs)
-		  from eval_c abr have "s2'=s1'" by auto
-		  moreover from calculation no_jmp 
-		  have "abupd (absorb (Cont l)) s2'=s2'"
-		    by (cases s1') (simp add: absorb_def)
-		  ultimately show ?thesis
-		    using eval_while abr
-		    by auto
-		next
-		  case False
-		  with Loop.hyps show ?thesis by simp
-		qed
-	      qed
-	      with P' False show ?thesis
-		by auto
-	    qed
-	  qed
-	next
-	  case (Abrupt abr s t' n' Y' T E)
-	  note t' = `t' = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
-	  note conf = `(Some abr, s)\<Colon>\<preceq>(G, L)`
-	  note P = `P Y' (Some abr, s) Z`
-	  note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
-	  show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) (undefined3 t') (Some abr, s) Z"
-	  proof -
-	    have eval_e: 
-	      "G\<turnstile>(Some abr,s) \<midarrow>\<langle>e\<rangle>\<^sub>e\<succ>\<midarrow>n'\<rightarrow> (undefined3 \<langle>e\<rangle>\<^sub>e,(Some abr,s))"
-	      by auto
-	    from valid_e P valid_A conf eval_e 
-	    have "P' (undefined3 \<langle>e\<rangle>\<^sub>e) (Some abr,s) Z"
-	      by (cases rule: validE [where ?P="P"]) simp+
-	    with t' show ?thesis
-	      by auto
-	  qed
-	qed simp_all
+                     simp
+                show ?thesis
+                proof (cases "the_Bool b")
+                  case True  
+                  with Loop.hyps obtain
+                    eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and 
+                    eval_while:  
+                     "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
+                    by (simp add: eqs)
+                  from eval_c abr have "s2'=s1'" by auto
+                  moreover from calculation no_jmp 
+                  have "abupd (absorb (Cont l)) s2'=s2'"
+                    by (cases s1') (simp add: absorb_def)
+                  ultimately show ?thesis
+                    using eval_while abr
+                    by auto
+                next
+                  case False
+                  with Loop.hyps show ?thesis by simp
+                qed
+              qed
+              with P' False show ?thesis
+                by auto
+            qed
+          qed
+        next
+          case (Abrupt abr s t' n' Y' T E)
+          note t' = `t' = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
+          note conf = `(Some abr, s)\<Colon>\<preceq>(G, L)`
+          note P = `P Y' (Some abr, s) Z`
+          note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
+          show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) (undefined3 t') (Some abr, s) Z"
+          proof -
+            have eval_e: 
+              "G\<turnstile>(Some abr,s) \<midarrow>\<langle>e\<rangle>\<^sub>e\<succ>\<midarrow>n'\<rightarrow> (undefined3 \<langle>e\<rangle>\<^sub>e,(Some abr,s))"
+              by auto
+            from valid_e P valid_A conf eval_e 
+            have "P' (undefined3 \<langle>e\<rangle>\<^sub>e) (Some abr,s) Z"
+              by (cases rule: validE [where ?P="P"]) simp+
+            with t' show ?thesis
+              by auto
+          qed
+        qed simp_all
       } note generalized=this
       from eval _ valid_A P conf_s0 wt da
       have "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3 Z"
-	by (rule generalized)  simp_all
+        by (rule generalized)  simp_all
       moreover
       have "s3\<Colon>\<preceq>(G, L)"
       proof (cases "normal s0")
-	case True
-	from eval wt [OF True] da [OF True] conf_s0 wf
-	show ?thesis
-	  by (rule evaln_type_sound [elim_format]) simp
+        case True
+        from eval wt [OF True] da [OF True] conf_s0 wf
+        show ?thesis
+          by (rule evaln_type_sound [elim_format]) simp
       next
-	case False
-	with eval have "s3=s0"
-	  by auto
-	with conf_s0 show ?thesis 
-	  by simp
+        case False
+        with eval have "s3=s0"
+          by auto
+        with conf_s0 show ?thesis 
+          by simp
       qed
       ultimately show ?thesis ..
     qed
@@ -2224,14 +2224,14 @@
     show "P \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G,L)"
     proof -
       from eval obtain s where  
-	s: "s0=Norm s" "s1=(Some (Jump j), s)" 
-	using normal_s0 by (auto elim: evaln_elim_cases)
+        s: "s0=Norm s" "s1=(Some (Jump j), s)" 
+        using normal_s0 by (auto elim: evaln_elim_cases)
       with P have "P \<diamondsuit> s1 Z"
-	by simp
+        by simp
       moreover 
       from eval wt da conf_s0 wf
       have "s1\<Colon>\<preceq>(G,L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -2252,24 +2252,24 @@
     show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
     proof -
       from eval obtain s1 a where
-	eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
-	s2: "s2 = abupd (throw a) s1"
-	using normal_s0 by (auto elim: evaln_elim_cases)
+        eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
+        s2: "s2 = abupd (throw a) s1"
+        using normal_s0 by (auto elim: evaln_elim_cases)
       from wt obtain T where
-	wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
-	by cases simp
+        wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
+        by cases simp
       from da obtain E where
-	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
-	by cases simp
+        da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+        by cases simp
       from valid_e P valid_A conf_s0 eval_e wt_e da_e 
       obtain "(\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>) \<lfloor>a\<rfloor>\<^sub>e s1 Z"
-	by (rule validE)
+        by (rule validE)
       with s2 have "Q \<diamondsuit> s2 Z"
-	by simp
+        by simp
       moreover 
       from eval wt da conf_s0 wf
       have "s2\<Colon>\<preceq>(G,L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -2294,114 +2294,114 @@
     show "R \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
     proof -
       from eval obtain s1 s2 where
-	eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1" and
+        eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1" and
         sxalloc: "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" and
         s3: "if G,s2\<turnstile>catch C 
                 then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3 
                 else s3 = s2"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from wt obtain
-	wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
-	wt_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>\<turnstile>c2\<Colon>\<surd>"
-	by cases simp
+        wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+        wt_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>\<turnstile>c2\<Colon>\<surd>"
+        by cases simp
       from da obtain C1 C2 where
-	da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
-	da_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
+        da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
+        da_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
                    \<turnstile> (dom (locals (store s0)) \<union> {VName vn}) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
-	by cases simp
+        by cases simp
       from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1
       obtain sxQ: "(SXAlloc G Q) \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
-	by (rule validE)
+        by (rule validE)
       from sxalloc sxQ
       have Q: "Q \<diamondsuit> s2 Z"
-	by auto
+        by auto
       have "R \<diamondsuit> s3 Z"
       proof (cases "\<exists> x. abrupt s1 = Some (Xcpt x)")
-	case False
-	from sxalloc wf
-	have "s2=s1"
-	  by (rule sxalloc_type_sound [elim_format])
-	     (insert False, auto split: option.splits abrupt.splits )
-	with False 
-	have no_catch: "\<not>  G,s2\<turnstile>catch C"
-	  by (simp add: catch_def)
-	moreover
-	from no_catch s3
-	have "s3=s2"
-	  by simp
-	ultimately show ?thesis
-	  using Q Q_R by simp
+        case False
+        from sxalloc wf
+        have "s2=s1"
+          by (rule sxalloc_type_sound [elim_format])
+             (insert False, auto split: option.splits abrupt.splits )
+        with False 
+        have no_catch: "\<not>  G,s2\<turnstile>catch C"
+          by (simp add: catch_def)
+        moreover
+        from no_catch s3
+        have "s3=s2"
+          by simp
+        ultimately show ?thesis
+          using Q Q_R by simp
       next
-	case True
-	note exception_s1 = this
-	show ?thesis
-	proof (cases "G,s2\<turnstile>catch C") 
-	  case False
-	  with s3
-	  have "s3=s2"
-	    by simp
-	  with False Q Q_R show ?thesis
-	    by simp
-	next
-	  case True
-	  with s3 have eval_c2: "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3"
-	    by simp
-	  from conf_s1 sxalloc wf 
-	  have conf_s2: "s2\<Colon>\<preceq>(G, L)" 
-	    by (auto dest: sxalloc_type_sound 
+        case True
+        note exception_s1 = this
+        show ?thesis
+        proof (cases "G,s2\<turnstile>catch C") 
+          case False
+          with s3
+          have "s3=s2"
+            by simp
+          with False Q Q_R show ?thesis
+            by simp
+        next
+          case True
+          with s3 have eval_c2: "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3"
+            by simp
+          from conf_s1 sxalloc wf 
+          have conf_s2: "s2\<Colon>\<preceq>(G, L)" 
+            by (auto dest: sxalloc_type_sound 
                     split: option.splits abrupt.splits)
-	  from exception_s1 sxalloc wf
-	  obtain a 
-	    where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
-	    by (auto dest!: sxalloc_type_sound 
+          from exception_s1 sxalloc wf
+          obtain a 
+            where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
+            by (auto dest!: sxalloc_type_sound 
                             split: option.splits abrupt.splits)
-	  with True
-	  have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class C"
-	    by (cases s2) simp
-	  with xcpt_s2 conf_s2 wf
-	  have conf_new_xcpt: "new_xcpt_var vn s2 \<Colon>\<preceq>(G, L(VName vn\<mapsto>Class C))"
-	    by (auto dest: Try_lemma)
-	  obtain C2' where
-	    da_c2':
-	    "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
+          with True
+          have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class C"
+            by (cases s2) simp
+          with xcpt_s2 conf_s2 wf
+          have conf_new_xcpt: "new_xcpt_var vn s2 \<Colon>\<preceq>(G, L(VName vn\<mapsto>Class C))"
+            by (auto dest: Try_lemma)
+          obtain C2' where
+            da_c2':
+            "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
               \<turnstile> (dom (locals (store (new_xcpt_var vn s2)))) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
-	  proof -
-	    have "(dom (locals (store s0)) \<union> {VName vn}) 
+          proof -
+            have "(dom (locals (store s0)) \<union> {VName vn}) 
                     \<subseteq> dom (locals (store (new_xcpt_var vn s2)))"
             proof -
-	      from eval_c1 
+              from eval_c1 
               have "dom (locals (store s0)) 
                       \<subseteq> dom (locals (store s1))"
-		by (rule dom_locals_evaln_mono_elim)
+                by (rule dom_locals_evaln_mono_elim)
               also
               from sxalloc
               have "\<dots> \<subseteq> dom (locals (store s2))"
-		by (rule dom_locals_sxalloc_mono)
+                by (rule dom_locals_sxalloc_mono)
               also 
               have "\<dots> \<subseteq> dom (locals (store (new_xcpt_var vn s2)))" 
-		by (cases s2) (simp add: new_xcpt_var_def, blast) 
+                by (cases s2) (simp add: new_xcpt_var_def, blast) 
               also
               have "{VName vn} \<subseteq> \<dots>"
-		by (cases s2) simp
+                by (cases s2) simp
               ultimately show ?thesis
-		by (rule Un_least)
+                by (rule Un_least)
             qed
-	    with da_c2 show thesis
-	      by (rule da_weakenE) (rule that)
-	  qed
-	  from Q eval_c2 True 
-	  have "(Q \<and>. (\<lambda>s. G,s\<turnstile>catch C) ;. new_xcpt_var vn) 
+            with da_c2 show thesis
+              by (rule da_weakenE) (rule that)
+          qed
+          from Q eval_c2 True 
+          have "(Q \<and>. (\<lambda>s. G,s\<turnstile>catch C) ;. new_xcpt_var vn) 
                    \<diamondsuit> (new_xcpt_var vn s2) Z"
-	    by auto
-	  from valid_c2 this valid_A conf_new_xcpt eval_c2 wt_c2 da_c2'
-	  show "R \<diamondsuit> s3 Z"
-	    by (rule validE)
-	qed
+            by auto
+          from valid_c2 this valid_A conf_new_xcpt eval_c2 wt_c2 da_c2'
+          show "R \<diamondsuit> s3 Z"
+            by (rule validE)
+        qed
       qed
       moreover 
       from eval wt da conf_s0 wf
       have "s3\<Colon>\<preceq>(G,L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -2426,56 +2426,56 @@
     show "R \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
     proof -
       from eval obtain s1 abr1 s2 where
-	eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> (abr1, s1)" and
+        eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> (abr1, s1)" and
         eval_c2: "G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n\<rightarrow> s2" and
         s3: "s3 = (if \<exists>err. abr1 = Some (Error err) 
                       then (abr1, s1)
                       else abupd (abrupt_if (abr1 \<noteq> None) abr1) s2)"
-	using normal_s0 by (fastsimp elim: evaln_elim_cases)
+        using normal_s0 by (fastsimp elim: evaln_elim_cases)
       from wt obtain
-	wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
-	wt_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c2\<Colon>\<surd>"
-	by cases simp
+        wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+        wt_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c2\<Colon>\<surd>"
+        by cases simp
       from da obtain C1 C2 where
-	da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
-	da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
-	by cases simp
+        da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
+        da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
+        by cases simp
       from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1
       obtain Q: "Q \<diamondsuit> (abr1,s1) Z" and conf_s1: "(abr1,s1)\<Colon>\<preceq>(G,L)" 
-	by (rule validE)
+        by (rule validE)
       from Q 
       have Q': "(Q \<and>. (\<lambda>s. abr1 = fst s) ;. abupd (\<lambda>x. None)) \<diamondsuit> (Norm s1) Z"
-	by auto
+        by auto
       from eval_c1 wt_c1 da_c1 conf_s0 wf
       have  "error_free (abr1,s1)"
-	by (rule evaln_type_sound  [elim_format]) (insert normal_s0,simp)
+        by (rule evaln_type_sound  [elim_format]) (insert normal_s0,simp)
       with s3 have s3': "s3 = abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
-	by (simp add: error_free_def)
+        by (simp add: error_free_def)
       from conf_s1 
       have conf_Norm_s1: "Norm s1\<Colon>\<preceq>(G,L)"
-	by (rule conforms_NormI)
+        by (rule conforms_NormI)
       obtain C2' where 
-	da_c2': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        da_c2': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                    \<turnstile> dom (locals (store ((Norm s1)::state))) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
       proof -
-	from eval_c1 
-	have "dom (locals (store s0)) \<subseteq> dom (locals (store (abr1,s1)))"
+        from eval_c1 
+        have "dom (locals (store s0)) \<subseteq> dom (locals (store (abr1,s1)))"
           by (rule dom_locals_evaln_mono_elim)
-	hence "dom (locals (store s0)) 
+        hence "dom (locals (store s0)) 
                  \<subseteq> dom (locals (store ((Norm s1)::state)))"
-	  by simp
-	with da_c2 show thesis
-	  by (rule da_weakenE) (rule that)
+          by simp
+        with da_c2 show thesis
+          by (rule da_weakenE) (rule that)
       qed
       from valid_c2 Q' valid_A conf_Norm_s1 eval_c2 wt_c2 da_c2'
       have "(abupd (abrupt_if (abr1 \<noteq> None) abr1) .; R) \<diamondsuit> s2 Z"
-	by (rule validE)
+        by (rule validE)
       with s3' have "R \<diamondsuit> s3 Z"
-	by simp
+        by simp
       moreover
       from eval wt da conf_s0 wf
       have "s3\<Colon>\<preceq>(G,L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
@@ -2495,11 +2495,11 @@
     show "P \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
     proof -
       from P have inited: "inited C (globs (store s0))"
-	by simp
+        by simp
       with eval have "s3=s0"
-	using normal_s0 by (auto elim: evaln_elim_cases)
+        using normal_s0 by (auto elim: evaln_elim_cases)
       with P conf_s0 show ?thesis
-	by simp
+        by simp
     qed
   qed
 next
@@ -2529,88 +2529,88 @@
     proof -
       from P have not_inited: "\<not> inited C (globs (store s0))" by simp
       with eval c obtain s1 s2 where
-	eval_super: 
-	"G\<turnstile>Norm ((init_class_obj G C) (store s0)) 
+        eval_super: 
+        "G\<turnstile>Norm ((init_class_obj G C) (store s0)) 
            \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1" and
-	eval_init: "G\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2" and
+        eval_init: "G\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2" and
         s3: "s3 = (set_lvars (locals (store s1))) s2"
-	using normal_s0 by (auto elim!: evaln_elim_cases)
+        using normal_s0 by (auto elim!: evaln_elim_cases)
       from wt c have
-	cls_C: "class G C = Some c"
-	by cases auto
+        cls_C: "class G C = Some c"
+        by cases auto
       from wf cls_C have
-	wt_super: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        wt_super: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                          \<turnstile>(if C = Object then Skip else Init (super c))\<Colon>\<surd>"
-	by (cases "C=Object")
+        by (cases "C=Object")
            (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
       obtain S where
-	da_super:
-	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        da_super:
+        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
           \<turnstile> dom (locals (store ((Norm 
                             ((init_class_obj G C) (store s0)))::state))) 
                \<guillemotright>\<langle>if C = Object then Skip else Init (super c)\<rangle>\<^sub>s\<guillemotright> S"
       proof (cases "C=Object")
-	case True 
-	with da_Skip show ?thesis
-	  using that by (auto intro: assigned.select_convs)
+        case True 
+        with da_Skip show ?thesis
+          using that by (auto intro: assigned.select_convs)
       next
-	case False 
-	with da_Init show ?thesis
-	  by - (rule that, auto intro: assigned.select_convs)
+        case False 
+        with da_Init show ?thesis
+          by - (rule that, auto intro: assigned.select_convs)
       qed
       from normal_s0 conf_s0 wf cls_C not_inited
       have conf_init_cls: "(Norm ((init_class_obj G C) (store s0)))\<Colon>\<preceq>(G, L)"
-	by (auto intro: conforms_init_class_obj)	
+        by (auto intro: conforms_init_class_obj)        
       from P 
       have P': "(Normal (P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C)))
                    Y (Norm ((init_class_obj G C) (store s0))) Z"
-	by auto
+        by auto
 
       from valid_super P' valid_A conf_init_cls eval_super wt_super da_super
       obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
-	by (rule validE)
+        by (rule validE)
       
       from cls_C wf have wt_init: "\<lparr>prg=G, cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd>"
-	by (rule wf_prog_cdecl [THEN wf_cdecl_wt_init])
+        by (rule wf_prog_cdecl [THEN wf_cdecl_wt_init])
       from cls_C wf obtain I where 
-	"\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<^sub>s\<guillemotright> I"
-	by (rule wf_prog_cdecl [THEN wf_cdeclE,simplified]) blast
+        "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<^sub>s\<guillemotright> I"
+        by (rule wf_prog_cdecl [THEN wf_cdeclE,simplified]) blast
        (*  simplified: to rewrite \<langle>init c\<rangle> to In1r (init c) *) 
       then obtain I' where
-	da_init:
-	"\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>dom (locals (store ((set_lvars empty) s1))) 
+        da_init:
+        "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>dom (locals (store ((set_lvars empty) s1))) 
             \<guillemotright>\<langle>init c\<rangle>\<^sub>s\<guillemotright> I'"
-	by (rule da_weakenE) simp
+        by (rule da_weakenE) simp
       have conf_s1_empty: "(set_lvars empty) s1\<Colon>\<preceq>(G, empty)"
       proof -
-	from eval_super have
-	  "G\<turnstile>Norm ((init_class_obj G C) (store s0)) 
+        from eval_super have
+          "G\<turnstile>Norm ((init_class_obj G C) (store s0)) 
              \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1"
-	  by (rule evaln_eval)
-	from this wt_super wf
-	have s1_no_ret: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
-	  by - (rule eval_statement_no_jump 
+          by (rule evaln_eval)
+        from this wt_super wf
+        have s1_no_ret: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+          by - (rule eval_statement_no_jump 
                  [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>"], auto split: split_if)
         with conf_s1
-	show ?thesis
-	  by (cases s1) (auto intro: conforms_set_locals)
+        show ?thesis
+          by (cases s1) (auto intro: conforms_set_locals)
       qed
       
       obtain l where l: "l = locals (store s1)"
-	by simp
+        by simp
       with Q 
       have Q': "(Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty)
                   \<diamondsuit> ((set_lvars empty) s1) Z"
-	by auto
+        by auto
       from valid_init Q' valid_A conf_s1_empty eval_init wt_init da_init
       have "(set_lvars l .; R) \<diamondsuit> s2 Z"
-	by (rule validE)
+        by (rule validE)
       with s3 l have "R \<diamondsuit> s3 Z"
-	by simp
+        by simp
       moreover 
       from eval wt da conf_s0 wf
       have "s3\<Colon>\<preceq>(G,L)"
-	by (rule evaln_type_sound [elim_format]) simp
+        by (rule evaln_type_sound [elim_format]) simp
       ultimately show ?thesis ..
     qed
   qed
--- a/src/HOL/Bali/Basis.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Basis.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -86,19 +86,19 @@
     proof (cases rule: converse_rtranclE)
       assume "a=y"
       with a_v_r v_x_rt have "(y,x) \<in> r\<^sup>*"
-	by (auto intro: r_into_rtrancl rtrancl_trans)
+        by (auto intro: r_into_rtrancl rtrancl_trans)
       then show ?thesis 
-	by blast
+        by blast
     next
       fix w 
       assume a_w_r: "(a, w) \<in> r" and
             w_y_rt: "(w, y) \<in> r\<^sup>*"
       from a_v_r a_w_r unique 
       have "v=w" 
-	by auto
+        by auto
       with w_y_rt hyp 
       show ?thesis
-	by blast
+        by blast
     qed
   qed
 qed
@@ -213,11 +213,11 @@
 primrec  "the_In3 (In3 c) = c"
 
 syntax
-	 In1l	:: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
-	 In1r	:: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
+         In1l   :: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
+         In1r   :: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
 translations
-	"In1l e" == "In1 (Inl e)"
-	"In1r c" == "In1 (Inr c)"
+        "In1l e" == "In1 (Inl e)"
+        "In1r c" == "In1 (Inr c)"
 
 syntax the_In1l :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'al"
        the_In1r :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'ar"
--- a/src/HOL/Bali/Conform.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Conform.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -17,7 +17,7 @@
 \end{itemize}
 *}
 
-types	env' = "prog \<times> (lname, ty) table" (* same as env of WellType.thy *)
+types env' = "prog \<times> (lname, ty) table" (* same as env of WellType.thy *)
 
 
 section "extension of global store"
@@ -102,7 +102,7 @@
 constdefs
 
   conf  :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool"    ("_,_\<turnstile>_\<Colon>\<preceq>_"   [71,71,71,71] 70)
-	   "G,s\<turnstile>v\<Colon>\<preceq>T \<equiv> \<exists>T'\<in>typeof (\<lambda>a. Option.map obj_ty (heap s a)) v:G\<turnstile>T'\<preceq>T"
+           "G,s\<turnstile>v\<Colon>\<preceq>T \<equiv> \<exists>T'\<in>typeof (\<lambda>a. Option.map obj_ty (heap s a)) v:G\<turnstile>T'\<preceq>T"
 
 lemma conf_cong [simp]: "G,set_locals l s\<turnstile>v\<Colon>\<preceq>T = G,s\<turnstile>v\<Colon>\<preceq>T"
 by (auto simp: conf_def)
@@ -250,7 +250,7 @@
 done
 
 lemma lconf_init_vals [intro!]: 
-	" \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<Colon>\<preceq>]fs"
+        " \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<Colon>\<preceq>]fs"
 apply (unfold lconf_def)
 apply force
 done
@@ -338,7 +338,7 @@
   by (simp add: wlconf_def)
 
 lemma wlconf_init_vals [intro!]: 
-	" \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<sim>\<Colon>\<preceq>]fs"
+        " \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<sim>\<Colon>\<preceq>]fs"
 apply (unfold wlconf_def)
 apply force
 done
@@ -352,9 +352,9 @@
 constdefs
 
   oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool"  ("_,_\<turnstile>_\<Colon>\<preceq>\<surd>_"  [71,71,71,71] 70)
-	   "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r \<equiv> G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and> 
+           "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r \<equiv> G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and> 
                            (case r of 
-		              Heap a \<Rightarrow> is_type G (obj_ty obj) 
+                              Heap a \<Rightarrow> is_type G (obj_ty obj) 
                             | Stat C \<Rightarrow> True)"
 
 
@@ -405,11 +405,11 @@
 by (auto simp: conforms_def Let_def)
 
 lemma conforms_XcptLocD: "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); x = Some (Xcpt (Loc a))\<rbrakk> \<Longrightarrow>  
-	  G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)"
+          G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)"
 by (auto simp: conforms_def Let_def)
 
 lemma conforms_RetD: "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); x = Some (Jump Ret)\<rbrakk> \<Longrightarrow>  
-	  (locals s) Result \<noteq> None"
+          (locals s) Result \<noteq> None"
 by (auto simp: conforms_def Let_def)
 
 lemma conforms_RefTD: 
--- a/src/HOL/Bali/Decl.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Decl.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -171,7 +171,7 @@
 
 types     
         fdecl           (* field declaration, cf. 8.3 *)
-	= "vname \<times> field"
+        = "vname \<times> field"
 
 
 translations
@@ -294,9 +294,9 @@
 
 record  iface = ibody + --{* interface *}
          isuperIfs:: "qtname list" --{* superinterface list *}
-types	
-	idecl           --{* interface declaration, cf. 9.1 *}
-	= "qtname \<times> iface"
+types
+        idecl           --{* interface declaration, cf. 9.1 *}
+        = "qtname \<times> iface"
 
 translations
   "ibody" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list\<rparr>"
@@ -326,9 +326,9 @@
 record "class" = cbody +           --{* class *}
         super   :: "qtname"      --{* superclass *}
         superIfs:: "qtname list" --{* implemented interfaces *}
-types	
-	cdecl           --{* class declaration, cf. 8.1 *}
-	= "qtname \<times> class"
+types
+        cdecl           --{* class declaration, cf. 8.1 *}
+        = "qtname \<times> class"
 
 translations
   "cbody" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
@@ -388,8 +388,8 @@
 
 constdefs standard_classes :: "cdecl list"
          "standard_classes \<equiv> [ObjectC, SXcptC Throwable,
-		SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast,
-		SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]"
+                SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast,
+                SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]"
 
 
 section "programs"
@@ -409,10 +409,10 @@
   is_class  :: "prog  \<Rightarrow> qtname  \<Rightarrow> bool"
 
 translations
-	   "iface G I" == "table_of (ifaces G) I"
-	   "class G C" == "table_of (classes G) C"
-	"is_iface G I" == "iface G I \<noteq> None"
-	"is_class G C" == "class G C \<noteq> None"
+           "iface G I" == "table_of (ifaces G) I"
+           "class G C" == "table_of (classes G) C"
+        "is_iface G I" == "iface G I \<noteq> None"
+        "is_class G C" == "class G C \<noteq> None"
 
 
 section "is type"
@@ -421,12 +421,12 @@
   is_type :: "prog \<Rightarrow>     ty \<Rightarrow> bool"
   isrtype :: "prog \<Rightarrow> ref_ty \<Rightarrow> bool"
 
-primrec	"is_type G (PrimT pt)  = True"
-	"is_type G (RefT  rt)  = isrtype G rt"
-	"isrtype G (NullT    ) = True"
-	"isrtype G (IfaceT tn) = is_iface G tn"
-	"isrtype G (ClassT tn) = is_class G tn"
-	"isrtype G (ArrayT T ) = is_type  G T"
+primrec "is_type G (PrimT pt)  = True"
+        "is_type G (RefT  rt)  = isrtype G rt"
+        "isrtype G (NullT    ) = True"
+        "isrtype G (IfaceT tn) = is_iface G tn"
+        "isrtype G (ClassT tn) = is_class G tn"
+        "isrtype G (ArrayT T ) = is_type  G T"
 
 lemma type_is_iface: "is_type G (Iface I) \<Longrightarrow> is_iface G I"
 by auto
@@ -457,7 +457,7 @@
 
 translations
         "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" == "(C,D) \<in> subcls1 G"
-	"G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" (* cf. 8.1.3 *)
+        "G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" (* cf. 8.1.3 *)
         "G\<turnstile>C \<prec>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^+"
  
 
@@ -516,7 +516,7 @@
   
   ws_prog  :: "prog \<Rightarrow> bool"
  "ws_prog G \<equiv> (\<forall>(I,i)\<in>set (ifaces  G). ws_idecl G I (isuperIfs i)) \<and> 
-	      (\<forall>(C,c)\<in>set (classes G). ws_cdecl G C (super c))"
+              (\<forall>(C,c)\<in>set (classes G). ws_cdecl G C (super c))"
 
 
 lemma ws_progI: 
@@ -710,12 +710,12 @@
     next
       case False
       with ws iscls obtain sc where
-	sc: "class G (super c) = Some sc"
-	by (auto dest: ws_prog_cdeclD)
+        sc: "class G (super c) = Some sc"
+        by (auto dest: ws_prog_cdeclD)
       from iscls False have "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 (super c)" by (rule subcls1I)
       with False ws step hyp iscls sc
       show "P C c" 
-	by (auto)  
+        by (auto)  
     qed
   qed
   with clsC show "P C c" by auto
@@ -742,12 +742,12 @@
       case Nil
       with if_I hyp_sub 
       show "P I" 
-	by auto
+        by auto
     next
       case (Cons hd tl)
       with hyp if_I hyp_sub 
       show "P I" 
-	by auto
+        by auto
     qed
   qed
 qed
@@ -802,7 +802,7 @@
   \<equiv> iface_rec (G,I)  
               (\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus> 
                         (Option.set \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
-	
+        
 
 
 end
--- a/src/HOL/Bali/DeclConcepts.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/DeclConcepts.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -922,20 +922,20 @@
            "memberid n = memberid m" 
            "G\<turnstile> mbr m declared_in C" 
            "declclass m = C"
-	by auto
+        by auto
       with member_n   
       show ?thesis
-	by (cases n, cases m) 
+        by (cases n, cases m) 
            (auto simp add: declared_in_def 
                            cdeclaredmethd_def cdeclaredfield_def
                     split: memberdecl.splits)
     next
       case (Inherited m' _ _)
       then have "G\<turnstile> memberid m undeclared_in C"
-	by simp
+        by simp
       with eqid member_n
       show ?thesis
-	by (cases n) (auto dest: declared_not_undeclared)
+        by (cases n) (auto dest: declared_not_undeclared)
     qed
   next
     case (Inherited n C S)
@@ -950,14 +950,14 @@
       then have "G\<turnstile> mbr m declared_in C" by simp
       with eqid undecl
       show ?thesis
-	by (cases m) (auto dest: declared_not_undeclared)
+        by (cases m) (auto dest: declared_not_undeclared)
     next
       case Inherited 
       with super have "G \<turnstile> m member_of S"
-	by (auto dest!: subcls1D)
+        by (auto dest!: subcls1D)
       with eqid hyp
       show ?thesis 
-	by blast
+        by blast
     qed
   qed
 qed
@@ -1227,15 +1227,15 @@
       assume "D=C" 
       with super member_of_D_props 
       show ?thesis
-	by (auto intro: members.Inherited)
+        by (auto intro: members.Inherited)
     next
       case Subcls
       assume "G\<turnstile>D\<prec>\<^sub>C C"
       with super 
       have "G\<turnstile>S\<preceq>\<^sub>C C"
-	by (auto dest: subcls1D subcls_superD)
+        by (auto dest: subcls1D subcls_superD)
       with subclseq_C_m hyp show ?thesis
-	by blast
+        by blast
     qed
   qed
 qed
@@ -1660,34 +1660,34 @@
       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)
+        by (cases m,auto)
       with clsC 
       have "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = Some m"
-	by (cases m)
-	   (auto simp add: declared_in_def cdeclaredmethd_def
-	            intro: table_of_mapconst_SomeI)
+        by (cases m)
+           (auto simp add: declared_in_def cdeclaredmethd_def
+                    intro: table_of_mapconst_SomeI)
       with clsC neq_C_Obj ws 
       show ?thesis
-	by (simp add: methd_rec)
+        by (simp add: methd_rec)
     next
       case (Inherited membr Ca S)
       with clsC
       have eq_Ca_C: "Ca=C" and
             undecl: "G\<turnstile>mid sig undeclared_in C" and
              super: "G \<turnstile>Methd sig m member_of (super c)"
-	by (auto dest: subcls1D)
+        by (auto dest: subcls1D)
       from eq_Ca_C clsC undecl 
       have "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = None"
-	by (auto simp add: undeclared_in_def cdeclaredmethd_def
-	            intro: table_of_mapconst_NoneI)
+        by (auto simp add: undeclared_in_def cdeclaredmethd_def
+                    intro: table_of_mapconst_NoneI)
       moreover
       from Inherited have "G \<turnstile> C inherits (method sig m)"
-	by (auto simp add: inherits_def)
+        by (auto simp add: inherits_def)
       moreover
       note clsC neq_C_Obj ws super hyp 
       ultimately
       show ?thesis
-	by (auto simp add: methd_rec intro: filter_tab_SomeI)
+        by (auto simp add: methd_rec intro: filter_tab_SomeI)
     qed
   qed
 qed
@@ -1771,12 +1771,12 @@
       then have eq_statC_Obj: "statC = Object" ..
       show ?thesis 
       proof (cases "methd G statC sig")
-	case None then show ?thesis by (simp add: dynmethd_def)
+        case None then show ?thesis by (simp add: dynmethd_def)
       next
-	case Some
-	with True Object ws eq_statC_Obj 
-	show ?thesis
-	  by (auto simp add: dynmethd_def class_rec
+        case Some
+        with True Object ws eq_statC_Obj 
+        show ?thesis
+          by (auto simp add: dynmethd_def class_rec
                       intro: filter_tab_SomeI)
       qed
     qed
@@ -1791,88 +1791,88 @@
       note subclseq_dynC_statC = True
       show ?thesis
       proof (cases "methd G statC sig")
-	case None then show ?thesis by (simp add: dynmethd_def)
+        case None then show ?thesis by (simp add: dynmethd_def)
       next
-	case (Some statM)
-	note statM = Some
-	let "?filter C" = 
+        case (Some statM)
+        note statM = Some
+        let "?filter C" = 
               "filter_tab
                 (\<lambda>_ dynM. G,sig \<turnstile> dynM overrides statM \<or> dynM = statM)
                 (methd G C)"
         let "?class_rec C" =
               "(class_rec (G, C) empty
                            (\<lambda>C c subcls_mthds. subcls_mthds ++ (?filter C)))"
-	from statM Subcls ws subclseq_dynC_statC
-	have dynmethd_dynC_def:
+        from statM Subcls ws subclseq_dynC_statC
+        have dynmethd_dynC_def:
              "?Dynmethd_def dynC sig =
                 ((?class_rec (super c)) 
                  ++
                 (?filter dynC)) sig"
          by (simp (no_asm_simp) only: dynmethd_def class_rec)
-	    auto
+            auto
        show ?thesis
        proof (cases "dynC = statC")
-	 case True
-	 with subclseq_dynC_statC statM dynmethd_dynC_def
-	 have "?Dynmethd_def dynC sig = Some statM"
-	   by (auto intro: map_add_find_right filter_tab_SomeI)
-	 with subclseq_dynC_statC True Some 
-	 show ?thesis
-	   by auto
+         case True
+         with subclseq_dynC_statC statM dynmethd_dynC_def
+         have "?Dynmethd_def dynC sig = Some statM"
+           by (auto intro: map_add_find_right filter_tab_SomeI)
+         with subclseq_dynC_statC True Some 
+         show ?thesis
+           by auto
        next
-	 case False
-	 with subclseq_dynC_statC Subcls 
-	 have subclseq_super_statC: "G\<turnstile>(super c) \<preceq>\<^sub>C statC"
-	   by (blast dest: subclseq_superD)
-	 show ?thesis
-	 proof (cases "methd G dynC sig") 
-	   case None
-	   then have "?filter dynC sig = None"
-	     by (rule filter_tab_None)
+         case False
+         with subclseq_dynC_statC Subcls 
+         have subclseq_super_statC: "G\<turnstile>(super c) \<preceq>\<^sub>C statC"
+           by (blast dest: subclseq_superD)
+         show ?thesis
+         proof (cases "methd G dynC sig") 
+           case None
+           then have "?filter dynC sig = None"
+             by (rule filter_tab_None)
            then have "?Dynmethd_def dynC sig=?class_rec (super c) sig"
-	     by (simp add: dynmethd_dynC_def)
-	   with  subclseq_super_statC statM None
-	   have "?Dynmethd_def dynC sig = ?Dynmethd_def (super c) sig"
-	     by (auto simp add: empty_def dynmethd_def)
+             by (simp add: dynmethd_dynC_def)
+           with  subclseq_super_statC statM None
+           have "?Dynmethd_def dynC sig = ?Dynmethd_def (super c) sig"
+             by (auto simp add: empty_def dynmethd_def)
            with None subclseq_dynC_statC statM
-	   show ?thesis 
-	     by simp
-	 next
-	   case (Some dynM)
-	   note dynM = Some
-	   let ?Termination = "G \<turnstile> qmdecl sig dynM overrides qmdecl sig statM \<or>
+           show ?thesis 
+             by simp
+         next
+           case (Some dynM)
+           note dynM = Some
+           let ?Termination = "G \<turnstile> qmdecl sig dynM overrides qmdecl sig statM \<or>
                                dynM = statM"
-	   show ?thesis
-	   proof (cases "?filter dynC sig")
-	     case None
-	     with dynM 
-	     have no_termination: "\<not> ?Termination"
-	       by (simp add: filter_tab_def)
-	     from None 
-	     have "?Dynmethd_def dynC sig=?class_rec (super c) sig"
-	       by (simp add: dynmethd_dynC_def)
-	     with subclseq_super_statC statM dynM no_termination
-	     show ?thesis
-	       by (auto simp add: empty_def dynmethd_def)
-	   next
-	     case Some
-	     with dynM
-	     have "termination": "?Termination"
-	       by (auto)
-	     with Some dynM
-	     have "?Dynmethd_def dynC sig=Some dynM"
-	      by (auto simp add: dynmethd_dynC_def)
-	     with subclseq_super_statC statM dynM "termination"
-	     show ?thesis
-	       by (auto simp add: dynmethd_def)
-	   qed
-	 qed
+           show ?thesis
+           proof (cases "?filter dynC sig")
+             case None
+             with dynM 
+             have no_termination: "\<not> ?Termination"
+               by (simp add: filter_tab_def)
+             from None 
+             have "?Dynmethd_def dynC sig=?class_rec (super c) sig"
+               by (simp add: dynmethd_dynC_def)
+             with subclseq_super_statC statM dynM no_termination
+             show ?thesis
+               by (auto simp add: empty_def dynmethd_def)
+           next
+             case Some
+             with dynM
+             have "termination": "?Termination"
+               by (auto)
+             with Some dynM
+             have "?Dynmethd_def dynC sig=Some dynM"
+              by (auto simp add: dynmethd_dynC_def)
+             with subclseq_super_statC statM dynM "termination"
+             show ?thesis
+               by (auto simp add: dynmethd_def)
+           qed
+         qed
        qed
      qed
    qed
  qed
 qed
-	       
+               
 lemma dynmethd_C_C:"\<lbrakk>is_class G C; ws_prog G\<rbrakk> 
 \<Longrightarrow> dynmethd G C C sig = methd G C sig"          
 apply (auto simp add: dynmethd_rec)
@@ -1993,17 +1993,17 @@
       case Static
       with is_cls_statC ws subclseq_dynC_statC 
       show ?thesis 
-	by (auto intro: rtrancl_trans dest: methd_declC)
+        by (auto intro: rtrancl_trans dest: methd_declC)
     next
       case Override
       with clsDynC ws 
       show ?thesis 
-	by (auto dest: methd_declC)
+        by (auto dest: methd_declC)
     next
       case Recursion
       with hyp subclseq_dynC_super 
       show ?thesis 
-	by (auto intro: rtrancl_trans) 
+        by (auto intro: rtrancl_trans) 
     qed
   qed
 qed
@@ -2041,7 +2041,7 @@
       case Eq
       with ws statM clsDynC' 
       show ?thesis
-	by (auto simp add: dynmethd_rec)
+        by (auto simp add: dynmethd_rec)
     next
       case Subcls
       assume "G\<turnstile>dynC\<prec>\<^sub>C statC"
@@ -2049,7 +2049,7 @@
       have "G\<turnstile>super dc\<preceq>\<^sub>C statC" by (rule subcls_superD)
       with hyp ws clsDynC' subclseq' statM
       show ?thesis
-	by (auto simp add: dynmethd_rec)
+        by (auto simp add: dynmethd_rec)
     qed
   qed
 qed
--- a/src/HOL/Bali/DefiniteAssignment.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/DefiniteAssignment.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -260,11 +260,11 @@
     proof (cases "the_Bool bv")
       case True
       with Cond show ?thesis using v bv
-	by (auto intro: hyp_CondL)
+        by (auto intro: hyp_CondL)
     next
       case False
       with Cond show ?thesis using v bv
-	by (auto intro: hyp_CondR)
+        by (auto intro: hyp_CondR)
     qed
   qed (simp_all)
   with const 
@@ -364,7 +364,7 @@
       hence "?Ass b e1" by (rule hyp_e1)
       with emptyC bv True
       show ?thesis
-	by simp
+        by simp
     next
       case False
       with const bv  
@@ -372,7 +372,7 @@
       hence "?Ass b e2" by (rule hyp_e2)
       with emptyC bv False
       show ?thesis
-	by simp
+        by simp
     qed
   qed (simp_all)
   with boolConst
@@ -411,7 +411,7 @@
       hence "?Ass b e1" by (rule hyp_e1)
       with bv True
       show ?thesis
-	by simp
+        by simp
     next
       case False
       with const bv  
@@ -419,7 +419,7 @@
       hence "?Ass b e2" by (rule hyp_e2)
       with bv False 
       show ?thesis
-	by simp
+        by simp
     qed
   qed (simp_all)
   with boolConst
@@ -959,14 +959,14 @@
       from `E\<turnstile>(Cast T e)\<Colon>- (PrimT Boolean)`
       obtain Te where "E\<turnstile>e\<Colon>-Te"
                            "prg E\<turnstile>Te\<preceq>? PrimT Boolean"
-	by cases simp
+        by cases simp
       thus ?thesis
-	by - (drule cast_Boolean2,simp)
+        by - (drule cast_Boolean2,simp)
     qed
     with Cast.hyps
     show ?case
       by simp
-  next	
+  next  
     case (Lit val) 
     thus ?case
       by - (erule wt_elim_cases, cases "val", auto simp add: empty_dt_def)
@@ -1000,17 +1000,17 @@
       case None
       with boolean_e1 boolean_e2
       show ?thesis
-	using hyp_e1 hyp_e2 
-	by (auto)
+        using hyp_e1 hyp_e2 
+        by (auto)
     next
       case (Some bv)
       show ?thesis
       proof (cases "the_Bool bv")
-	case True
-	with Some show ?thesis using hyp_e1 boolean_e1 by auto
+        case True
+        with Some show ?thesis using hyp_e1 boolean_e1 by auto
       next
-	case False
-	with Some show ?thesis using hyp_e2 boolean_e2 by auto
+        case False
+        with Some show ?thesis using hyp_e2 boolean_e2 by auto
       qed
     qed
   qed simp_all
@@ -1085,7 +1085,7 @@
       fix l'
       from hyp_brk
       have "rmlab l (brk C) l'  \<subseteq> rmlab l (brk C') l'"
-	by  (cases "l=l'") simp_all
+        by  (cases "l=l'") simp_all
     }
     moreover note A A'
     ultimately show ?case
@@ -1154,15 +1154,15 @@
     { fix l'
       have  "brk A l' \<subseteq> brk A' l'"
       proof (cases "constVal e")
-	case None
-	with A A' C' 
-	show ?thesis
-	   by (cases "l=l'") auto
+        case None
+        with A A' C' 
+        show ?thesis
+           by (cases "l=l'") auto
       next
-	case (Some bv)
-	with A A' C'
-	show ?thesis
-	  by (cases "the_Bool bv", cases "l=l'") auto
+        case (Some bv)
+        with A A' C'
+        show ?thesis
+          by (cases "the_Bool bv", cases "l=l'") auto
       qed
     }
     ultimately show ?case
@@ -1381,7 +1381,7 @@
     proof -
       from B'
       have "(B \<union> assigns_if True e) \<subseteq> (B' \<union> assigns_if True e)"
-	by blast
+        by blast
       moreover
       have "PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c1\<rangle>" by (rule If.hyps)
       ultimately 
@@ -1391,7 +1391,7 @@
     obtain C2' where "Env\<turnstile> (B' \<union> assigns_if False e) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'"
     proof - 
       from B' have "(B \<union> assigns_if False e) \<subseteq> (B' \<union> assigns_if False e)"
-	by blast
+        by blast
       moreover
       have "PROP ?Hyp Env (B \<union> assigns_if False e) \<langle>c2\<rangle>" by (rule If.hyps)
       ultimately
@@ -1415,7 +1415,7 @@
     proof -
       from B'
       have "(B \<union> assigns_if True e) \<subseteq> (B' \<union> assigns_if True e)"
-	by blast
+        by blast
       moreover
       have "PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c\<rangle>" by (rule Loop.hyps)
       ultimately 
@@ -1461,7 +1461,7 @@
       moreover
       have "PROP ?Hyp (Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>) 
                       (B \<union> {VName vn}) \<langle>c2\<rangle>" 
-	by (rule Try.hyps)
+        by (rule Try.hyps)
       ultimately
       show ?thesis using that by iprover
     qed
@@ -1516,7 +1516,7 @@
     obtain E2' where "Env\<turnstile> B' \<union>  assigns_if True e1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2'"
     proof -
       from B' have "B \<union> assigns_if True e1 \<subseteq> B' \<union>  assigns_if True e1"
-	by blast
+        by blast
       moreover
       have "PROP ?Hyp Env (B \<union> assigns_if True e1) \<langle>e2\<rangle>" by (rule CondAnd.hyps)
       ultimately show ?thesis using that by iprover
@@ -1538,7 +1538,7 @@
     obtain E2' where "Env\<turnstile> B' \<union> assigns_if False e1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2'"
     proof -
       from B' have "B \<union> assigns_if False e1 \<subseteq> B' \<union>  assigns_if False e1"
-	by blast
+        by blast
       moreover
       have "PROP ?Hyp Env (B \<union> assigns_if False e1) \<langle>e2\<rangle>" by (rule CondOr.hyps)
       ultimately show ?thesis using that by iprover
@@ -1562,7 +1562,7 @@
       have "Env\<turnstile> B \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1" by (rule BinOp.hyps)
       from this B' E1'
       have "nrm E1 \<subseteq> nrm E1'"
-	by (rule da_monotone [THEN conjE])
+        by (rule da_monotone [THEN conjE])
       moreover 
       have "PROP ?Hyp Env (nrm E1) \<langle>e2\<rangle>" by (rule BinOp.hyps)
       ultimately show ?thesis using that by iprover
@@ -1612,7 +1612,7 @@
       have "Env\<turnstile> B \<guillemotright>\<langle>v\<rangle>\<guillemotright> V" by (rule Ass.hyps)
       from this B' V'
       have "nrm V \<subseteq> nrm V'"
-	by (rule da_monotone [THEN conjE])
+        by (rule da_monotone [THEN conjE])
       moreover 
       have "PROP ?Hyp Env (nrm V) \<langle>e\<rangle>" by (rule Ass.hyps)
       ultimately show ?thesis using that by iprover
@@ -1636,7 +1636,7 @@
     proof -
       from B'
       have "(B \<union> assigns_if True c) \<subseteq> (B' \<union> assigns_if True c)"
-	by blast
+        by blast
       moreover
       have "PROP ?Hyp Env (B \<union> assigns_if True c) \<langle>e1\<rangle>" by (rule CondBool.hyps)
       ultimately 
@@ -1647,7 +1647,7 @@
     proof -
       from B'
       have "(B \<union> assigns_if False c) \<subseteq> (B' \<union> assigns_if False c)"
-	by blast
+        by blast
       moreover
       have "PROP ?Hyp Env (B \<union> assigns_if False c) \<langle>e2\<rangle>" by(rule CondBool.hyps)
       ultimately 
@@ -1672,7 +1672,7 @@
     proof -
       from B'
       have "(B \<union> assigns_if True c) \<subseteq> (B' \<union> assigns_if True c)"
-	by blast
+        by blast
       moreover
       have "PROP ?Hyp Env (B \<union> assigns_if True c) \<langle>e1\<rangle>" by (rule Cond.hyps)
       ultimately 
@@ -1683,7 +1683,7 @@
     proof -
       from B'
       have "(B \<union> assigns_if False c) \<subseteq> (B' \<union> assigns_if False c)"
-	by blast
+        by blast
       moreover
       have "PROP ?Hyp Env (B \<union> assigns_if False c) \<langle>e2\<rangle>" by (rule Cond.hyps)
       ultimately 
@@ -1708,7 +1708,7 @@
       have "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E" by (rule Call.hyps)
       from this B' E'
       have "nrm E \<subseteq> nrm E'"
-	by (rule da_monotone [THEN conjE])
+        by (rule da_monotone [THEN conjE])
       moreover 
       have "PROP ?Hyp Env (nrm E) \<langle>args\<rangle>" by (rule Call.hyps)
       ultimately show ?thesis using that by iprover
@@ -1728,10 +1728,10 @@
       moreover note B'
       moreover
       from B' obtain C' where da_c: "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
-	by (rule Body.hyps [elim_format]) blast
+        by (rule Body.hyps [elim_format]) blast
       ultimately
       have "nrm C \<subseteq> nrm C'"
-	by (rule da_monotone [THEN conjE])
+        by (rule da_monotone [THEN conjE])
       with da_c that show ?thesis by iprover
     qed
     moreover 
@@ -1762,7 +1762,7 @@
       have "Env\<turnstile> B \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1" by (rule AVar.hyps)
       from this B' E1'
       have "nrm E1 \<subseteq> nrm E1'"
-	by (rule da_monotone [THEN conjE])
+        by (rule da_monotone [THEN conjE])
       moreover 
       have "PROP ?Hyp Env (nrm E1) \<langle>e2\<rangle>" by (rule AVar.hyps)
       ultimately show ?thesis using that by iprover
@@ -1788,7 +1788,7 @@
       have "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E" by (rule Cons.hyps)
       from this B' E'
       have "nrm E \<subseteq> nrm E'"
-	by (rule da_monotone [THEN conjE])
+        by (rule da_monotone [THEN conjE])
       moreover 
       have "PROP ?Hyp Env (nrm E) \<langle>es\<rangle>" by (rule Cons.hyps)
       ultimately show ?thesis using that by iprover
--- a/src/HOL/Bali/DefiniteAssignmentCorrect.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/DefiniteAssignmentCorrect.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -407,7 +407,7 @@
         with wt_e G have jmp_s1: "?Jmp jmps s1" 
           by simp
         show ?thesis
-	proof (cases "the_Bool b")
+        proof (cases "the_Bool b")
           case False
           from Loop.hyps
           have "s3=s1"
@@ -434,20 +434,20 @@
           have "?Jmp jmps (abupd (absorb (Cont l)) s2)"
           proof -
             {
-      	      fix j'
-      	      assume abs: "abrupt (abupd (absorb (Cont l)) s2)=Some (Jump j')"
-      	      have "j' \<in> jmps"
-      	      proof (cases "j' = Cont l")
-      		case True
-      		with abs show ?thesis
-      		  by (cases s2) (simp add: absorb_def)
-      	      next
-      		case False 
-      		with abs have "abrupt s2 = Some (Jump j')"
-      		  by (cases s2) (simp add: absorb_def) 
-      		with jmp_s2 False show ?thesis
-      		  by simp
-      	      qed
+              fix j'
+              assume abs: "abrupt (abupd (absorb (Cont l)) s2)=Some (Jump j')"
+              have "j' \<in> jmps"
+              proof (cases "j' = Cont l")
+                case True
+                with abs show ?thesis
+                  by (cases s2) (simp add: absorb_def)
+              next
+                case False 
+                with abs have "abrupt s2 = Some (Jump j')"
+                  by (cases s2) (simp add: absorb_def) 
+                with jmp_s2 False show ?thesis
+                  by simp
+              qed
             }
             thus ?thesis by simp
           qed
@@ -512,7 +512,7 @@
         proof (cases "G,s2\<turnstile>catch C")
           case False
           from Try.hyps have "s3=s2" 
-	    by (simp (no_asm_use) only: False if_False)
+            by (simp (no_asm_use) only: False if_False)
           with jmp have "abrupt s1 = Some (Jump j)"
             using sxalloc_no_jump' [OF s2] by simp
           with jmp_s1 
@@ -666,7 +666,7 @@
       assume jmp: "abrupt s3 = Some (Jump j)"
       have "j\<in>jmps"
       proof -
-	note G = `prg Env = G`
+        note G = `prg Env = G`
         from NewA.prems 
         obtain wt_init: "Env\<turnstile>init_comp_ty elT\<Colon>\<surd>" and 
                wt_size: "Env\<turnstile>e\<Colon>-PrimT Integer"
@@ -1315,18 +1315,18 @@
       case True
       with Loop.hyps
       obtain
-	s0_s1: 
-	"dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))" and
-	s1_s2: "dom (locals (store s1)) \<subseteq> dom (locals (store s2))" and
-	s2_s3: "dom (locals (store s2)) \<subseteq> dom (locals (store s3))"
-	by simp
+        s0_s1: 
+        "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))" and
+        s1_s2: "dom (locals (store s1)) \<subseteq> dom (locals (store s2))" and
+        s2_s3: "dom (locals (store s2)) \<subseteq> dom (locals (store s3))"
+        by simp
       note s0_s1 also note s1_s2 also note s2_s3
       finally show ?thesis
-	by simp
+        by simp
     next
       case False
       with Loop.hyps show ?thesis
-	by simp
+        by simp
     qed
   next
     case Jmp thus ?case by simp
@@ -1348,16 +1348,16 @@
       from True Try.hyps 
       have "dom (locals (store (new_xcpt_var vn s2))) 
               \<subseteq> dom (locals (store s3))"
-	by simp
+        by simp
       hence "dom (locals (store s2)) \<subseteq> dom (locals (store s3))"
-	by (cases s2, simp )
+        by (cases s2, simp )
       finally show ?thesis by simp
     next
       case False
       note s0_s1 also note s1_s2
       finally 
       show ?thesis 
-	using False Try.hyps by simp
+        using False Try.hyps by simp
     qed
   next
     case (Fin s0 c1 x1 s1 c2 s2 s3) 
@@ -1365,22 +1365,22 @@
     proof (cases "\<exists>err. x1 = Some (Error err)")
       case True
       with Fin.hyps show ?thesis
-	by simp
+        by simp
     next
       case False
       from Fin.hyps
       have "dom (locals (store ((Norm s0)::state))) 
              \<subseteq> dom (locals (store (x1, s1)))" 
-	by simp
+        by simp
       hence "dom (locals (store ((Norm s0)::state))) 
               \<subseteq> dom (locals (store ((Norm s1)::state)))"
-	by simp
+        by simp
       also
       from Fin.hyps
       have "\<dots> \<subseteq> dom (locals (store s2))"
-	by simp
+        by simp
       finally show ?thesis 
-	using Fin.hyps by simp
+        using Fin.hyps by simp
     qed
   next
     case (Init C c s0 s3 s1 s2)
@@ -1393,14 +1393,14 @@
       with Init.hyps 
       obtain s0_s1: "dom (locals (store (Norm ((init_class_obj G C) s0))))
                      \<subseteq> dom (locals (store s1))" and
-	     s3: "s3 = (set_lvars (locals (snd s1))) s2"
-	by simp
+             s3: "s3 = (set_lvars (locals (snd s1))) s2"
+        by simp
       from s0_s1
       have "dom (locals (store (Norm s0))) \<subseteq> dom (locals (store s1))"
-	by (cases s0) simp
+        by (cases s0) simp
       with s3
       have "dom (locals (store (Norm s0))) \<subseteq> dom (locals (store s3))"
-	by (cases s2) simp
+        by (cases s2) simp
       thus ?thesis by simp
     qed
   next
@@ -1460,28 +1460,28 @@
       with Ass.hyps 
       have ass_ok:
         "\<And> s val. dom (locals (store s)) \<subseteq> dom (locals (store (f val s)))"
-	by simp
+        by simp
       note s0_s1
       also
       from Ass.hyps
       have "dom (locals (store s1)) \<subseteq> dom (locals (store s2))" 
-	by simp
+        by simp
       also
       from ass_ok
       have "\<dots> \<subseteq> dom (locals (store (assign f v s2)))"
-	by (rule dom_locals_assign_mono)
+        by (rule dom_locals_assign_mono)
       finally show ?thesis by simp
     next
       case False
       with `G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2`
       have "s2=s1"
-	by auto
+        by auto
       with s0_s1 False
       have "dom (locals (store ((Norm s0)::state))) 
             \<subseteq> dom (locals (store (assign f v s2)))"
-	by simp
+        by simp
       thus ?thesis
-	by simp
+        by simp
     qed
   next
     case (Cond s0 e0 b s1 e1 e2 v s2)
@@ -1529,7 +1529,7 @@
                  abrupt s2 = Some (Jump (Cont l))
              then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)`
       show ?thesis
-	by simp
+        by simp
     qed
     finally show ?case by simp
   next
@@ -1693,7 +1693,7 @@
     proof -
       from NewA
       have "normal (abupd (check_neg i) s2)"
-	by - (erule halloc_no_abrupt [rule_format])
+        by - (erule halloc_no_abrupt [rule_format])
       hence "normal s2" by (cases s2) simp
       with NewA.hyps
       show ?thesis by iprover
@@ -1734,7 +1734,7 @@
       note `G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
                       else In1r Skip)\<succ>\<rightarrow> (In1 v2, s2)`
       thus ?thesis
-	by (rule dom_locals_eval_mono_elim)
+        by (rule dom_locals_eval_mono_elim)
     qed
     finally have s2: "assigns (In1l e1) \<subseteq> dom (locals (store s2))" .
     show ?case
@@ -1745,7 +1745,7 @@
       case False
       with BinOp 
       have "assigns (In1l e2) \<subseteq> dom (locals (store s2))"
-	by (simp add: need_second_arg_def)
+        by (simp add: need_second_arg_def)
       with s2
       show ?thesis using False by simp
     qed
@@ -1789,19 +1789,19 @@
     proof (cases "\<exists> n. va = LVar n")
       case False
       with va_e show ?thesis 
-	by (simp add: Un_assoc)
+        by (simp add: Un_assoc)
     next
       case True
       then obtain n where va: "va = LVar n"
-	by blast
+        by blast
       with Ass.hyps 
       have "G\<turnstile>Norm s0 \<midarrow>LVar n=\<succ>(w,f)\<rightarrow> s1" 
-	by simp
+        by simp
       hence "(w,f) = lvar n s0"
-	by (rule eval_elim_cases) simp
+        by (rule eval_elim_cases) simp
       with nrm_ass_s2
       have "n \<in> dom (locals (store (assign f v s2)))"
-	by (cases s2) (simp add: assign_def Let_def lvar_def)
+        by (cases s2) (simp add: assign_def Let_def lvar_def)
       with va_e True va 
       show ?thesis by (simp add: Un_assoc)
     qed
@@ -1821,25 +1821,25 @@
       case True
       with Cond 
       have "assigns (In1l e1) \<subseteq> dom (locals (store s2))"
-	by simp
+        by simp
       hence "assigns (In1l e1) \<inter> assigns (In1l e2) \<subseteq> \<dots>" 
-	by blast
+        by blast
       with e0
       have "assigns (In1l e0) \<union> assigns (In1l e1) \<inter> assigns (In1l e2) 
              \<subseteq> dom (locals (store s2))"
-	by (rule Un_least)
+        by (rule Un_least)
       thus ?thesis using True by simp 
     next
       case False
       with Cond 
       have " assigns (In1l e2) \<subseteq> dom (locals (store s2))"
-	by simp
+        by simp
       hence "assigns (In1l e1) \<inter> assigns (In1l e2) \<subseteq> \<dots>"
-	by blast
+        by blast
       with e0
       have "assigns (In1l e0) \<union> assigns (In1l e1) \<inter> assigns (In1l e2) 
              \<subseteq> dom (locals (store s2))"
-	by (rule Un_least)
+        by (rule Un_least)
       thus ?thesis using False by simp 
     qed
   next
@@ -1849,15 +1849,15 @@
       from `normal ((set_lvars (locals (snd s2))) s4)`
       have normal_s4: "normal s4" by simp
       hence "normal s3'" using Call.hyps
-	by - (erule eval_no_abrupt_lemma [rule_format]) 
+        by - (erule eval_no_abrupt_lemma [rule_format]) 
       moreover note
        `s3' = check_method_access G accC statT mode \<lparr>name=mn, parTs=pTs\<rparr> a' s3`
       ultimately have "normal s3"
-	by (cases s3) (simp add: check_method_access_def Let_def) 
+        by (cases s3) (simp add: check_method_access_def Let_def) 
       moreover
       note s3 = `s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a' vs s2`
       ultimately show "normal s2"
-	by (cases s2) (simp add: init_lvars_def2)
+        by (cases s2) (simp add: init_lvars_def2)
     qed
     hence "normal s1" using Call.hyps
       by - (erule eval_no_abrupt_lemma [rule_format])
@@ -1892,9 +1892,9 @@
     proof -
       note `normal s3`
       with s3 have "normal s2'"
-	by (cases s2') (simp add: check_field_access_def Let_def)
+        by (cases s2') (simp add: check_field_access_def Let_def)
       with avar show "normal s2"
-	by (cases s2) (simp add: fvar_def2)
+        by (cases s2) (simp add: fvar_def2)
     qed
     with FVar.hyps 
     have "assigns (In1l e) \<subseteq> dom (locals (store s2))"
@@ -1904,9 +1904,9 @@
     proof -
       from avar
       have "s2' = snd (fvar statDeclC stat fn a s2)"
-	by (cases "fvar statDeclC stat fn a s2") simp
+        by (cases "fvar statDeclC stat fn a s2") simp
       thus ?thesis
-	by simp (rule dom_locals_fvar_mono)
+        by simp (rule dom_locals_fvar_mono)
     qed
     also from s3
     have "\<dots> \<subseteq>  dom (locals (store s3))"
@@ -1941,9 +1941,9 @@
     have "dom (locals (store s2)) \<subseteq>  dom (locals (store s2'))"
     proof -
       from avar have "s2' = snd (avar G i a s2)"
-	by (cases "avar G i a s2") simp
+        by (cases "avar G i a s2") simp
       thus ?thesis
-	by simp (rule dom_locals_avar_mono)
+        by simp (rule dom_locals_avar_mono)
     qed
     finally  
     show ?case
@@ -2065,24 +2065,24 @@
     proof (cases "need_second_arg binop v1")
       case True
       with v2 nrm_s1  obtain s1' 
-	where "G\<turnstile>Norm s1' \<midarrow>e2-\<succ>v2\<rightarrow> s" 
-	by (cases s1) simp
+        where "G\<turnstile>Norm s1' \<midarrow>e2-\<succ>v2\<rightarrow> s" 
+        by (cases s1) simp
       with c2 obtain "v2 = c2" "normal s"
-	by (rule BinOp.hyps [elim_format]) iprover
+        by (rule BinOp.hyps [elim_format]) iprover
       with c c1 c2 eq_v1_c1 v 
       show ?thesis by simp
     next
       case False
       with nrm_s1 v2
       have "s=s1"
-	by (cases s1) (auto elim!: eval_elim_cases)
+        by (cases s1) (auto elim!: eval_elim_cases)
       moreover
       from False c v eq_v1_c1  
       have "v = c"
-	by (simp add: eval_binop_arg2_indep)
+        by (simp add: eval_binop_arg2_indep)
       ultimately 
       show ?thesis
-	using nrm_s1 by simp
+        using nrm_s1 by simp
     qed  (* hallo ehco *)
   next
     case Super hence False by simp thus ?case ..
@@ -2111,25 +2111,25 @@
       case True
       with c cb c1 eq_vb_cb 
       have "c = c1"
-	by simp
+        by simp
       moreover
       from True eval_v nrm_s1 obtain s1' 
-	where "G\<turnstile>Norm s1' \<midarrow>e1-\<succ>v\<rightarrow> s"
-	by (cases s1) simp
+        where "G\<turnstile>Norm s1' \<midarrow>e1-\<succ>v\<rightarrow> s"
+        by (cases s1) simp
       with c1 obtain "c1 = v" "normal s"
-	by (rule Cond.hyps [elim_format]) iprover 
+        by (rule Cond.hyps [elim_format]) iprover 
       ultimately show ?thesis by simp
     next
       case False
       with c cb c2 eq_vb_cb 
       have "c = c2"
-	by simp
+        by simp
       moreover
       from False eval_v nrm_s1 obtain s1' 
-	where "G\<turnstile>Norm s1' \<midarrow>e2-\<succ>v\<rightarrow> s"
-	by (cases s1) simp
+        where "G\<turnstile>Norm s1' \<midarrow>e2-\<succ>v\<rightarrow> s"
+        by (cases s1) simp
       with c2 obtain "c2 = v" "normal s"
-	by (rule Cond.hyps [elim_format]) iprover 
+        by (rule Cond.hyps [elim_format]) iprover 
       ultimately show ?thesis by simp
     qed
   next
@@ -2218,13 +2218,13 @@
       case True
       from c1 wt_e1 
       have  "typeof empty_dt c1 = Some (PrimT Boolean)" 
-	by (rule Cond.hyps)
+        by (rule Cond.hyps)
       with True c cb c1 show ?thesis by simp
     next
       case False
       from c2 wt_e2 
       have "typeof empty_dt c2 = Some (PrimT Boolean)" 
-	by (rule Cond.hyps)
+        by (rule Cond.hyps)
       with False c cb c2 show ?thesis by simp
     qed
   next
@@ -2233,7 +2233,7 @@
   with const wt
   show ?thesis
     by iprover
-qed	
+qed     
       
 lemma assigns_if_good_approx:
   assumes
@@ -2272,13 +2272,13 @@
      proof -
        from s2 and `normal s2`
        have "normal s1"
-	 by (cases s1) simp
+         by (cases s1) simp
        moreover
        from `Env\<turnstile>Cast T e\<Colon>-PrimT Boolean`
        have "Env\<turnstile>e\<Colon>- PrimT Boolean" 
-	 by cases (auto dest: cast_Boolean2)
+         by cases (auto dest: cast_Boolean2)
        ultimately show ?thesis 
-	 by (rule Cast.hyps [elim_format]) auto
+         by (rule Cast.hyps [elim_format]) auto
      qed
      also from s2 
      have "\<dots> \<subseteq> dom (locals (store s2))"
@@ -2311,33 +2311,33 @@
        note `normal s1`
        moreover note bool_e
        ultimately have "assigns_if (the_Bool v) e \<subseteq> dom (locals (store s1))"
-	 by (rule UnOp.hyps [elim_format]) auto
+         by (rule UnOp.hyps [elim_format]) auto
        moreover
        from bool have "unop = UNot"
-	 by cases (cases unop, simp_all)
+         by cases (cases unop, simp_all)
        moreover note None
        ultimately 
        have "assigns_if (the_Bool (eval_unop unop v)) (UnOp unop e) 
               \<subseteq> dom (locals (store s1))"
-	 by simp
+         by simp
        thus ?thesis by simp
      next
        case (Some c)
        moreover
        from `prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1`
        have "prg Env\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>eval_unop unop v\<rightarrow> s1" 
-	 by (rule eval.UnOp)
+         by (rule eval.UnOp)
        with Some
        have "eval_unop unop v=c"
-	 by (rule constVal_eval_elim) simp
+         by (rule constVal_eval_elim) simp
        moreover
        from Some bool
        obtain b where "c=Bool b"
-	 by (rule constVal_Boolean [elim_format])
-	    (cases c, simp_all add: empty_dt_def)
+         by (rule constVal_Boolean [elim_format])
+            (cases c, simp_all add: empty_dt_def)
        ultimately
        have "assigns_if (the_Bool (eval_unop unop v)) (UnOp unop e) = {}"
-	 by simp
+         by simp
        thus ?thesis by simp
      qed
    next
@@ -2349,155 +2349,155 @@
        moreover
        from BinOp.hyps
        have
-	 "prg Env\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>eval_binop binop v1 v2\<rightarrow> s2"
-	 by - (rule eval.BinOp) 
+         "prg Env\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>eval_binop binop v1 v2\<rightarrow> s2"
+         by - (rule eval.BinOp) 
        with Some
        have "eval_binop binop v1 v2=c"
-	 by (rule constVal_eval_elim) simp
+         by (rule constVal_eval_elim) simp
        moreover
        from Some bool
        obtain b where "c = Bool b"
-	 by (rule constVal_Boolean [elim_format])
-	    (cases c, simp_all add: empty_dt_def)
+         by (rule constVal_Boolean [elim_format])
+            (cases c, simp_all add: empty_dt_def)
        ultimately
        have "assigns_if (the_Bool (eval_binop binop v1 v2)) (BinOp binop e1 e2) 
              = {}"
-	 by simp 
+         by simp 
        thus ?thesis by simp
      next
        case None
        show ?thesis
        proof (cases "binop=CondAnd \<or> binop=CondOr")
-	 case True
-	 from bool obtain bool_e1: "Env\<turnstile>e1\<Colon>-PrimT Boolean" and
+         case True
+         from bool obtain bool_e1: "Env\<turnstile>e1\<Colon>-PrimT Boolean" and
                           bool_e2: "Env\<turnstile>e2\<Colon>-PrimT Boolean"
-	   using True by cases auto
-	 have "assigns_if (the_Bool v1) e1 \<subseteq> dom (locals (store s1))"
-	 proof -
-	   from BinOp have "normal s1"
-	     by - (erule eval_no_abrupt_lemma [rule_format])
-	   from this bool_e1
-	   show ?thesis
-	     by (rule BinOp.hyps [elim_format]) auto
-	 qed
-	 also
-	 from BinOp.hyps
-	 have "\<dots> \<subseteq> dom (locals (store s2))"
-	   by - (erule dom_locals_eval_mono_elim,simp)
-	 finally
-	 have e1_s2: "assigns_if (the_Bool v1) e1 \<subseteq> dom (locals (store s2))".
-	 from True show ?thesis
-	 proof
-	   assume condAnd: "binop = CondAnd"
-	   show ?thesis
-	   proof (cases "the_Bool (eval_binop binop v1 v2)")
-	     case True
-	     with condAnd 
-	     have need_second: "need_second_arg binop v1"
-	       by (simp add: need_second_arg_def)
-	     from `normal s2`
-	     have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
-	       by (rule BinOp.hyps [elim_format]) 
+           using True by cases auto
+         have "assigns_if (the_Bool v1) e1 \<subseteq> dom (locals (store s1))"
+         proof -
+           from BinOp have "normal s1"
+             by - (erule eval_no_abrupt_lemma [rule_format])
+           from this bool_e1
+           show ?thesis
+             by (rule BinOp.hyps [elim_format]) auto
+         qed
+         also
+         from BinOp.hyps
+         have "\<dots> \<subseteq> dom (locals (store s2))"
+           by - (erule dom_locals_eval_mono_elim,simp)
+         finally
+         have e1_s2: "assigns_if (the_Bool v1) e1 \<subseteq> dom (locals (store s2))".
+         from True show ?thesis
+         proof
+           assume condAnd: "binop = CondAnd"
+           show ?thesis
+           proof (cases "the_Bool (eval_binop binop v1 v2)")
+             case True
+             with condAnd 
+             have need_second: "need_second_arg binop v1"
+               by (simp add: need_second_arg_def)
+             from `normal s2`
+             have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
+               by (rule BinOp.hyps [elim_format]) 
                   (simp add: need_second bool_e2)+
-	     with e1_s2
-	     have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
-		    \<subseteq> dom (locals (store s2))"
-	       by (rule Un_least)
-	     with True condAnd None show ?thesis
-	       by simp
-	   next
-	     case False
-	     note binop_False = this
-	     show ?thesis
-	     proof (cases "need_second_arg binop v1")
-	       case True
-	       with binop_False condAnd
-	       obtain "the_Bool v1=True" and "the_Bool v2 = False"
-		 by (simp add: need_second_arg_def)
-	       moreover
-	       from `normal s2`
-	       have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
-		 by (rule BinOp.hyps [elim_format]) (simp add: True bool_e2)+
-	       with e1_s2
-	       have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
-		      \<subseteq> dom (locals (store s2))"
-		 by (rule Un_least)
-	       moreover note binop_False condAnd None
-	       ultimately show ?thesis
-		 by auto
-	     next
-	       case False
-	       with binop_False condAnd
-	       have "the_Bool v1=False"
-		 by (simp add: need_second_arg_def)
-	       with e1_s2 
-	       show ?thesis
-		 using binop_False condAnd None by auto
-	     qed
-	   qed
-	 next
-	   assume condOr: "binop = CondOr"
-	   show ?thesis
-	   proof (cases "the_Bool (eval_binop binop v1 v2)")
-	     case False
-	     with condOr 
-	     have need_second: "need_second_arg binop v1"
-	       by (simp add: need_second_arg_def)
-	     from `normal s2`
-	     have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
-	       by (rule BinOp.hyps [elim_format]) 
+             with e1_s2
+             have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
+                    \<subseteq> dom (locals (store s2))"
+               by (rule Un_least)
+             with True condAnd None show ?thesis
+               by simp
+           next
+             case False
+             note binop_False = this
+             show ?thesis
+             proof (cases "need_second_arg binop v1")
+               case True
+               with binop_False condAnd
+               obtain "the_Bool v1=True" and "the_Bool v2 = False"
+                 by (simp add: need_second_arg_def)
+               moreover
+               from `normal s2`
+               have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
+                 by (rule BinOp.hyps [elim_format]) (simp add: True bool_e2)+
+               with e1_s2
+               have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
+                      \<subseteq> dom (locals (store s2))"
+                 by (rule Un_least)
+               moreover note binop_False condAnd None
+               ultimately show ?thesis
+                 by auto
+             next
+               case False
+               with binop_False condAnd
+               have "the_Bool v1=False"
+                 by (simp add: need_second_arg_def)
+               with e1_s2 
+               show ?thesis
+                 using binop_False condAnd None by auto
+             qed
+           qed
+         next
+           assume condOr: "binop = CondOr"
+           show ?thesis
+           proof (cases "the_Bool (eval_binop binop v1 v2)")
+             case False
+             with condOr 
+             have need_second: "need_second_arg binop v1"
+               by (simp add: need_second_arg_def)
+             from `normal s2`
+             have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
+               by (rule BinOp.hyps [elim_format]) 
                   (simp add: need_second bool_e2)+
-	     with e1_s2
-	     have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
-		    \<subseteq> dom (locals (store s2))"
-	       by (rule Un_least)
-	     with False condOr None show ?thesis
-	       by simp
-	   next
-	     case True
-	     note binop_True = this
-	     show ?thesis
-	     proof (cases "need_second_arg binop v1")
-	       case True
-	       with binop_True condOr
-	       obtain "the_Bool v1=False" and "the_Bool v2 = True"
-		 by (simp add: need_second_arg_def)
-	       moreover
-	       from `normal s2`
-	       have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
-		 by (rule BinOp.hyps [elim_format]) (simp add: True bool_e2)+
-	       with e1_s2
-	       have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
-		      \<subseteq> dom (locals (store s2))"
-		 by (rule Un_least)
-	       moreover note binop_True condOr None
-	       ultimately show ?thesis
-		 by auto
-	     next
-	       case False
-	       with binop_True condOr
-	       have "the_Bool v1=True"
-		 by (simp add: need_second_arg_def)
-	       with e1_s2 
-	       show ?thesis
-		 using binop_True condOr None by auto
-	     qed
-	   qed
-	 qed  
+             with e1_s2
+             have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
+                    \<subseteq> dom (locals (store s2))"
+               by (rule Un_least)
+             with False condOr None show ?thesis
+               by simp
+           next
+             case True
+             note binop_True = this
+             show ?thesis
+             proof (cases "need_second_arg binop v1")
+               case True
+               with binop_True condOr
+               obtain "the_Bool v1=False" and "the_Bool v2 = True"
+                 by (simp add: need_second_arg_def)
+               moreover
+               from `normal s2`
+               have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
+                 by (rule BinOp.hyps [elim_format]) (simp add: True bool_e2)+
+               with e1_s2
+               have "assigns_if (the_Bool v1) e1 \<union> assigns_if (the_Bool v2) e2
+                      \<subseteq> dom (locals (store s2))"
+                 by (rule Un_least)
+               moreover note binop_True condOr None
+               ultimately show ?thesis
+                 by auto
+             next
+               case False
+               with binop_True condOr
+               have "the_Bool v1=True"
+                 by (simp add: need_second_arg_def)
+               with e1_s2 
+               show ?thesis
+                 using binop_True condOr None by auto
+             qed
+           qed
+         qed  
        next
-	 case False
-	 note `\<not> (binop = CondAnd \<or> binop = CondOr)`
-	 from BinOp.hyps
-	 have
-	   "prg Env\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>eval_binop binop v1 v2\<rightarrow> s2"
-	   by - (rule eval.BinOp)  
-	 moreover note `normal s2`
-	 ultimately
-	 have "assignsE (BinOp binop e1 e2) \<subseteq> dom (locals (store s2))"
-	   by (rule assignsE_good_approx)
-	 with False None
-	 show ?thesis 
-	   by simp
+         case False
+         note `\<not> (binop = CondAnd \<or> binop = CondOr)`
+         from BinOp.hyps
+         have
+           "prg Env\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>eval_binop binop v1 v2\<rightarrow> s2"
+           by - (rule eval.BinOp)  
+         moreover note `normal s2`
+         ultimately
+         have "assignsE (BinOp binop e1 e2) \<subseteq> dom (locals (store s2))"
+           by (rule assignsE_good_approx)
+         with False None
+         show ?thesis 
+           by simp
        qed
      qed
    next     
@@ -2533,68 +2533,68 @@
        note eval_e0 
        moreover
        from Cond.hyps and `normal s2` have "normal s1"
-	 by - (erule eval_no_abrupt_lemma [rule_format],simp)
+         by - (erule eval_no_abrupt_lemma [rule_format],simp)
        ultimately
        have "assignsE e0 \<subseteq> dom (locals (store s1))"
-	 by (rule assignsE_good_approx)
+         by (rule assignsE_good_approx)
        also
        from Cond
        have "\<dots> \<subseteq> dom (locals (store s2))"
-	 by - (erule dom_locals_eval_mono [elim_format],simp)
+         by - (erule dom_locals_eval_mono [elim_format],simp)
        finally show ?thesis .
      qed
      show ?case
      proof (cases "constVal e0")
        case None
        have "assigns_if (the_Bool v) e1 \<inter>  assigns_if (the_Bool v) e2 
-	      \<subseteq> dom (locals (store s2))"
+              \<subseteq> dom (locals (store s2))"
        proof (cases "the_Bool b")
-	 case True
-	 from `normal s2`
-	 have "assigns_if (the_Bool v) e1 \<subseteq> dom (locals (store s2))"    
-	   by (rule Cond.hyps [elim_format]) (simp_all add: wt_e1 True)
-	 thus ?thesis
-	   by blast
+         case True
+         from `normal s2`
+         have "assigns_if (the_Bool v) e1 \<subseteq> dom (locals (store s2))"    
+           by (rule Cond.hyps [elim_format]) (simp_all add: wt_e1 True)
+         thus ?thesis
+           by blast
        next
-	 case False
-	 from `normal s2`
-	 have "assigns_if (the_Bool v) e2 \<subseteq> dom (locals (store s2))"    
-	   by (rule Cond.hyps [elim_format]) (simp_all add: wt_e2 False)
-	 thus ?thesis
-	   by blast
+         case False
+         from `normal s2`
+         have "assigns_if (the_Bool v) e2 \<subseteq> dom (locals (store s2))"    
+           by (rule Cond.hyps [elim_format]) (simp_all add: wt_e2 False)
+         thus ?thesis
+           by blast
        qed
        with e0_s2
        have "assignsE e0 \<union> 
              (assigns_if (the_Bool v) e1 \<inter>  assigns_if (the_Bool v) e2)
-	       \<subseteq> dom (locals (store s2))"
-	 by (rule Un_least)
+               \<subseteq> dom (locals (store s2))"
+         by (rule Un_least)
        with None show ?thesis
-	 by simp
+         by simp
      next
        case (Some c)
        from this eval_e0 have eq_b_c: "b=c" 
-	 by (rule constVal_eval_elim) 
+         by (rule constVal_eval_elim) 
        show ?thesis
        proof (cases "the_Bool c")
-	 case True
-	 from `normal s2`
-	 have "assigns_if (the_Bool v) e1 \<subseteq> dom (locals (store s2))"
-	   by (rule Cond.hyps [elim_format]) (simp_all add: eq_b_c True wt_e1)
-	 with e0_s2
-	 have "assignsE e0 \<union> assigns_if (the_Bool v) e1  \<subseteq> \<dots>"
-	   by (rule Un_least)
-	 with Some True show ?thesis
-	   by simp
+         case True
+         from `normal s2`
+         have "assigns_if (the_Bool v) e1 \<subseteq> dom (locals (store s2))"
+           by (rule Cond.hyps [elim_format]) (simp_all add: eq_b_c True wt_e1)
+         with e0_s2
+         have "assignsE e0 \<union> assigns_if (the_Bool v) e1  \<subseteq> \<dots>"
+           by (rule Un_least)
+         with Some True show ?thesis
+           by simp
        next
-	 case False
-	 from `normal s2`
-	 have "assigns_if (the_Bool v) e2 \<subseteq> dom (locals (store s2))"    
-	   by (rule Cond.hyps [elim_format]) (simp_all add: eq_b_c False wt_e2)
-	 with e0_s2
-	 have "assignsE e0 \<union> assigns_if (the_Bool v) e2  \<subseteq> \<dots>"
-	   by (rule Un_least)
-	 with Some False show ?thesis
-	   by simp
+         case False
+         from `normal s2`
+         have "assigns_if (the_Bool v) e2 \<subseteq> dom (locals (store s2))"    
+           by (rule Cond.hyps [elim_format]) (simp_all add: eq_b_c False wt_e2)
+         with e0_s2
+         have "assignsE e0 \<union> assigns_if (the_Bool v) e2  \<subseteq> \<dots>"
+           by (rule Un_least)
+         with Some False show ?thesis
+           by simp
        qed
      qed
    next
@@ -2715,38 +2715,38 @@
       assume normal: "normal (abupd (absorb j) s1)"
       show "nrm A \<subseteq> dom (locals (store (abupd (absorb j) s1)))"
       proof (cases "abrupt s1")
-	case None
-	with norm_c A 
-	show ?thesis
-	  by auto
+        case None
+        with norm_c A 
+        show ?thesis
+          by auto
       next
-	case Some
-	with normal j 
-	have "abrupt s1 = Some (Jump (Break l))"
-	  by (auto dest: absorb_Some_NoneD)
-	with brk_c A
-  	show ?thesis
-	  by auto
+        case Some
+        with normal j 
+        have "abrupt s1 = Some (Jump (Break l))"
+          by (auto dest: absorb_Some_NoneD)
+        with brk_c A
+        show ?thesis
+          by auto
       qed
     qed
     moreover 
     have "?BreakAssigned (Norm s0) (abupd (absorb j) s1) A"
     proof -
       {
-	fix l'
-	assume break: "abrupt (abupd (absorb j) s1) = Some (Jump (Break l'))"
-	with j 
-	have "l\<noteq>l'"
-	  by (cases s1) (auto dest!: absorb_Some_JumpD)
-	hence "(rmlab l (brk C)) l'= (brk C) l'"
-	  by (simp)
-	with break brk_c A
-	have 
+        fix l'
+        assume break: "abrupt (abupd (absorb j) s1) = Some (Jump (Break l'))"
+        with j 
+        have "l\<noteq>l'"
+          by (cases s1) (auto dest!: absorb_Some_JumpD)
+        hence "(rmlab l (brk C)) l'= (brk C) l'"
+          by (simp)
+        with break brk_c A
+        have 
           "(brk A l' \<subseteq> dom (locals (store (abupd (absorb j) s1))))"
-	  by (cases s1) auto
+          by (cases s1) auto
       }
       then show ?thesis
-	by simp
+        by simp
     qed
     moreover
     from res_c have "?ResAssigned (Norm s0) (abupd (absorb j) s1)"
@@ -2776,7 +2776,7 @@
       case True
       with nrm_c1 have "nrm C1 \<subseteq> dom (locals (snd s1))" by iprover
       with da_c2 obtain C2' 
-	where  da_c2': "Env\<turnstile> dom (locals (snd s1)) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
+        where  da_c2': "Env\<turnstile> dom (locals (snd s1)) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
                nrm_c2: "nrm C2 \<subseteq> nrm C2'"                  and
                brk_c2: "\<forall> l. brk C2 l \<subseteq> brk C2' l"
         by (rule da_weakenE) iprover
@@ -2785,18 +2785,18 @@
       obtain nrm_c2': "?NormalAssigned s2 C2'" and 
              brk_c2': "?BreakAssigned s1 s2 C2'" and
              res_c2 : "?ResAssigned s1 s2" 
-	by simp
+        by simp
       from nrm_c2' nrm_c2 A
       have "?NormalAssigned s2 A" 
-	by blast
+        by blast
       moreover from brk_c2' brk_c2 A
       have "?BreakAssigned s1 s2 A" 
-	by fastsimp
+        by fastsimp
       with True 
       have "?BreakAssigned (Norm s0) s2 A" by simp
       moreover from res_c2 True
       have "?ResAssigned (Norm s0) s2"
-	by simp
+        by simp
       ultimately show ?thesis by (intro conjI)
     next
       case False
@@ -2806,23 +2806,23 @@
       moreover
       have "?BreakAssigned (Norm s0) s2 A"
       proof (cases "\<exists> l. abrupt s1 = Some (Jump (Break l))")
-	case True
-	then obtain l where l: "abrupt s1 = Some (Jump (Break l))" .. 
-	with brk_c1
-	have "brk C1 l \<subseteq> dom (locals (store s1))"
-	  by simp
-	with A eq_s1_s2 
-	have "brk A l \<subseteq> dom (locals (store s2))" 
-	  by auto
-	with l eq_s1_s2
-	show ?thesis by simp
+        case True
+        then obtain l where l: "abrupt s1 = Some (Jump (Break l))" .. 
+        with brk_c1
+        have "brk C1 l \<subseteq> dom (locals (store s1))"
+          by simp
+        with A eq_s1_s2 
+        have "brk A l \<subseteq> dom (locals (store s2))" 
+          by auto
+        with l eq_s1_s2
+        show ?thesis by simp
       next
-	case False
-	with eq_s1_s2 show ?thesis by simp
+        case False
+        with eq_s1_s2 show ?thesis by simp
       qed
       moreover from False res_c1 eq_s1_s2
       have "?ResAssigned (Norm s0) s2"
-	by simp
+        by simp
       ultimately show ?thesis by (intro conjI)
     qed
   next
@@ -2853,79 +2853,79 @@
       note normal_s1 = this 
       show ?thesis
       proof (cases "the_Bool b")
-	case True
-	from eval_e normal_s1 wt_e
-	have "assigns_if True e \<subseteq> dom (locals (store s1))"
-	  by (rule assigns_if_good_approx [elim_format]) (simp add: True)
-	with s0_s1
-	have "dom (locals (store ((Norm s0)::state))) \<union> assigns_if True e \<subseteq> \<dots>"
-	  by (rule Un_least)
-	with da_c1 obtain C1'
-	  where da_c1': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
-	        nrm_c1: "nrm C1 \<subseteq> nrm C1'"                  and
+        case True
+        from eval_e normal_s1 wt_e
+        have "assigns_if True e \<subseteq> dom (locals (store s1))"
+          by (rule assigns_if_good_approx [elim_format]) (simp add: True)
+        with s0_s1
+        have "dom (locals (store ((Norm s0)::state))) \<union> assigns_if True e \<subseteq> \<dots>"
+          by (rule Un_least)
+        with da_c1 obtain C1'
+          where da_c1': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
+                nrm_c1: "nrm C1 \<subseteq> nrm C1'"                  and
                 brk_c1: "\<forall> l. brk C1 l \<subseteq> brk C1' l"
           by (rule da_weakenE) iprover
-	from If.hyps True have "PROP ?Hyp (In1r c1) s1 s2" by simp
-	with wt_c1 da_c1'
-	obtain nrm_c1': "?NormalAssigned s2 C1'" and 
+        from If.hyps True have "PROP ?Hyp (In1r c1) s1 s2" by simp
+        with wt_c1 da_c1'
+        obtain nrm_c1': "?NormalAssigned s2 C1'" and 
                brk_c1': "?BreakAssigned s1 s2 C1'" and
                res_c1: "?ResAssigned s1 s2"
-	  using G by simp
-	from nrm_c1' nrm_c1 A
-	have "?NormalAssigned s2 A" 
-	  by blast
-	moreover from brk_c1' brk_c1 A
-	have "?BreakAssigned s1 s2 A" 
-	  by fastsimp
-	with normal_s1
-	have "?BreakAssigned (Norm s0) s2 A" by simp
-	moreover from res_c1 normal_s1 have "?ResAssigned (Norm s0) s2"
-	  by simp
-	ultimately show ?thesis by (intro conjI)
+          using G by simp
+        from nrm_c1' nrm_c1 A
+        have "?NormalAssigned s2 A" 
+          by blast
+        moreover from brk_c1' brk_c1 A
+        have "?BreakAssigned s1 s2 A" 
+          by fastsimp
+        with normal_s1
+        have "?BreakAssigned (Norm s0) s2 A" by simp
+        moreover from res_c1 normal_s1 have "?ResAssigned (Norm s0) s2"
+          by simp
+        ultimately show ?thesis by (intro conjI)
       next
-	case False
-	from eval_e normal_s1 wt_e
-	have "assigns_if False e \<subseteq> dom (locals (store s1))"
-	  by (rule assigns_if_good_approx [elim_format]) (simp add: False)
-	with s0_s1
-	have "dom (locals (store ((Norm s0)::state)))\<union> assigns_if False e \<subseteq> \<dots>"
-	  by (rule Un_least)
-	with da_c2 obtain C2'
-	  where da_c2': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
-	        nrm_c2: "nrm C2 \<subseteq> nrm C2'"                  and
+        case False
+        from eval_e normal_s1 wt_e
+        have "assigns_if False e \<subseteq> dom (locals (store s1))"
+          by (rule assigns_if_good_approx [elim_format]) (simp add: False)
+        with s0_s1
+        have "dom (locals (store ((Norm s0)::state)))\<union> assigns_if False e \<subseteq> \<dots>"
+          by (rule Un_least)
+        with da_c2 obtain C2'
+          where da_c2': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
+                nrm_c2: "nrm C2 \<subseteq> nrm C2'"                  and
                 brk_c2: "\<forall> l. brk C2 l \<subseteq> brk C2' l"
           by (rule da_weakenE) iprover
-	from If.hyps False have "PROP ?Hyp (In1r c2) s1 s2" by simp
-	with wt_c2 da_c2'
-	obtain nrm_c2': "?NormalAssigned s2 C2'" and 
+        from If.hyps False have "PROP ?Hyp (In1r c2) s1 s2" by simp
+        with wt_c2 da_c2'
+        obtain nrm_c2': "?NormalAssigned s2 C2'" and 
                brk_c2': "?BreakAssigned s1 s2 C2'" and
-	       res_c2: "?ResAssigned s1 s2"
-	  using G by simp
-	from nrm_c2' nrm_c2 A
-	have "?NormalAssigned s2 A" 
-	  by blast
-	moreover from brk_c2' brk_c2 A
-	have "?BreakAssigned s1 s2 A" 
-	  by fastsimp
-	with normal_s1
-	have "?BreakAssigned (Norm s0) s2 A" by simp
-	moreover from res_c2 normal_s1 have "?ResAssigned (Norm s0) s2"
-	  by simp
-	ultimately show ?thesis by (intro conjI)
+               res_c2: "?ResAssigned s1 s2"
+          using G by simp
+        from nrm_c2' nrm_c2 A
+        have "?NormalAssigned s2 A" 
+          by blast
+        moreover from brk_c2' brk_c2 A
+        have "?BreakAssigned s1 s2 A" 
+          by fastsimp
+        with normal_s1
+        have "?BreakAssigned (Norm s0) s2 A" by simp
+        moreover from res_c2 normal_s1 have "?ResAssigned (Norm s0) s2"
+          by simp
+        ultimately show ?thesis by (intro conjI)
       qed
     next
       case False 
       then obtain abr where abr: "abrupt s1 = Some abr"
-	by (cases s1) auto
+        by (cases s1) auto
       moreover
       from eval_e _ wt_e have "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
-	by (rule eval_expression_no_jump) (simp_all add: G wf)
+        by (rule eval_expression_no_jump) (simp_all add: G wf)
       moreover
       have "s2 = s1"
       proof -
-	from abr and `G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2`
+        from abr and `G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2`
         show ?thesis  
-	  by (cases s1) simp
+          by (cases s1) simp
       qed
       ultimately show ?thesis by simp
     qed
@@ -2960,162 +2960,162 @@
       note normal_s1 = this
       show ?thesis
       proof (cases "the_Bool b")
-	case True  
-	with Loop.hyps obtain
+        case True  
+        with Loop.hyps obtain
           eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" and 
           eval_while: "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
-	  by simp
-	from Loop.hyps True 
-	have "?HypObj (In1r c) s1 s2" by simp
-	note hyp_c = this [rule_format]
-	from Loop.hyps True
-	have "?HypObj (In1r (l\<bullet> While(e) c)) (abupd (absorb (Cont l)) s2) s3"
-	  by simp
-	note hyp_while = this [rule_format]
-	from eval_e normal_s1 wt_e
-	have "assigns_if True e \<subseteq> dom (locals (store s1))"
-	  by (rule assigns_if_good_approx [elim_format]) (simp add: True)
-	with s0_s1
-	have "dom (locals (store ((Norm s0)::state))) \<union> assigns_if True e \<subseteq> \<dots>"
-	  by (rule Un_least)
-	with da_c obtain C'
-	  where da_c': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'" and
-	     nrm_C_C': "nrm C \<subseteq> nrm C'"                  and
+          by simp
+        from Loop.hyps True 
+        have "?HypObj (In1r c) s1 s2" by simp
+        note hyp_c = this [rule_format]
+        from Loop.hyps True
+        have "?HypObj (In1r (l\<bullet> While(e) c)) (abupd (absorb (Cont l)) s2) s3"
+          by simp
+        note hyp_while = this [rule_format]
+        from eval_e normal_s1 wt_e
+        have "assigns_if True e \<subseteq> dom (locals (store s1))"
+          by (rule assigns_if_good_approx [elim_format]) (simp add: True)
+        with s0_s1
+        have "dom (locals (store ((Norm s0)::state))) \<union> assigns_if True e \<subseteq> \<dots>"
+          by (rule Un_least)
+        with da_c obtain C'
+          where da_c': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'" and
+             nrm_C_C': "nrm C \<subseteq> nrm C'"                  and
              brk_C_C': "\<forall> l. brk C l \<subseteq> brk C' l"
           by (rule da_weakenE) iprover
-	from hyp_c wt_c da_c'
-	obtain nrm_C': "?NormalAssigned s2 C'" and 
+        from hyp_c wt_c da_c'
+        obtain nrm_C': "?NormalAssigned s2 C'" and 
           brk_C': "?BreakAssigned s1 s2 C'" and
           res_s2: "?ResAssigned s1 s2"
-	  using G by simp
-	show ?thesis
-	proof (cases "normal s2 \<or> abrupt s2 = Some (Jump (Cont l))")
-	  case True
-	  from Loop.prems obtain
-	    wt_while: "Env\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T" and
+          using G by simp
+        show ?thesis
+        proof (cases "normal s2 \<or> abrupt s2 = Some (Jump (Cont l))")
+          case True
+          from Loop.prems obtain
+            wt_while: "Env\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T" and
             da_while: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) 
                            \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A"
-	    by simp
-	  have "dom (locals (store ((Norm s0)::state)))
+            by simp
+          have "dom (locals (store ((Norm s0)::state)))
                   \<subseteq> dom (locals (store (abupd (absorb (Cont l)) s2)))"
-	  proof -
-	    note s0_s1
-	    also from eval_c 
-	    have "dom (locals (store s1)) \<subseteq> dom (locals (store s2))"
-	      by (rule dom_locals_eval_mono_elim)
-	    also have "\<dots> \<subseteq> dom (locals (store (abupd (absorb (Cont l)) s2)))"
-	      by simp
-	    finally show ?thesis .
-	  qed
-	  with da_while obtain A'
-	    where 
-	    da_while': "Env\<turnstile> dom (locals (store (abupd (absorb (Cont l)) s2))) 
+          proof -
+            note s0_s1
+            also from eval_c 
+            have "dom (locals (store s1)) \<subseteq> dom (locals (store s2))"
+              by (rule dom_locals_eval_mono_elim)
+            also have "\<dots> \<subseteq> dom (locals (store (abupd (absorb (Cont l)) s2)))"
+              by simp
+            finally show ?thesis .
+          qed
+          with da_while obtain A'
+            where 
+            da_while': "Env\<turnstile> dom (locals (store (abupd (absorb (Cont l)) s2))) 
                        \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A'" 
-	    and  nrm_A_A': "nrm A \<subseteq> nrm A'"                  
+            and  nrm_A_A': "nrm A \<subseteq> nrm A'"                  
             and  brk_A_A': "\<forall> l. brk A l \<subseteq> brk A' l"
-	    by (rule da_weakenE) simp
-	  with wt_while hyp_while
-	  obtain nrm_A': "?NormalAssigned s3 A'" and
+            by (rule da_weakenE) simp
+          with wt_while hyp_while
+          obtain nrm_A': "?NormalAssigned s3 A'" and
                  brk_A': "?BreakAssigned (abupd (absorb (Cont l)) s2) s3 A'" and
                  res_s3: "?ResAssigned (abupd (absorb (Cont l)) s2) s3"
-	    using G by simp 
-	  from nrm_A_A' nrm_A'
-	  have "?NormalAssigned s3 A" 
-	    by blast
-	  moreover 
-	  have "?BreakAssigned (Norm s0) s3 A" 
-	  proof -
-	    from brk_A_A' brk_A' 
-	    have "?BreakAssigned (abupd (absorb (Cont l)) s2) s3 A" 
-	      by fastsimp
-	    moreover
-	    from True have "normal (abupd (absorb (Cont l)) s2)"
-	      by (cases s2) auto
-	    ultimately show ?thesis
-	      by simp
-	  qed
-	  moreover from res_s3 True have "?ResAssigned (Norm s0) s3"
-	    by auto
-	  ultimately show ?thesis by (intro conjI)
-	next
-	  case False
-	  then obtain abr where 
-	    "abrupt s2 = Some abr" and
-	    "abrupt (abupd (absorb (Cont l)) s2) = Some abr"
-	    by auto
-	  with eval_while 
-	  have eq_s3_s2: "s3=s2"
-	    by auto
-	  with nrm_C_C' nrm_C' A
-	  have "?NormalAssigned s3 A"
-	    by fastsimp
-	  moreover
-	  from eq_s3_s2 brk_C_C' brk_C' normal_s1 A
-	  have "?BreakAssigned (Norm s0) s3 A"
-	    by fastsimp
-	  moreover 
-	  from eq_s3_s2 res_s2 normal_s1 have "?ResAssigned (Norm s0) s3"
-	    by simp
-	  ultimately show ?thesis by (intro conjI)
-	qed
+            using G by simp 
+          from nrm_A_A' nrm_A'
+          have "?NormalAssigned s3 A" 
+            by blast
+          moreover 
+          have "?BreakAssigned (Norm s0) s3 A" 
+          proof -
+            from brk_A_A' brk_A' 
+            have "?BreakAssigned (abupd (absorb (Cont l)) s2) s3 A" 
+              by fastsimp
+            moreover
+            from True have "normal (abupd (absorb (Cont l)) s2)"
+              by (cases s2) auto
+            ultimately show ?thesis
+              by simp
+          qed
+          moreover from res_s3 True have "?ResAssigned (Norm s0) s3"
+            by auto
+          ultimately show ?thesis by (intro conjI)
+        next
+          case False
+          then obtain abr where 
+            "abrupt s2 = Some abr" and
+            "abrupt (abupd (absorb (Cont l)) s2) = Some abr"
+            by auto
+          with eval_while 
+          have eq_s3_s2: "s3=s2"
+            by auto
+          with nrm_C_C' nrm_C' A
+          have "?NormalAssigned s3 A"
+            by fastsimp
+          moreover
+          from eq_s3_s2 brk_C_C' brk_C' normal_s1 A
+          have "?BreakAssigned (Norm s0) s3 A"
+            by fastsimp
+          moreover 
+          from eq_s3_s2 res_s2 normal_s1 have "?ResAssigned (Norm s0) s3"
+            by simp
+          ultimately show ?thesis by (intro conjI)
+        qed
       next
-	case False
-	with Loop.hyps have eq_s3_s1: "s3=s1"
-	  by simp
-	from eq_s3_s1 res_s1 
-	have res_s3: "?ResAssigned (Norm s0) s3"
-	  by simp
-	from eval_e True wt_e
-	have "assigns_if False e \<subseteq> dom (locals (store s1))"
-	  by (rule assigns_if_good_approx [elim_format]) (simp add: False)
-	with s0_s1
-	have "dom (locals (store ((Norm s0)::state)))\<union>assigns_if False e \<subseteq> \<dots>"
-	  by (rule Un_least)
-	hence "nrm C \<inter> 
+        case False
+        with Loop.hyps have eq_s3_s1: "s3=s1"
+          by simp
+        from eq_s3_s1 res_s1 
+        have res_s3: "?ResAssigned (Norm s0) s3"
+          by simp
+        from eval_e True wt_e
+        have "assigns_if False e \<subseteq> dom (locals (store s1))"
+          by (rule assigns_if_good_approx [elim_format]) (simp add: False)
+        with s0_s1
+        have "dom (locals (store ((Norm s0)::state)))\<union>assigns_if False e \<subseteq> \<dots>"
+          by (rule Un_least)
+        hence "nrm C \<inter> 
                (dom (locals (store ((Norm s0)::state))) \<union> assigns_if False e)
                \<subseteq> dom (locals (store s1))"
-	  by (rule subset_Intr)
-	with normal_s1 A eq_s3_s1
-	have "?NormalAssigned s3 A"
-	  by simp
-	moreover
-	from normal_s1 eq_s3_s1
-	have "?BreakAssigned (Norm s0) s3 A"
-	  by simp
-	moreover note res_s3
-	ultimately show ?thesis by (intro conjI)
+          by (rule subset_Intr)
+        with normal_s1 A eq_s3_s1
+        have "?NormalAssigned s3 A"
+          by simp
+        moreover
+        from normal_s1 eq_s3_s1
+        have "?BreakAssigned (Norm s0) s3 A"
+          by simp
+        moreover note res_s3
+        ultimately show ?thesis by (intro conjI)
       qed
     next
       case False
       then obtain abr where abr: "abrupt s1 = Some abr"
-	by (cases s1) auto
+        by (cases s1) auto
       moreover
       from eval_e _ wt_e have no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
-	by (rule eval_expression_no_jump) (simp_all add: wf G)
+        by (rule eval_expression_no_jump) (simp_all add: wf G)
       moreover
       have eq_s3_s1: "s3=s1"
       proof (cases "the_Bool b")
-	case True  
-	with Loop.hyps obtain
+        case True  
+        with Loop.hyps obtain
           eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" and 
           eval_while: "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
-	  by simp
-	from eval_c abr have "s2=s1" by auto
-	moreover from calculation no_jmp have "abupd (absorb (Cont l)) s2=s2"
-	  by (cases s1) (simp add: absorb_def)
-	ultimately show ?thesis
-	  using eval_while abr
-	  by auto
+          by simp
+        from eval_c abr have "s2=s1" by auto
+        moreover from calculation no_jmp have "abupd (absorb (Cont l)) s2=s2"
+          by (cases s1) (simp add: absorb_def)
+        ultimately show ?thesis
+          using eval_while abr
+          by auto
       next
-	case False
-	with Loop.hyps show ?thesis by simp
+        case False
+        with Loop.hyps show ?thesis by simp
       qed
       moreover
       from eq_s3_s1 res_s1 
       have res_s3: "?ResAssigned (Norm s0) s3"
-	by simp
+        by simp
       ultimately show ?thesis
-	by simp 
+        by simp 
     qed
   next 
     case (Jmp s j Env T A)
@@ -3143,21 +3143,21 @@
       by (elim da_elim_cases)
     from Throw.prems
       obtain eT where wt_e: "Env\<turnstile>e\<Colon>-eT"
-	by (elim wt_elim_cases)
+        by (elim wt_elim_cases)
     have "?NormalAssigned (abupd (throw a) s1) A"
       by (cases s1) (simp add: throw_def)
     moreover
     have "?BreakAssigned (Norm s0) (abupd (throw a) s1) A"
     proof -
       from G Throw.hyps have eval_e: "prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1" 
-	by (simp (no_asm_simp))
+        by (simp (no_asm_simp))
       from eval_e _ wt_e 
       have "\<And> l. abrupt s1 \<noteq> Some (Jump (Break l))"
-	by (rule eval_expression_no_jump) (simp_all add: wf G) 
+        by (rule eval_expression_no_jump) (simp_all add: wf G) 
       hence "\<And> l. abrupt (abupd (throw a) s1) \<noteq> Some (Jump (Break l))"
-	by (cases s1) (simp add: throw_def abrupt_if_def)
+        by (cases s1) (simp add: throw_def abrupt_if_def)
       thus ?thesis
-	by simp
+        by simp
     qed
     moreover
     from wt_e da_e G have "?ResAssigned (Norm s0) s1"
@@ -3191,49 +3191,49 @@
     proof (cases "normal s1")
       case True
       with nrm_C1 have "nrm C1 \<inter> nrm C2 \<subseteq> dom (locals (store s1))"
-	by auto
+        by auto
       moreover
       have "s3=s1"
       proof -
-	from sxalloc True have eq_s2_s1: "s2=s1"
-	  by (cases s1) (auto elim: sxalloc_elim_cases)
-	with True have "\<not>  G,s2\<turnstile>catch C"
-	  by (simp add: catch_def)
-	with Try.hyps have "s3=s2"
-	  by simp
-	with eq_s2_s1 show ?thesis by simp
+        from sxalloc True have eq_s2_s1: "s2=s1"
+          by (cases s1) (auto elim: sxalloc_elim_cases)
+        with True have "\<not>  G,s2\<turnstile>catch C"
+          by (simp add: catch_def)
+        with Try.hyps have "s3=s2"
+          by simp
+        with eq_s2_s1 show ?thesis by simp
       qed
       ultimately show ?thesis
-	using True A res_s1 by simp
+        using True A res_s1 by simp
     next
       case False
       note not_normal_s1 = this
       show ?thesis
       proof (cases "\<exists> l. abrupt s1 = Some (Jump (Break l))")
-	case True
-	then obtain l where l: "abrupt s1 = Some (Jump (Break l))"
-	  by auto
-	with brk_C1 have "(brk C1 \<Rightarrow>\<inter> brk C2) l \<subseteq> dom (locals (store s1))"
-	  by auto
-	moreover have "s3=s1"
-	proof -
-	  from sxalloc l have eq_s2_s1: "s2=s1"
-	    by (cases s1) (auto elim: sxalloc_elim_cases)
-	  with l have "\<not>  G,s2\<turnstile>catch C"
-	    by (simp add: catch_def)
-	  with Try.hyps have "s3=s2"
-	    by simp
-	  with eq_s2_s1 show ?thesis by simp
-	qed
-	ultimately show ?thesis
-	  using l A res_s1 by simp
+        case True
+        then obtain l where l: "abrupt s1 = Some (Jump (Break l))"
+          by auto
+        with brk_C1 have "(brk C1 \<Rightarrow>\<inter> brk C2) l \<subseteq> dom (locals (store s1))"
+          by auto
+        moreover have "s3=s1"
+        proof -
+          from sxalloc l have eq_s2_s1: "s2=s1"
+            by (cases s1) (auto elim: sxalloc_elim_cases)
+          with l have "\<not>  G,s2\<turnstile>catch C"
+            by (simp add: catch_def)
+          with Try.hyps have "s3=s2"
+            by simp
+          with eq_s2_s1 show ?thesis by simp
+        qed
+        ultimately show ?thesis
+          using l A res_s1 by simp
       next
-	case False
-	note abrupt_no_break = this
-	show ?thesis
-	proof (cases "G,s2\<turnstile>catch C")
-	  case True
-	  with Try.hyps have "?HypObj (In1r c2) (new_xcpt_var vn s2) s3"
+        case False
+        note abrupt_no_break = this
+        show ?thesis
+        proof (cases "G,s2\<turnstile>catch C")
+          case True
+          with Try.hyps have "?HypObj (In1r c2) (new_xcpt_var vn s2) s3"
             by simp
           note hyp_c2 = this [rule_format]
           have "(dom (locals (store ((Norm s0)::state))) \<union> {VName vn}) 
@@ -3279,9 +3279,9 @@
             with brk_C2' A show ?thesis 
               by fastsimp
           qed
-	  moreover
-	  from resAss_s3 have "?ResAssigned (Norm s0) s3"
-	    by (cases s2) ( simp add: new_xcpt_var_def)
+          moreover
+          from resAss_s3 have "?ResAssigned (Norm s0) s3"
+            by (cases s2) ( simp add: new_xcpt_var_def)
           ultimately show ?thesis by (intro conjI)
         next
           case False
@@ -3291,23 +3291,23 @@
           obtain "\<not> normal s2" 
                  "\<forall> l. abrupt s2 \<noteq> Some (Jump (Break l))"
             by - (rule sxalloc_cases,auto)
-	  ultimately obtain 
-	    "?NormalAssigned s3 A" and "?BreakAssigned (Norm s0) s3 A"
-	    by (cases s2) auto
-	  moreover have "?ResAssigned (Norm s0) s3"
-	  proof (cases "abrupt s1 = Some (Jump Ret)")
-	    case True
-	    with sxalloc have "s2=s1"
-	      by (elim sxalloc_cases) auto
-	    with res_s1 eq_s3_s2 show ?thesis by simp
-	  next
-	    case False
-	    with sxalloc
-	    have "abrupt s2 \<noteq> Some (Jump Ret)"
-	      by (rule sxalloc_no_jump) 
-	    with eq_s3_s2 show ?thesis
-	      by simp
-	  qed
+          ultimately obtain 
+            "?NormalAssigned s3 A" and "?BreakAssigned (Norm s0) s3 A"
+            by (cases s2) auto
+          moreover have "?ResAssigned (Norm s0) s3"
+          proof (cases "abrupt s1 = Some (Jump Ret)")
+            case True
+            with sxalloc have "s2=s1"
+              by (elim sxalloc_cases) auto
+            with res_s1 eq_s3_s2 show ?thesis by simp
+          next
+            case False
+            with sxalloc
+            have "abrupt s2 \<noteq> Some (Jump Ret)"
+              by (rule sxalloc_no_jump) 
+            with eq_s3_s2 show ?thesis
+              by simp
+          qed
           ultimately show ?thesis by (intro conjI)
         qed
       qed
@@ -3349,7 +3349,7 @@
       with wt_c2 da_C2' G
       obtain nrmAss_C2': "?NormalAssigned s2 C2'" and
              brkAss_C2': "?BreakAssigned (Norm s1) s2 C2'" and
-	     resAss_s2': "?ResAssigned (Norm s1) s2"
+             resAss_s2': "?ResAssigned (Norm s1) s2"
         by simp
       from nrmAss_C2' nrm_C2' have "?NormalAssigned s2 C2"
         by blast
@@ -3449,23 +3449,23 @@
       assume abr_s3: "abrupt s3 = Some (Jump Ret)"
       have "Result \<in> dom (locals (store s3))"
       proof (cases "x1 = Some (Jump Ret)")
-	case True
-	note ret_x1 = this
-	with resAss_s1 have res_s1: "Result \<in> dom (locals s1)"
-	  by simp
-	moreover have "dom (locals (store ((Norm s1)::state))) 
+        case True
+        note ret_x1 = this
+        with resAss_s1 have res_s1: "Result \<in> dom (locals s1)"
+          by simp
+        moreover have "dom (locals (store ((Norm s1)::state))) 
                          \<subseteq> dom (locals (store s2))"
-	  by (rule dom_locals_eval_mono_elim) (rule Fin.hyps)
-	ultimately have "Result \<in> dom (locals (store s2))"
-	  by - (rule subsetD,auto)
-	with res_s1 s3 show ?thesis
-	  by simp
+          by (rule dom_locals_eval_mono_elim) (rule Fin.hyps)
+        ultimately have "Result \<in> dom (locals (store s2))"
+          by - (rule subsetD,auto)
+        with res_s1 s3 show ?thesis
+          by simp
       next
-	case False
-	with s3 abr_s3 obtain  "abrupt s2 = Some (Jump Ret)" and "s3=s2"
-	  by (cases s2) (simp add: abrupt_if_def)
-	with resAss_s2 show ?thesis
-	  by simp
+        case False
+        with s3 abr_s3 obtain  "abrupt s2 = Some (Jump Ret)" and "s3=s2"
+          by (cases s2) (simp add: abrupt_if_def)
+        with resAss_s2 show ?thesis
+          by simp
       qed
     }
     hence "?ResAssigned (Norm s0) s3"
@@ -3494,7 +3494,7 @@
       case True
       with Init.hyps have "s3=Norm s0" by simp
       thus ?thesis
-	using nrm_A by simp
+        using nrm_A by simp
     next
       case False
       from Init.hyps False G
@@ -3503,27 +3503,27 @@
                        \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1" and
              eval_init: "prg Env\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<rightarrow> s2" and
              s3: "s3=(set_lvars (locals (store s1))) s2"
-	by simp
+        by simp
       have "?NormalAssigned s3 A"
       proof
-	show "nrm A \<subseteq> dom (locals (store s3))"
-	proof -
-	  from nrm_A have "nrm A \<subseteq> dom (locals (init_class_obj G C s0))"
-	    by simp
-	  also from eval_initC have "\<dots> \<subseteq> dom (locals (store s1))"
-	    by (rule dom_locals_eval_mono_elim) simp
-	  also from s3 have "\<dots> \<subseteq> dom (locals (store s3))"
-	    by (cases s1) (cases s2, simp add: init_lvars_def2)
-	  finally show ?thesis .
-	qed
+        show "nrm A \<subseteq> dom (locals (store s3))"
+        proof -
+          from nrm_A have "nrm A \<subseteq> dom (locals (init_class_obj G C s0))"
+            by simp
+          also from eval_initC have "\<dots> \<subseteq> dom (locals (store s1))"
+            by (rule dom_locals_eval_mono_elim) simp
+          also from s3 have "\<dots> \<subseteq> dom (locals (store s3))"
+            by (cases s1) (cases s2, simp add: init_lvars_def2)
+          finally show ?thesis .
+        qed
       qed
       moreover
       from eval 
       have "\<And> j. abrupt s3 \<noteq> Some (Jump j)"
-	by (rule eval_statement_no_jump) (auto simp add: wf c G)
+        by (rule eval_statement_no_jump) (auto simp add: wf c G)
       then obtain "?BreakAssigned (Norm s0) s3 A" 
               and "?ResAssigned (Norm s0) s3"
-	by simp
+        by simp
       ultimately show ?thesis by (intro conjI)
     qed
   next 
@@ -3552,7 +3552,7 @@
       fix j have "abrupt s2 \<noteq> Some (Jump j)"
       proof -
         have eval: "prg Env\<turnstile> Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
-	  unfolding G by (rule eval.NewC NewC.hyps)+
+          unfolding G by (rule eval.NewC NewC.hyps)+
         from NewC.prems 
         obtain T' where  "T=Inl T'"
           by (elim wt_elim_cases) simp
@@ -3617,7 +3617,7 @@
       fix j have "abrupt s3 \<noteq> Some (Jump j)"
       proof -
         have eval: "prg Env\<turnstile> Norm s0 \<midarrow>New elT[e]-\<succ>Addr a\<rightarrow> s3"
-	  unfolding G by (rule eval.NewA NewA.hyps)+
+          unfolding G by (rule eval.NewA NewA.hyps)+
         from NewA.prems 
         obtain T' where  "T=Inl T'"
           by (elim wt_elim_cases) simp
@@ -3662,7 +3662,7 @@
       fix j have "abrupt s2 \<noteq> Some (Jump j)"
       proof -
         have eval: "prg Env\<turnstile> Norm s0 \<midarrow>Cast cT e-\<succ>v\<rightarrow> s2"
-	  unfolding G by (rule eval.Cast Cast.hyps)+
+          unfolding G by (rule eval.Cast Cast.hyps)+
         from Cast.prems 
         obtain T' where  "T=Inl T'"
           by (elim wt_elim_cases) simp
@@ -3764,7 +3764,7 @@
           with ass_if CondAnd  
           have "assigns_if True (BinOp CondAnd e1 e2) 
                  \<subseteq> dom (locals (store s2))"
-	    by simp
+            by simp
           thus ?thesis by blast
         next
           case False
@@ -3900,7 +3900,7 @@
       with wt_v da_v G obtain 
         "?NormalAssigned s1 A" and
         "?BreakAssigned (Norm s0) s1 A" and
-	"?ResAssigned (Norm s0) s1"
+        "?ResAssigned (Norm s0) s1"
         by simp
       thus ?thesis by (intro conjI)
     qed
@@ -3925,56 +3925,56 @@
       note hyp_e = `PROP ?Hyp (In1l e) s1 s2`
       show "nrm A \<subseteq> dom (locals (store (assign upd v s2)))"
       proof (cases "\<exists> vn. var = LVar vn")
-	case True
-	then obtain vn where vn: "var=LVar vn"..
-	from Ass.prems obtain E where
-	  da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E" and
-	  nrm_A: "nrm A = nrm E \<union> {vn}"
-	  by (elim da_elim_cases) (insert vn,auto)
-	obtain E' where
-	  da_e': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'" and
-	  E_E': "nrm E \<subseteq> nrm E'"
-	proof -
-	  have "dom (locals (store ((Norm s0)::state))) 
+        case True
+        then obtain vn where vn: "var=LVar vn"..
+        from Ass.prems obtain E where
+          da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E" and
+          nrm_A: "nrm A = nrm E \<union> {vn}"
+          by (elim da_elim_cases) (insert vn,auto)
+        obtain E' where
+          da_e': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'" and
+          E_E': "nrm E \<subseteq> nrm E'"
+        proof -
+          have "dom (locals (store ((Norm s0)::state))) 
                   \<subseteq> dom (locals (store s1))"
-	    by (rule dom_locals_eval_mono_elim) (rule Ass.hyps)
-	  with da_e show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	from G eval_var vn 
-	have eval_lvar: "G\<turnstile>Norm s0 \<midarrow>LVar vn=\<succ>(w, upd)\<rightarrow> s1" 
-	  by simp
-	then have upd: "upd = snd (lvar vn (store s1))"
-	  by cases (cases "lvar vn (store s1)",simp)
-	have "nrm E \<subseteq> dom (locals (store (assign upd v s2)))"
-	proof -
-	  from hyp_e wt_e da_e' G normal_s2
-	  have "nrm E' \<subseteq> dom (locals (store s2))"
-	    by simp
-	  also
-	  from upd
-	  have "dom (locals (store s2))  \<subseteq> dom (locals (store (upd v s2)))"
-	    by (simp add: lvar_def) blast
-	  hence "dom (locals (store s2)) 
-	             \<subseteq> dom (locals (store (assign upd v s2)))"
-	    by (rule dom_locals_assign_mono)
-	  finally
-	  show ?thesis using E_E' 
-	    by blast
-	qed
-	moreover
-	from upd normal_s2
-	have "{vn} \<subseteq> dom (locals (store (assign upd v s2)))"
-	  by (auto simp add: assign_def Let_def lvar_def upd split: split_split)
-	ultimately
-	show "nrm A \<subseteq> \<dots>"
-	  by (rule Un_least [elim_format]) (simp add: nrm_A)
+            by (rule dom_locals_eval_mono_elim) (rule Ass.hyps)
+          with da_e show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        from G eval_var vn 
+        have eval_lvar: "G\<turnstile>Norm s0 \<midarrow>LVar vn=\<succ>(w, upd)\<rightarrow> s1" 
+          by simp
+        then have upd: "upd = snd (lvar vn (store s1))"
+          by cases (cases "lvar vn (store s1)",simp)
+        have "nrm E \<subseteq> dom (locals (store (assign upd v s2)))"
+        proof -
+          from hyp_e wt_e da_e' G normal_s2
+          have "nrm E' \<subseteq> dom (locals (store s2))"
+            by simp
+          also
+          from upd
+          have "dom (locals (store s2))  \<subseteq> dom (locals (store (upd v s2)))"
+            by (simp add: lvar_def) blast
+          hence "dom (locals (store s2)) 
+                     \<subseteq> dom (locals (store (assign upd v s2)))"
+            by (rule dom_locals_assign_mono)
+          finally
+          show ?thesis using E_E' 
+            by blast
+        qed
+        moreover
+        from upd normal_s2
+        have "{vn} \<subseteq> dom (locals (store (assign upd v s2)))"
+          by (auto simp add: assign_def Let_def lvar_def upd split: split_split)
+        ultimately
+        show "nrm A \<subseteq> \<dots>"
+          by (rule Un_least [elim_format]) (simp add: nrm_A)
       next
-	case False
-	from Ass.prems obtain V where
-	  da_var: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>var\<rangle>\<guillemotright> V" and
-	  da_e:   "Env\<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
-	  by (elim da_elim_cases) (insert False,simp+)
+        case False
+        from Ass.prems obtain V where
+          da_var: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>var\<rangle>\<guillemotright> V" and
+          da_e:   "Env\<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
+          by (elim da_elim_cases) (insert False,simp+)
         from hyp_var wt_var da_var G normal_s1
         have "nrm V \<subseteq> dom (locals (store s1))"
           by simp
@@ -4077,7 +4077,7 @@
                        \<subseteq> dom (locals (store s1))"
           by (rule dom_locals_eval_mono_elim) (rule Cond.hyps)
         have eval_e0: "prg Env\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1"
-	  unfolding G by (rule Cond.hyps)
+          unfolding G by (rule Cond.hyps)
         have normal_s1: "normal s1"
           by (rule eval_no_abrupt_lemma [rule_format]) (rule Cond.hyps, rule normal_s2)
         show ?thesis
--- a/src/HOL/Bali/Eval.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Eval.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -241,7 +241,7 @@
 
   assign     :: "('a \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> 'a \<Rightarrow> state \<Rightarrow> state"
  "assign f v \<equiv> \<lambda>(x,s). let (x',s') = (if x = None then f v else id) (x,s)
-		   in  (x',if x' = None then s' else s)"
+                   in  (x',if x' = None then s' else s)"
 
 (*
 lemma assign_Norm_Norm [simp]: 
@@ -343,7 +343,7 @@
 
 constdefs
   init_lvars :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> inv_mode \<Rightarrow> val \<Rightarrow> val list \<Rightarrow>
-		   state \<Rightarrow> state"
+                   state \<Rightarrow> state"
  "init_lvars G C sig mode a' pvs 
    \<equiv> \<lambda> (x,s). 
       let m = mthd (the (methd G C sig));
@@ -399,7 +399,7 @@
  "fvar C stat fn a' s 
     \<equiv> let (oref,xf) = if stat then (Stat C,id)
                               else (Heap (the_Addr a'),np a');
-	          n = Inl (fn,C); 
+                  n = Inl (fn,C); 
                   f = (\<lambda>v. supd (upd_gobj oref n v)) 
       in ((the (values (the (globs (store s) oref)) n),f),abupd xf s)"
 
@@ -481,17 +481,17 @@
   Abrupt: 
   "G\<turnstile>(Some x,s) \<midarrow>halloc oi\<succ>undefined\<rightarrow> (Some x,s)"
 
-| New:	  "\<lbrakk>new_Addr (heap s) = Some a; 
-	    (x,oi') = (if atleast_free (heap s) (Suc (Suc 0)) then (None,oi)
-		       else (Some (Xcpt (Loc a)),CInst (SXcpt OutOfMemory)))\<rbrakk>
+| 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)"
+            G\<turnstile>Norm s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,init_obj G oi' (Heap a) s)"
 
 inductive 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"
+  Norm:  "G\<turnstile> Norm              s   \<midarrow>sxalloc\<rightarrow>  Norm             s"
 
 | Jmp:   "G\<turnstile>(Some (Jump j), s)  \<midarrow>sxalloc\<rightarrow> (Some (Jump j), s)"
 
@@ -500,7 +500,7 @@
 | XcptL: "G\<turnstile>(Some (Xcpt (Loc a) ),s)  \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s)"
 
 | SXcpt: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>halloc (CInst (SXcpt xn))\<succ>a\<rightarrow> (x,s1)\<rbrakk> \<Longrightarrow>
-	  G\<turnstile>(Some (Xcpt (Std xn)),s0) \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s1)"
+          G\<turnstile>(Some (Xcpt (Std xn)),s0) \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s1)"
 
 
 inductive
@@ -509,7 +509,7 @@
   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)
+                    val  list ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
   for G::prog
 where
 
@@ -528,23 +528,23 @@
 --{* 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>
-				  G\<turnstile>Norm s0 \<midarrow>Expr e\<rightarrow> s1"
+| Expr: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+                                  G\<turnstile>Norm s0 \<midarrow>Expr e\<rightarrow> s1"
 
 | Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<rightarrow> s1\<rbrakk> \<Longrightarrow>
                                 G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<rightarrow> abupd (absorb l) s1"
   --{* cf. 14.2 *}
-| 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"
+| Comp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<rightarrow> s1;
+          G\<turnstile>     s1 \<midarrow>c2 \<rightarrow> s2\<rbrakk> \<Longrightarrow>
+                                 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<rightarrow> s2"
 
   --{* cf. 14.8.2 *}
-| If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
-	  G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<rbrakk> \<Longrightarrow>
-		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<rightarrow> s2"
+| If:   "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
+          G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+                       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<rightarrow> s2"
 
   --{* cf. 14.10, 14.10.1 *}
   
@@ -554,40 +554,40 @@
      the state before the iterative evaluation of the while statement.
      A break jump is handled by the Lab Statement @{text "Lab l (while\<dots>)"}.
   *}
-| Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
-	  if the_Bool b 
+| 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"
+             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)"
    
   --{* cf. 14.16 *}
 | Throw: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1\<rbrakk> \<Longrightarrow>
-				 G\<turnstile>Norm s0 \<midarrow>Throw e\<rightarrow> abupd (throw a') s1"
+                                 G\<turnstile>Norm s0 \<midarrow>Throw e\<rightarrow> abupd (throw a') s1"
 
   --{* cf. 14.18.1 *}
-| Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; 
-	  if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2\<rbrakk> \<Longrightarrow>
-		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(C vn) c2\<rightarrow> s3"
+| Try:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; 
+          if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2\<rbrakk> \<Longrightarrow>
+                  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(C vn) c2\<rightarrow> s3"
 
   --{* cf. 14.18.2 *}
-| Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1,s1);
-	  G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2;
+| 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) 
               else abupd (abrupt_if (x1\<noteq>None) x1) s2) \<rbrakk> 
           \<Longrightarrow>
           G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<rightarrow> s3"
   --{* cf. 12.4.2, 8.5 *}
-| Init:	"\<lbrakk>the (class G C) = c;
-	  if inited C (globs s0) then s3 = Norm s0
-	  else (G\<turnstile>Norm (init_class_obj G C s0) 
-		  \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1 \<and>
-	       G\<turnstile>set_lvars empty s1 \<midarrow>init c\<rightarrow> s2 \<and> s3 = restore_lvars s1 s2)\<rbrakk> 
+| Init: "\<lbrakk>the (class G C) = c;
+          if inited C (globs s0) then s3 = Norm s0
+          else (G\<turnstile>Norm (init_class_obj G C s0) 
+                  \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1 \<and>
+               G\<turnstile>set_lvars empty s1 \<midarrow>init c\<rightarrow> s2 \<and> s3 = restore_lvars s1 s2)\<rbrakk> 
               \<Longrightarrow>
-		 G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s3"
+                 G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s3"
    --{* This class initialisation rule is a little bit inaccurate. Look at the
       exact sequence:
       (1) The current class object (the static fields) are initialised
@@ -613,27 +613,27 @@
 --{* evaluation of expressions *}
 
   --{* cf. 15.8.1, 12.4.1 *}
-| NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1;
-	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
-	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
+| NewC: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1;
+          G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+                                  G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
 
   --{* cf. 15.9.1, 12.4.1 *}
-| NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<rightarrow> s2; 
-	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
-	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<rightarrow> s3"
+| NewA: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<rightarrow> s2; 
+          G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
+                                G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<rightarrow> s3"
 
   --{* cf. 15.15 *}
-| Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
-	  s2 = abupd (raise_if (\<not>G,store s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
-			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<rightarrow> s2"
+| Cast: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
+          s2 = abupd (raise_if (\<not>G,store s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
+                                G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<rightarrow> s2"
 
   --{* cf. 15.19.2 *}
-| Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
-	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
-			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<rightarrow> s1"
+| 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> 
          \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>(eval_unop unop v)\<rightarrow> s1"
@@ -648,18 +648,18 @@
 | Super: "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<rightarrow> Norm s"
 
   --{* cf. 15.2 *}
-| Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<rightarrow> s1\<rbrakk> \<Longrightarrow>
-	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<rightarrow> s1"
+| 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"
+                                   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"
+                            G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<rightarrow> s2"
 
 
 -- {* The interplay of  @{term Call}, @{term Methd} and @{term Body}:
@@ -681,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;
@@ -694,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>
-				G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<rightarrow> s1"
+| 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 
@@ -713,34 +713,34 @@
 --{* 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;
-	  (v,s2') = fvar statDeclC stat fn a 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"
+          G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v\<rightarrow> s3"
  --{* The accessibility check is after @{term fvar}, to keep it simple. 
     @{term fvar} already tests for the absence of a null-pointer reference 
     in case of an instance field
   *}
 
   --{* cf. 15.12.1, 15.25.1 *}
-| 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'"
+| AVar: "\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2;
+          (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
+                      G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<rightarrow> s2'"
 
 
 --{* evaluation of expression lists *}
 
   --{* cf. 15.11.4.2 *}
 | Nil:
-				    "G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<rightarrow> Norm s0"
+                                    "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"
+                                   G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<rightarrow> s2"
 
 (* Rearrangement of premisses:
 [0,1(Abrupt),2(Skip),8(Jmp),4(Lab),30(Nil),31(Cons),27(LVar),17(Cast),18(Inst),
@@ -770,11 +770,11 @@
   "G\<turnstile>(Norm    s) \<midarrow>halloc oi\<succ>a\<rightarrow> s'"
 
 inductive_cases sxalloc_elim_cases:
- 	"G\<turnstile> Norm                 s  \<midarrow>sxalloc\<rightarrow> s'"
+        "G\<turnstile> 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 (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'"
+        "G\<turnstile>(Some (Xcpt (Std xn)),s) \<midarrow>sxalloc\<rightarrow> s'"
 inductive_cases sxalloc_cases: "G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'"
 
 lemma sxalloc_elim_cases2: "\<lbrakk>G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s';  
@@ -796,36 +796,36 @@
 
 inductive_cases 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 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>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')"
+        "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]
 declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *}
--- a/src/HOL/Bali/Evaln.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Evaln.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -29,7 +29,7 @@
 *}
 
 inductive
-  evaln	:: "[prog, state, term, nat, vals, state] \<Rightarrow> bool"
+  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)
@@ -37,7 +37,7 @@
     ("_\<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"
+  and execn     :: "[prog, state, stmt, nat, state] \<Rightarrow> bool"
     ("_\<turnstile>_ \<midarrow>_\<midarrow>_\<rightarrow> _"     [61,61,65,   61,61] 60)
   for G :: prog
 where
@@ -54,39 +54,39 @@
 
 --{* 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;
-	  (v,s2') = fvar statDeclC stat fn a 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"
+          G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v\<midarrow>n\<rightarrow> s3"
 
-| AVar:	"\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1 ; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2; 
-	  (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
-	              G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<midarrow>n\<rightarrow> s2'"
+| AVar: "\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1 ; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2; 
+          (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
+                      G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<midarrow>n\<rightarrow> s2'"
 
 
 
 
 --{* evaluation of expressions *}
 
-| NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1;
-	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
-	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<midarrow>n\<rightarrow> s2"
+| NewC: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1;
+          G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+                                  G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<midarrow>n\<rightarrow> s2"
 
-| NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<midarrow>n\<rightarrow> s2; 
-	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
-	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<midarrow>n\<rightarrow> s3"
+| NewA: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<midarrow>n\<rightarrow> s2; 
+          G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
+                                G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<midarrow>n\<rightarrow> s3"
 
-| Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
-	  s2 = abupd (raise_if (\<not>G,snd s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
-			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<midarrow>n\<rightarrow> s2"
+| Cast: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
+          s2 = abupd (raise_if (\<not>G,snd s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
+                                G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<midarrow>n\<rightarrow> s2"
 
-| Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1;
-	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
-			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<midarrow>n\<rightarrow> s1"
+| 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> 
          \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>(eval_unop unop v)\<midarrow>n\<rightarrow> s1"
@@ -96,20 +96,20 @@
             \<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>
-	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<midarrow>n\<rightarrow> s1"
+| 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"
+                                   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"
+                            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;
@@ -120,9 +120,9 @@
     G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<midarrow>n\<rightarrow> (restore_lvars s2 s4)"
 
 | Methd:"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
-				G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<midarrow>Suc n\<rightarrow> s1"
+                                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 
@@ -133,63 +133,63 @@
 --{* evaluation of expression lists *}
 
 | Nil:
-				"G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<midarrow>n\<rightarrow> Norm s0"
+                                "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"
+                             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>
-				  G\<turnstile>Norm s0 \<midarrow>Expr e\<midarrow>n\<rightarrow> s1"
+| Expr: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
+                                  G\<turnstile>Norm s0 \<midarrow>Expr e\<midarrow>n\<rightarrow> s1"
 
 | Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<midarrow>n\<rightarrow> s1\<rbrakk> \<Longrightarrow>
                              G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<midarrow>n\<rightarrow> abupd (absorb l) s1"
 
-| Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1;
-	  G\<turnstile>     s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
-				 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<midarrow>n\<rightarrow> s2"
+| Comp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1;
+          G\<turnstile>     s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+                                 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<midarrow>n\<rightarrow> s2"
 
-| If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
-	  G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
-		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<midarrow>n\<rightarrow> s2"
+| If:   "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
+          G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2\<rbrakk> \<Longrightarrow>
+                       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<midarrow>n\<rightarrow> s2"
 
-| Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1;
-	  if the_Bool b 
+| 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"
+             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)"
   
 | 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"
+                                 G\<turnstile>Norm s0 \<midarrow>Throw e\<midarrow>n\<rightarrow> abupd (throw a') s1"
 
-| Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2;
-	  if G,s2\<turnstile>catch tn then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3 else s3 = s2\<rbrakk>
+| 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"
+                  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(tn vn) c2\<midarrow>n\<rightarrow> s3"
 
-| Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n\<rightarrow> (x1,s1);
-	  G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n\<rightarrow> s2;
+| 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;
-	  if inited C (globs s0) then s3 = Norm s0
-	  else (G\<turnstile>Norm (init_class_obj G C s0)
-	          \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1 \<and>
-	        G\<turnstile>set_lvars empty s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2 \<and> 
+| Init: "\<lbrakk>the (class G C) = c;
+          if inited C (globs s0) then s3 = Norm s0
+          else (G\<turnstile>Norm (init_class_obj G C s0)
+                  \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1 \<and>
+                G\<turnstile>set_lvars empty s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2 \<and> 
                 s3 = restore_lvars s1 s2)\<rbrakk>
           \<Longrightarrow>
-		 G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s3"
+                 G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s3"
 monos
   if_bool_eq_conj
 
@@ -203,37 +203,37 @@
 inductive_cases 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> (v, s')"
-	"G\<turnstile>Norm s \<midarrow>In1r Skip                      \<succ>\<midarrow>n\<rightarrow> (x, s')"
+        "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>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>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] 
@@ -569,7 +569,7 @@
     "if the_Bool b 
         then (G\<turnstile>s1 \<midarrow>c\<midarrow>n2\<rightarrow> s2 \<and> 
               G\<turnstile>(abupd (absorb (Cont l)) s2)\<midarrow>l\<bullet> While(e) c\<midarrow>n2\<rightarrow> s3)
-	else s3 = s1"
+        else s3 = s1"
     by simp (iprover intro: evaln_nonstrict le_maxI1 le_maxI2)
   ultimately
   have "G\<turnstile>Norm s0 \<midarrow>l\<bullet> While(e) c\<midarrow>max n1 n2\<rightarrow> s3"
@@ -629,8 +629,8 @@
   moreover from Init.hyps obtain n where
       "if inited C (globs s0) then s3 = Norm s0
        else (G\<turnstile>Norm (init_class_obj G C s0)
-	      \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1 \<and>
-	           G\<turnstile>set_lvars empty s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2 \<and> 
+              \<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1 \<and>
+                   G\<turnstile>set_lvars empty s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2 \<and> 
                    s3 = restore_lvars s1 s2)"
     by (auto intro: evaln_nonstrict le_maxI1 le_maxI2)
   ultimately have "G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s3"
--- a/src/HOL/Bali/Example.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Example.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -170,11 +170,11 @@
   Ext_foo  :: mdecl
  "Ext_foo  \<equiv> (foo_sig, 
               \<lparr>access=Public,static=False,pars=[z],resT=Class Ext,
-	       mbody=\<lparr>lcls=[]
+               mbody=\<lparr>lcls=[]
                      ,stmt=Expr({Ext,Ext,False}Cast (Class Ext) (!!z)..vee := 
-       	                                                     Lit (Intg 1)) ;;
+                                                             Lit (Intg 1)) ;;
                                 Return (Lit Null)\<rparr>
-	      \<rparr>)"
+              \<rparr>)"
 
 constdefs
   
@@ -184,7 +184,7 @@
 BaseCl :: "class"
 "BaseCl \<equiv> \<lparr>access=Public,
            cfields=[(arr, \<lparr>access=Public,static=True ,type=PrimT Boolean.[]\<rparr>),
-	            (vee, \<lparr>access=Public,static=False,type=Iface HasFoo    \<rparr>)],
+                    (vee, \<lparr>access=Public,static=False,type=Iface HasFoo    \<rparr>)],
            methods=[Base_foo],
            init=Expr(arr_viewed_from Base Base 
                      :=New (PrimT Boolean)[Lit (Intg 2)]),
--- a/src/HOL/Bali/Name.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Name.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -7,7 +7,7 @@
 theory Name imports Basis begin
 
 (* cf. 6.5 *) 
-typedecl tnam	--{* ordinary type name, i.e. class or interface name *}
+typedecl tnam   --{* ordinary type name, i.e. class or interface name *}
 typedecl pname  --{* package name *}
 typedecl mname  --{* method name *}
 typedecl vname  --{* variable or field name *}
@@ -28,10 +28,10 @@
   "VName n" == "EName (VNam n)"
   "Result"  == "EName Res"
 
-datatype xname		--{* names of standard exceptions *}
-	= Throwable
-	| NullPointer | OutOfMemory | ClassCast   
-	| NegArrSize  | IndOutBound | ArrStore
+datatype xname          --{* names of standard exceptions *}
+        = Throwable
+        | NullPointer | OutOfMemory | ClassCast   
+        | NegArrSize  | IndOutBound | ArrStore
 
 lemma xn_cases: 
   "xn = Throwable   \<or> xn = NullPointer \<or>  
@@ -42,10 +42,10 @@
 done
 
 
-datatype tname	--{* type names for standard classes and other type names *}
-	= Object'
-	| SXcpt'   xname
-	| TName   tnam
+datatype tname  --{* type names for standard classes and other type names *}
+        = Object'
+        | SXcpt'   xname
+        | TName   tnam
 
 record   qtname = --{* qualified tname cf. 6.5.3, 6.5.4*}
           pid :: pname  
--- a/src/HOL/Bali/State.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/State.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -19,15 +19,15 @@
 section "objects"
 
 datatype  obj_tag =     --{* tag for generic object   *}
-	  CInst qtname  --{* class instance           *}
-	| Arr  ty int   --{* array with component type and length *}
+          CInst qtname  --{* class instance           *}
+        | Arr  ty int   --{* array with component type and length *}
     --{* | CStat qtname   the tag is irrelevant for a class object,
-			   i.e. the static fields of a class,
+                           i.e. the static fields of a class,
                            since its type is given already by the reference to 
                            it (see below) *}
 
-types	vn   = "fspec + int"                    --{* variable name      *}
-record	obj  = 
+types   vn   = "fspec + int"                    --{* variable name      *}
+record  obj  = 
           tag :: "obj_tag"                      --{* generalized object *}
           "values" :: "(vn, val) table"      
 
@@ -214,12 +214,12 @@
 
 section "stores"
 
-types	globs               --{* global variables: heap and static variables *}
-	= "(oref , obj) table"
-	heap
-	= "(loc  , obj) table"
-(*	locals                   
-	= "(lname, val) table" *) (* defined in Value.thy local variables *)
+types   globs               --{* global variables: heap and static variables *}
+        = "(oref , obj) table"
+        heap
+        = "(loc  , obj) table"
+(*      locals                   
+        = "(lname, val) table" *) (* defined in Value.thy local variables *)
 
 translations
  "globs"  <= (type) "(oref , obj) table"
@@ -227,7 +227,7 @@
 (*  "locals" <= (type) "(lname, val) table" *)
 
 datatype st = (* pure state, i.e. contents of all variables *)
-	 st globs locals
+         st globs locals
 
 subsection "access"
 
@@ -481,7 +481,7 @@
 primrec "the_Std (Std x) = x"
 
 
-	
+        
 
 constdefs
   abrupt_if    :: "bool \<Rightarrow> abopt \<Rightarrow> abopt \<Rightarrow> abopt"
--- a/src/HOL/Bali/Term.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Term.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -67,8 +67,8 @@
         | Ret         --{* return from method *}
 
 datatype xcpt        --{* exception *}
-	= Loc loc    --{* location of allocated execption object *}
-	| Std xname  --{* intermediate standard exception, see Eval.thy *}
+        = Loc loc    --{* location of allocated execption object *}
+        | Std xname  --{* intermediate standard exception, see Eval.thy *}
 
 datatype error
        =  AccessViolation  --{* Access to a member that isn't permitted *}
@@ -92,12 +92,12 @@
  "locals" <= (type) "(lname, val) table"
 
 datatype inv_mode                  --{* invocation mode for method calls *}
-	= Static                   --{* static *}
-	| SuperM                   --{* super  *}
-	| IntVir                   --{* interface or virtual *}
+        = Static                   --{* static *}
+        | SuperM                   --{* super  *}
+        | IntVir                   --{* interface or virtual *}
 
 record  sig =              --{* signature of a method, cf. 8.4.2  *}
-	  name ::"mname"   --{* acutally belongs to Decl.thy *}
+          name ::"mname"   --{* acutally belongs to Decl.thy *}
           parTs::"ty list"        
 
 translations
@@ -142,7 +142,7 @@
 *}
 
 datatype var
-	= LVar lname --{* local variable (incl. parameters) *}
+        = LVar lname --{* local variable (incl. parameters) *}
         | FVar qtname qtname bool expr vname ("{_,_,_}_.._"[10,10,10,85,99]90)
                      --{* class field *}
                      --{* @{term "{accC,statDeclC,stat}e..fn"}   *}
@@ -153,7 +153,7 @@
                      --{* @{text stat}: static or instance field?*}
                      --{* @{text e}: reference to object*}
                      --{* @{text fn}: field name*}
-	| AVar expr expr ("_.[_]"[90,10   ]90)
+        | AVar expr expr ("_.[_]"[90,10   ]90)
                      --{* array component *}
                      --{* @{term "e1.[e2]"}: e1 array reference; e2 index *}
         | InsInitV stmt var 
@@ -161,21 +161,21 @@
                      --{* of var (technical term for smallstep semantics.)*}
 
 and expr
-	= NewC qtname         --{* class instance creation *}
-	| NewA ty expr ("New _[_]"[99,10   ]85) 
+        = NewC qtname         --{* class instance creation *}
+        | NewA ty expr ("New _[_]"[99,10   ]85) 
                               --{* array creation *} 
-	| Cast ty expr        --{* type cast  *}
-	| Inst expr ref_ty ("_ InstOf _"[85,99] 85)   
+        | Cast ty expr        --{* type cast  *}
+        | Inst expr ref_ty ("_ InstOf _"[85,99] 85)   
                               --{* instanceof *}     
-	| Lit  val              --{* literal value, references not allowed *}
+        | Lit  val              --{* literal value, references not allowed *}
         | UnOp unop expr        --{* unary operation *}
         | BinOp binop expr expr --{* binary operation *}
         
-	| Super               --{* special Super keyword *}
-	| Acc  var            --{* variable access *}
-	| Ass  var expr       ("_:=_"   [90,85   ]85)
+        | Super               --{* special Super keyword *}
+        | Acc  var            --{* variable access *}
+        | Ass  var expr       ("_:=_"   [90,85   ]85)
                               --{* variable assign *} 
-	| Cond expr expr expr ("_ ? _ : _" [85,85,80]80) --{* conditional *}  
+        | Cond expr expr expr ("_ ? _ : _" [85,85,80]80) --{* conditional *}  
         | Call qtname ref_ty inv_mode expr mname "(ty list)" "(expr list)"  
             ("{_,_,_}_\<cdot>_'( {_}_')"[10,10,10,85,99,10,10]85) 
                     --{* method call *} 
@@ -198,25 +198,25 @@
         | Callee locals expr  --{* save callers locals in callee-Frame *}
                               --{* (technical term for smallstep semantics) *}
 and  stmt
-	= Skip                  --{* empty      statement *}
-	| Expr  expr            --{* expression statement *}
+        = Skip                  --{* empty      statement *}
+        | Expr  expr            --{* expression statement *}
         | Lab   jump stmt       ("_\<bullet> _" [      99,66]66)
                                 --{* labeled statement; handles break *}
-	| Comp  stmt stmt       ("_;; _"                  [      66,65]65)
-	| If'   expr stmt stmt  ("If'(_') _ Else _"       [   80,79,79]70)
-	| Loop  label expr stmt ("_\<bullet> While'(_') _"        [   99,80,79]70)
+        | Comp  stmt stmt       ("_;; _"                  [      66,65]65)
+        | If'   expr stmt stmt  ("If'(_') _ Else _"       [   80,79,79]70)
+        | Loop  label expr stmt ("_\<bullet> While'(_') _"        [   99,80,79]70)
         | Jmp jump              --{* break, continue, return *}
-	| Throw expr
+        | Throw expr
         | TryC  stmt qtname vname stmt ("Try _ Catch'(_ _') _"  [79,99,80,79]70)
              --{* @{term "Try c1 Catch(C vn) c2"} *} 
              --{* @{text c1}: block were exception may be thrown *}
              --{* @{text C}:  execption class to catch *}
              --{* @{text vn}: local name for exception used in @{text c2}*}
              --{* @{text c2}: block to execute when exception is cateched*}
-	| Fin  stmt  stmt        ("_ Finally _"               [      79,79]70)
+        | Fin  stmt  stmt        ("_ Finally _"               [      79,79]70)
         | FinA abopt stmt       --{* Save abruption of first statement *} 
                                 --{* technical term  for smallstep sem.) *}
-	| Init  qtname          --{* class initialization *}
+        | Init  qtname          --{* class initialization *}
 
 
 text {*
--- a/src/HOL/Bali/Trans.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Trans.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -124,11 +124,11 @@
     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')"
@@ -143,7 +143,7 @@
           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)"*)
+(*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> 
--- a/src/HOL/Bali/Type.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Type.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -15,21 +15,21 @@
 \end{itemize}
 *}
 
-datatype prim_ty       	--{* primitive type, cf. 4.2 *}
-	= Void    	--{* result type of void methods *}
-	| Boolean
-	| Integer
+datatype prim_ty        --{* primitive type, cf. 4.2 *}
+        = Void          --{* result type of void methods *}
+        | Boolean
+        | Integer
 
 
-datatype ref_ty		--{* reference type, cf. 4.3 *}
-	= NullT		--{* null type, cf. 4.1 *}
-	| IfaceT qtname	--{* interface type *}
-	| ClassT qtname	--{* class type *}
-	| ArrayT ty	--{* array type *}
+datatype ref_ty         --{* reference type, cf. 4.3 *}
+        = NullT         --{* null type, cf. 4.1 *}
+        | IfaceT qtname --{* interface type *}
+        | ClassT qtname --{* class type *}
+        | ArrayT ty     --{* array type *}
 
-and ty	    		--{* any type, cf. 4.1 *}
-	= PrimT prim_ty	--{* primitive type *}
-	| RefT  ref_ty	--{* reference type *}
+and ty                  --{* any type, cf. 4.1 *}
+        = PrimT prim_ty --{* primitive type *}
+        | RefT  ref_ty  --{* reference type *}
 
 translations
   "prim_ty" <= (type) "Type.prim_ty"
@@ -37,16 +37,16 @@
   "ty"      <= (type) "Type.ty"
 
 syntax
-	 NT	:: "       \<spacespace> ty"
-	 Iface	:: "qtname  \<Rightarrow> ty"
-	 Class	:: "qtname  \<Rightarrow> ty"
-	 Array	:: "ty     \<Rightarrow> ty"	("_.[]" [90] 90)
+         NT     :: "       \<spacespace> ty"
+         Iface  :: "qtname  \<Rightarrow> ty"
+         Class  :: "qtname  \<Rightarrow> ty"
+         Array  :: "ty     \<Rightarrow> ty"    ("_.[]" [90] 90)
 
 translations
-	"NT"      == "RefT   NullT"
-	"Iface I" == "RefT (IfaceT I)"
-	"Class C" == "RefT (ClassT C)"
-	"T.[]"    == "RefT (ArrayT T)"
+        "NT"      == "RefT   NullT"
+        "Iface I" == "RefT (IfaceT I)"
+        "Class C" == "RefT (ClassT C)"
+        "T.[]"    == "RefT (ArrayT T)"
 
 constdefs
   the_Class :: "ty \<Rightarrow> qtname"
--- a/src/HOL/Bali/TypeRel.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/TypeRel.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -59,12 +59,12 @@
 
 translations
 
-	"G\<turnstile>I \<prec>I1 J" == "(I,J) \<in> subint1 G"
-	"G\<turnstile>I \<preceq>I  J" == "(I,J) \<in>(subint1 G)^*" --{* cf. 9.1.3 *}
-  	(* Defined in Decl.thy:
+        "G\<turnstile>I \<prec>I1 J" == "(I,J) \<in> subint1 G"
+        "G\<turnstile>I \<preceq>I  J" == "(I,J) \<in>(subint1 G)^*" --{* cf. 9.1.3 *}
+        (* Defined in Decl.thy:
         "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" == "(C,D) \<in> subcls1 G"
-	"G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" 
-	*)
+        "G\<turnstile>C \<preceq>\<^sub>C  D" == "(C,D) \<in>(subcls1 G)^*" 
+        *)
         "G\<turnstile>C \<leadsto>1 I" == "(C,I) \<in> implmt1 G"
 
 
@@ -202,12 +202,12 @@
       assume subcls_D_C: "G\<turnstile>D \<prec>\<^sub>C C"
       show "False"
       proof -
-	from subcls_D_C subcls1_C_Z
-	have "G\<turnstile>D \<prec>\<^sub>C Z"
-	  by (auto dest: r_into_trancl trancl_trans)
-	with nsubcls_D_Z
-	show ?thesis
-	  by (rule notE)
+        from subcls_D_C subcls1_C_Z
+        have "G\<turnstile>D \<prec>\<^sub>C Z"
+          by (auto dest: r_into_trancl trancl_trans)
+        with nsubcls_D_Z
+        show ?thesis
+          by (rule notE)
       qed
     qed
   qed  
@@ -243,16 +243,16 @@
       assume eq_C_D: "C=D"
       show "False"
       proof -
-	from subcls1_C_Z eq_C_D 
-	have "G\<turnstile>D \<prec>\<^sub>C Z"
-	  by (auto)
-	also
-	from subcls_Z_D ws   
-	have "\<not> G\<turnstile>D \<prec>\<^sub>C Z"
-	  by (rule subcls_acyclic)
-	ultimately 
-	show ?thesis 
-	  by - (rule notE)
+        from subcls1_C_Z eq_C_D 
+        have "G\<turnstile>D \<prec>\<^sub>C Z"
+          by (auto)
+        also
+        from subcls_Z_D ws   
+        have "\<not> G\<turnstile>D \<prec>\<^sub>C Z"
+          by (rule subcls_acyclic)
+        ultimately 
+        show ?thesis 
+          by - (rule notE)
       qed
     qed
   qed
@@ -298,12 +298,12 @@
       case Eq
       with ws subcls_C_D
       show ?thesis 
-	by (auto dest: subcls_irrefl)
+        by (auto dest: subcls_irrefl)
     next
       case Subcls
       with nsubcls_D_C
       show ?thesis
-	by blast
+        by blast
     qed
   qed
 qed
@@ -387,7 +387,7 @@
 
 inductive
  --{*widening, viz. method invocation conversion, cf. 5.3
-			    i.e. kind of syntactic subtyping *}
+                            i.e. kind of syntactic subtyping *}
   widen :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<preceq>_" [71,71,71] 70)
   for G :: prog
 where
@@ -619,8 +619,8 @@
 | 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"
+           (\<lambda>(md, mh   ) (md',mh').\<spacespace>G\<turnstile>mrt mh\<preceq>mrt mh') \<Longrightarrow>
+           \<not>G\<turnstile>I\<preceq>I J         \<spacespace>\<spacespace>\<spacespace>\<Longrightarrow> G\<turnstile>     Iface I\<succ>Iface J"
 | array:   "G\<turnstile>RefT S\<succ>RefT T   \<spacespace>\<Longrightarrow> G\<turnstile>   RefT S.[]\<succ>RefT T.[]"
 
 (*unused*)
@@ -637,7 +637,7 @@
 by (ind_cases "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"
+                                  \<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
 by (ind_cases "G\<turnstile>S\<succ>PrimT pt")
 
 text {* 
--- a/src/HOL/Bali/TypeSafe.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/TypeSafe.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -366,12 +366,12 @@
       case CInst
       with modeIntVir obj_props
       show ?thesis 
-	by (auto dest!: widen_Array2 split add: split_if)
+        by (auto dest!: widen_Array2 split add: split_if)
     next
       case Arr
       from Arr obtain T where "obj_ty obj = T.[]" by (blast dest: obj_ty_Arr1)
       moreover from Arr have "obj_class obj = Object" 
-	by (blast dest: obj_class_Arr1)
+        by (blast dest: obj_class_Arr1)
       moreover note modeIntVir obj_props wf 
       ultimately show ?thesis by (auto dest!: widen_Array )
     qed
@@ -411,22 +411,22 @@
       case Static
       with wf dynlookup statI is_iface 
       show ?thesis
-	by (auto simp add: invocation_declclass_def dynlookup_def 
+        by (auto simp add: invocation_declclass_def dynlookup_def 
                            dynimethd_def dynmethd_C_C 
-	            intro: dynmethd_declclass
-	            dest!: wf_imethdsD
+                    intro: dynmethd_declclass
+                    dest!: wf_imethdsD
                      dest: table_of_map_SomeI
                     split: split_if_asm)
-    next	
+    next        
       case SuperM
       with not_SuperM show ?thesis ..
     next
       case IntVir
       with wf dynlookup IfaceT invC_prop show ?thesis 
-	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
+        by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
                            DynT_prop_def
-	            intro: methd_declclass dynmethd_declclass
-	            split: split_if_asm)
+                    intro: methd_declclass dynmethd_declclass
+                    split: split_if_asm)
     qed
   next
     case ClassT
@@ -445,9 +445,9 @@
       case IntVir
       with wf ClassT dynlookup statC_prop invC_prop 
       show ?thesis
-	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
+        by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
                            DynT_prop_def
-	            intro: dynmethd_declclass)
+                    intro: dynmethd_declclass)
     qed
   next
     case ArrayT
@@ -455,7 +455,7 @@
     proof (cases mode)
       case Static
       with wf ArrayT dynlookup show ?thesis
-	by (auto simp add: invocation_declclass_def dynlookup_def 
+        by (auto simp add: invocation_declclass_def dynlookup_def 
                            dynimethd_def dynmethd_C_C
                     intro: dynmethd_declclass
                      dest: table_of_map_SomeI)
@@ -465,7 +465,7 @@
     next
       case IntVir
       with wf ArrayT dynlookup invC_prop show ?thesis
-	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
+        by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
                            DynT_prop_def dynmethd_C_C
                     intro: dynmethd_declclass
                      dest: table_of_map_SomeI)
@@ -564,8 +564,8 @@
                                             dynimethd_def)
     from sm wf wt_e not_Null False invC_prop' obtain "dm" where
                     dm: "methd G invC sig = Some dm"          and
-	eq_declC_sm_dm:"statDeclT = ClassT (declclass dm)"  and
-	     eq_mheads:"sm=mhead (mthd dm) "
+        eq_declC_sm_dm:"statDeclT = ClassT (declclass dm)"  and
+             eq_mheads:"sm=mhead (mthd dm) "
       by - (drule static_mheadsD, (force dest: accmethd_SomeD)+)
     then have static: "is_static dm = is_static sm" by - (auto)
     with declC invC dynlookup_static dm
@@ -1339,24 +1339,24 @@
     proof (cases "(tab(x\<mapsto>y)) vn \<noteq> Some z")
       case True
       with some ys have "(tab'(x\<mapsto>y)(xs[\<mapsto>]tl)) vn = Some z"
-	by (fastsimp intro: Cons.hyps)
+        by (fastsimp intro: Cons.hyps)
       with ys show ?thesis 
-	by simp
+        by simp
     next
       case False
       hence tabx_z: "(tab(x\<mapsto>y)) vn = Some z" by blast
       moreover
       from tabx_z tab_not_z
       have "(tab'(x\<mapsto>y)) vn = Some z" 
-	by (rule map_upd_in_expansion_map_swap)
+        by (rule map_upd_in_expansion_map_swap)
       ultimately
       have "(tab(x\<mapsto>y)) vn =(tab'(x\<mapsto>y)) vn"
-	by simp
+        by simp
       hence "(tab(x\<mapsto>y)(xs[\<mapsto>]tl)) vn = (tab'(x\<mapsto>y)(xs[\<mapsto>]tl)) vn"
-	by (rule map_upds_cong_ext)
+        by (rule map_upds_cong_ext)
       with some ys
       show ?thesis 
-	by simp
+        by simp
     qed
   qed
 qed
@@ -1460,12 +1460,12 @@
       case (EName en)
       show ?thesis
       proof (cases en)
-	case Res
-	with EName el_in_list show ?thesis by (simp add: dom_def)
+        case Res
+        with EName el_in_list show ?thesis by (simp add: dom_def)
       next
-	case (VNam vn)
-	with EName el_in_list show ?thesis 
-	  by (auto simp add: dom_def dest: map_upds_cut_irrelevant)
+        case (VNam vn)
+        with EName el_in_list show ?thesis 
+          by (auto simp add: dom_def dest: map_upds_cut_irrelevant)
       qed
     qed
   qed
@@ -1482,12 +1482,12 @@
       case (EName en)
       show ?thesis
       proof (cases en)
-	case Res
-	with EName el_in_hd_tl show ?thesis by (simp add: dom_def)
+        case Res
+        with EName el_in_hd_tl show ?thesis by (simp add: dom_def)
       next
-	case (VNam vn)
-	with EName el_in_hd_tl show ?thesis 
-	  by (auto simp add: dom_def intro: map_upds_Some_expand 
+        case (VNam vn)
+        with EName el_in_hd_tl show ?thesis 
+          by (auto simp add: dom_def intro: map_upds_Some_expand 
                                             map_upds_Some_insert)
       qed
     qed
@@ -1595,7 +1595,7 @@
     and wt_binop: "wt_binop G binop e1T e2T" 
     and conf_s0: "s0\<Colon>\<preceq>(G,L)"  
     and normal_s1: "normal s1"
-    and	eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
+    and eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
     and conf_v1: "G,store s1\<turnstile>v1\<Colon>\<preceq>e1T"
     and wf: "wf_prog G"
   shows "\<exists> E2. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) 
@@ -1625,62 +1625,62 @@
       assume condAnd: "binop=CondAnd"
       have ?thesis
       proof -
-	from da obtain E2' where
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        from da obtain E2' where
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
              \<turnstile> dom (locals (store s0)) \<union> assigns_if True e1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> E2'"
-	  by cases (simp add: condAnd)+
-	moreover
-	have "dom (locals (store s0)) 
+          by cases (simp add: condAnd)+
+        moreover
+        have "dom (locals (store s0)) 
           \<union> assigns_if True e1 \<subseteq> dom (locals (store s1))"
-	proof -
-	  from condAnd wt_binop have e1T: "e1T=PrimT Boolean"
-	    by simp
-	  with normal_s1 conf_v1 obtain b where "v1=Bool b"
-	    by (auto dest: conf_Boolean)
-	  with True condAnd
-	  have v1: "v1=Bool True"
-	    by simp
-	  from eval_e1 normal_s1 
-	  have "assigns_if True e1 \<subseteq> dom (locals (store s1))"
-	    by (rule assigns_if_good_approx' [elim_format])
-	       (insert wt_e1, simp_all add: e1T v1)
-	  with s0_s1 show ?thesis by (rule Un_least)
-	qed
-	ultimately
-	show ?thesis
-	  using that by (cases rule: da_weakenE) (simp add: True)
+        proof -
+          from condAnd wt_binop have e1T: "e1T=PrimT Boolean"
+            by simp
+          with normal_s1 conf_v1 obtain b where "v1=Bool b"
+            by (auto dest: conf_Boolean)
+          with True condAnd
+          have v1: "v1=Bool True"
+            by simp
+          from eval_e1 normal_s1 
+          have "assigns_if True e1 \<subseteq> dom (locals (store s1))"
+            by (rule assigns_if_good_approx' [elim_format])
+               (insert wt_e1, simp_all add: e1T v1)
+          with s0_s1 show ?thesis by (rule Un_least)
+        qed
+        ultimately
+        show ?thesis
+          using that by (cases rule: da_weakenE) (simp add: True)
       qed
     }
     moreover
     { 
       assume condOr: "binop=CondOr"
       have ?thesis
-	(* Beweis durch Analogie/Example/Pattern?, True\<rightarrow>False; And\<rightarrow>Or *)
+        (* Beweis durch Analogie/Example/Pattern?, True\<rightarrow>False; And\<rightarrow>Or *)
       proof -
-	from da obtain E2' where
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        from da obtain E2' where
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
               \<turnstile> dom (locals (store s0)) \<union> assigns_if False e1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> E2'"
-	  by cases (simp add: condOr)+
-	moreover
-	have "dom (locals (store s0)) 
+          by cases (simp add: condOr)+
+        moreover
+        have "dom (locals (store s0)) 
                      \<union> assigns_if False e1 \<subseteq> dom (locals (store s1))"
-	proof -
-	  from condOr wt_binop have e1T: "e1T=PrimT Boolean"
-	    by simp
-	  with normal_s1 conf_v1 obtain b where "v1=Bool b"
-	    by (auto dest: conf_Boolean)
-	  with True condOr
-	  have v1: "v1=Bool False"
-	    by simp
-	  from eval_e1 normal_s1 
-	  have "assigns_if False e1 \<subseteq> dom (locals (store s1))"
-	    by (rule assigns_if_good_approx' [elim_format])
-	       (insert wt_e1, simp_all add: e1T v1)
-	  with s0_s1 show ?thesis by (rule Un_least)
-	qed
-	ultimately
-	show ?thesis
-	  using that by (rule da_weakenE) (simp add: True)
+        proof -
+          from condOr wt_binop have e1T: "e1T=PrimT Boolean"
+            by simp
+          with normal_s1 conf_v1 obtain b where "v1=Bool b"
+            by (auto dest: conf_Boolean)
+          with True condOr
+          have v1: "v1=Bool False"
+            by simp
+          from eval_e1 normal_s1 
+          have "assigns_if False e1 \<subseteq> dom (locals (store s1))"
+            by (rule assigns_if_good_approx' [elim_format])
+               (insert wt_e1, simp_all add: e1T v1)
+          with s0_s1 show ?thesis by (rule Un_least)
+        qed
+        ultimately
+        show ?thesis
+          using that by (rule da_weakenE) (simp add: True)
       qed
     }
     moreover
@@ -1688,16 +1688,16 @@
       assume notAndOr: "binop\<noteq>CondAnd" "binop\<noteq>CondOr"
       have ?thesis
       proof -
-	from da notAndOr obtain E1' where
+        from da notAndOr obtain E1' where
           da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                   \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1'"
-	  and da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1' \<guillemotright>In1l e2\<guillemotright> A"
-	  by cases simp+
-	from eval_e1 wt_e1 da_e1 wf normal_s1 
-	have "nrm E1' \<subseteq> dom (locals (store s1))"
-	  by (cases rule: da_good_approxE') iprover
-	with da_e2 show ?thesis
-	  using that by (rule da_weakenE) (simp add: True)
+          and da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1' \<guillemotright>In1l e2\<guillemotright> A"
+          by cases simp+
+        from eval_e1 wt_e1 da_e1 wf normal_s1 
+        have "nrm E1' \<subseteq> dom (locals (store s1))"
+          by (cases rule: da_good_approxE') iprover
+        with da_e2 show ?thesis
+          using that by (rule da_weakenE) (simp add: True)
       qed
     }
     ultimately show ?thesis
@@ -1822,23 +1822,23 @@
       case False
       with eval_c2 have "s2=s1" by auto
       with conf_s1 error_free_s1 False wt show ?thesis
-	by simp
+        by simp
     next
       case True
       obtain C2' where 
-	"\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In1r c2\<guillemotright> C2'"
+        "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In1r c2\<guillemotright> C2'"
       proof -
-	from eval_c1 wt_c1 da_c1 wf True
-	have "nrm C1 \<subseteq> dom (locals (store s1))"
-	  by (cases rule: da_good_approxE') iprover
-	with da_c2 show thesis
-	  by (rule da_weakenE) (rule that)
+        from eval_c1 wt_c1 da_c1 wf True
+        have "nrm C1 \<subseteq> dom (locals (store s1))"
+          by (cases rule: da_good_approxE') iprover
+        with da_c2 show thesis
+          by (rule da_weakenE) (rule that)
       qed
       with conf_s1 wt_c2 
       obtain "s2\<Colon>\<preceq>(G, L)" and "error_free s2"
-	by (rule hyp_c2 [elim_format]) (simp add: error_free_s1)
+        by (rule hyp_c2 [elim_format]) (simp add: error_free_s1)
       thus ?thesis
-	using wt by simp
+        using wt by simp
     qed
   next
     case (If s0 e b s1 c1 c2 s2 L accC T A)
@@ -1879,32 +1879,32 @@
       case False
       with eval_then_else have "s2=s1" by auto
       with conf_s1 error_free_s1 False wt show ?thesis
-	by simp
+        by simp
     next
       case True
       obtain C' where
-	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
           (dom (locals (store s1)))\<guillemotright>In1r (if the_Bool b then c1 else c2)\<guillemotright> C'"
       proof -
-	from eval_e have 
-	  "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
-	  by (rule dom_locals_eval_mono_elim)
+        from eval_e have 
+          "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+          by (rule dom_locals_eval_mono_elim)
         moreover
-	from eval_e True wt_e 
-	have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
-	  by (rule assigns_if_good_approx')
-	ultimately 
-	have "dom (locals (store ((Norm s0)::state))) 
+        from eval_e True wt_e 
+        have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+          by (rule assigns_if_good_approx')
+        ultimately 
+        have "dom (locals (store ((Norm s0)::state))) 
                 \<union> assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
-	  by (rule Un_least)
-	with da_then_else show thesis
-	  by (rule da_weakenE) (rule that)
+          by (rule Un_least)
+        with da_then_else show thesis
+          by (rule da_weakenE) (rule that)
       qed
       with conf_s1 wt_then_else  
       obtain "s2\<Colon>\<preceq>(G, L)" and "error_free s2"
-	by (rule hyp_then_else [elim_format]) (simp add: error_free_s1)
+        by (rule hyp_then_else [elim_format]) (simp add: error_free_s1)
       with wt show ?thesis
-	by simp
+        by simp
     qed
     -- {* Note that we don't have to show that @{term b} really is a boolean 
           value. With @{term the_Bool} we enforce to get a value of boolean 
@@ -1945,106 +1945,106 @@
       note normal_s1 = this
       show ?thesis
       proof (cases "the_Bool b")
-	case True
-	with Loop.hyps  obtain
+        case True
+        with Loop.hyps  obtain
           eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" and 
           eval_while: "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
-	  by simp 
-	have "?TypeSafeObj s1 s2 (In1r c) \<diamondsuit>"
-	  using Loop.hyps True by simp
-	note hyp_c = this [rule_format]
-	have "?TypeSafeObj (abupd (absorb (Cont l)) s2)
+          by simp 
+        have "?TypeSafeObj s1 s2 (In1r c) \<diamondsuit>"
+          using Loop.hyps True by simp
+        note hyp_c = this [rule_format]
+        have "?TypeSafeObj (abupd (absorb (Cont l)) s2)
           s3 (In1r (l\<bullet> While(e) c)) \<diamondsuit>"
-	  using Loop.hyps True by simp
-	note hyp_w = this [rule_format]
-	from eval_e have 
-	  s0_s1: "dom (locals (store ((Norm s0)::state)))
+          using Loop.hyps True by simp
+        note hyp_w = this [rule_format]
+        from eval_e have 
+          s0_s1: "dom (locals (store ((Norm s0)::state)))
                     \<subseteq> dom (locals (store s1))"
-	  by (rule dom_locals_eval_mono_elim)
-	obtain C' where
-	  "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(dom (locals (store s1)))\<guillemotright>In1r c\<guillemotright> C'" 
-	proof -
-	  note s0_s1
+          by (rule dom_locals_eval_mono_elim)
+        obtain C' where
+          "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(dom (locals (store s1)))\<guillemotright>In1r c\<guillemotright> C'" 
+        proof -
+          note s0_s1
           moreover
-	  from eval_e normal_s1 wt_e 
-	  have "assigns_if True e \<subseteq> dom (locals (store s1))"
-	    by (rule assigns_if_good_approx' [elim_format]) (simp add: True)
-	  ultimately 
-	  have "dom (locals (store ((Norm s0)::state))) 
+          from eval_e normal_s1 wt_e 
+          have "assigns_if True e \<subseteq> dom (locals (store s1))"
+            by (rule assigns_if_good_approx' [elim_format]) (simp add: True)
+          ultimately 
+          have "dom (locals (store ((Norm s0)::state))) 
                  \<union> assigns_if True e \<subseteq> dom (locals (store s1))"
-	    by (rule Un_least)
-	  with da_c show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	with conf_s1 wt_c  
-	obtain conf_s2:  "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
-	  by (rule hyp_c [elim_format]) (simp add: error_free_s1)
-	from error_free_s2 
-	have error_free_ab_s2: "error_free (abupd (absorb (Cont l)) s2)"
-	  by simp
-	from conf_s2 have "abupd (absorb (Cont l)) s2 \<Colon>\<preceq>(G, L)"
-	  by (cases s2) (auto intro: conforms_absorb)
-	moreover note wt
-	moreover
-	obtain A' where 
+            by (rule Un_least)
+          with da_c show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        with conf_s1 wt_c  
+        obtain conf_s2:  "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
+          by (rule hyp_c [elim_format]) (simp add: error_free_s1)
+        from error_free_s2 
+        have error_free_ab_s2: "error_free (abupd (absorb (Cont l)) s2)"
+          by simp
+        from conf_s2 have "abupd (absorb (Cont l)) s2 \<Colon>\<preceq>(G, L)"
+          by (cases s2) (auto intro: conforms_absorb)
+        moreover note wt
+        moreover
+        obtain A' where 
           "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
               dom (locals(store (abupd (absorb (Cont l)) s2)))
                 \<guillemotright>In1r (l\<bullet> While(e) c)\<guillemotright> A'"
-	proof -
-	  note s0_s1
-	  also from eval_c 
-	  have "dom (locals (store s1)) \<subseteq> dom (locals (store s2))"
-	    by (rule dom_locals_eval_mono_elim)
-	  also have "\<dots> \<subseteq> dom (locals (store (abupd (absorb (Cont l)) s2)))"
-	    by simp
-	  finally
+        proof -
+          note s0_s1
+          also from eval_c 
+          have "dom (locals (store s1)) \<subseteq> dom (locals (store s2))"
+            by (rule dom_locals_eval_mono_elim)
+          also have "\<dots> \<subseteq> dom (locals (store (abupd (absorb (Cont l)) s2)))"
+            by simp
+          finally
           have "dom (locals (store ((Norm s0)::state))) \<subseteq> \<dots>" .
-	  with da show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	ultimately obtain "s3\<Colon>\<preceq>(G, L)" and "error_free s3"
-	  by (rule hyp_w [elim_format]) (simp add: error_free_ab_s2)
-	with wt show ?thesis
-	  by simp
+          with da show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        ultimately obtain "s3\<Colon>\<preceq>(G, L)" and "error_free s3"
+          by (rule hyp_w [elim_format]) (simp add: error_free_ab_s2)
+        with wt show ?thesis
+          by simp
       next
-	case False
-	with Loop.hyps have "s3=s1" by simp
-	with conf_s1 error_free_s1 wt
-	show ?thesis
-	  by simp
+        case False
+        with Loop.hyps have "s3=s1" by simp
+        with conf_s1 error_free_s1 wt
+        show ?thesis
+          by simp
       qed
     next
       case False
       have "s3=s1"
       proof -
-	from False obtain abr where abr: "abrupt s1 = Some abr"
-	  by (cases s1) auto
-	from eval_e _ wt_e have no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
-	  by (rule eval_expression_no_jump 
+        from False obtain abr where abr: "abrupt s1 = Some abr"
+          by (cases s1) auto
+        from eval_e _ wt_e have no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+          by (rule eval_expression_no_jump 
                [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified]) 
              (simp_all add: wf)
-	    
-	show ?thesis
-	proof (cases "the_Bool b")
-	  case True  
-	  with Loop.hyps obtain
+            
+        show ?thesis
+        proof (cases "the_Bool b")
+          case True  
+          with Loop.hyps obtain
             eval_c: "G\<turnstile>s1 \<midarrow>c\<rightarrow> s2" and 
             eval_while: "G\<turnstile>abupd (absorb (Cont l)) s2 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
-	    by simp
-	  from eval_c abr have "s2=s1" by auto
-	  moreover from calculation no_jmp have "abupd (absorb (Cont l)) s2=s2"
-	    by (cases s1) (simp add: absorb_def)
-	  ultimately show ?thesis
-	    using eval_while abr
-	    by auto
-	next
-	  case False
-	  with Loop.hyps show ?thesis by simp
-	qed
+            by simp
+          from eval_c abr have "s2=s1" by auto
+          moreover from calculation no_jmp have "abupd (absorb (Cont l)) s2=s2"
+            by (cases s1) (simp add: absorb_def)
+          ultimately show ?thesis
+            using eval_while abr
+            by auto
+        next
+          case False
+          with Loop.hyps show ?thesis by simp
+        qed
       qed
       with conf_s1 error_free_s1 wt
       show ?thesis
-	by simp
+        by simp
     qed
   next
     case (Jmp s j L accC T A)
@@ -2123,53 +2123,53 @@
       case False
       from sx_alloc wf
       have eq_s2_s1: "s2=s1"
-	by (rule sxalloc_type_sound [elim_format])
-	   (insert False, auto split: option.splits abrupt.splits )
+        by (rule sxalloc_type_sound [elim_format])
+           (insert False, auto split: option.splits abrupt.splits )
       with False 
       have "\<not>  G,s2\<turnstile>catch catchC"
-	by (simp add: catch_def)
+        by (simp add: catch_def)
       with Try
       have "s3=s2"
-	by simp
+        by simp
       with wt conf_s1 error_free_s1 eq_s2_s1
       show ?thesis
-	by simp
+        by simp
     next
       case True
       note exception_s1 = this
       show ?thesis
       proof (cases "G,s2\<turnstile>catch catchC") 
-	case False
-	with Try
-	have "s3=s2"
-	  by simp
-	with wt conf_s2 error_free_s2 
-	show ?thesis
-	  by simp
+        case False
+        with Try
+        have "s3=s2"
+          by simp
+        with wt conf_s2 error_free_s2 
+        show ?thesis
+          by simp
       next
-	case True
-	with Try have "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3" by simp
-	from True Try.hyps
-	have "?TypeSafeObj (new_xcpt_var vn s2) s3 (In1r c2) \<diamondsuit>"
-	  by simp
-	note hyp_c2 = this [rule_format]
-	from exception_s1 sx_alloc wf
-	obtain a 
-	  where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
-	  by (auto dest!: sxalloc_type_sound split: option.splits abrupt.splits)
-	with True
-	have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class catchC"
-	  by (cases s2) simp
-	with xcpt_s2 conf_s2 wf
-	have "new_xcpt_var vn s2 \<Colon>\<preceq>(G, L(VName vn\<mapsto>Class catchC))"
-	  by (auto dest: Try_lemma)
-	moreover note wt_c2
-	moreover
-	obtain C2' where
-	  "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class catchC)\<rparr>
+        case True
+        with Try have "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3" by simp
+        from True Try.hyps
+        have "?TypeSafeObj (new_xcpt_var vn s2) s3 (In1r c2) \<diamondsuit>"
+          by simp
+        note hyp_c2 = this [rule_format]
+        from exception_s1 sx_alloc wf
+        obtain a 
+          where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
+          by (auto dest!: sxalloc_type_sound split: option.splits abrupt.splits)
+        with True
+        have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class catchC"
+          by (cases s2) simp
+        with xcpt_s2 conf_s2 wf
+        have "new_xcpt_var vn s2 \<Colon>\<preceq>(G, L(VName vn\<mapsto>Class catchC))"
+          by (auto dest: Try_lemma)
+        moreover note wt_c2
+        moreover
+        obtain C2' where
+          "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class catchC)\<rparr>
           \<turnstile> (dom (locals (store (new_xcpt_var vn s2)))) \<guillemotright>In1r c2\<guillemotright> C2'"
-	proof -
-	  have "(dom (locals (store ((Norm s0)::state))) \<union> {VName vn}) 
+        proof -
+          have "(dom (locals (store ((Norm s0)::state))) \<union> {VName vn}) 
                   \<subseteq> dom (locals (store (new_xcpt_var vn s2)))"
           proof -
             from `G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1`
@@ -2189,20 +2189,20 @@
             ultimately show ?thesis
               by (rule Un_least)
           qed
-	  with da_c2 show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	ultimately
-	obtain       conf_s3: "s3\<Colon>\<preceq>(G, L(VName vn\<mapsto>Class catchC))" and
+          with da_c2 show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        ultimately
+        obtain       conf_s3: "s3\<Colon>\<preceq>(G, L(VName vn\<mapsto>Class catchC))" and
                error_free_s3: "error_free s3"
-	  by (rule hyp_c2 [elim_format])
+          by (rule hyp_c2 [elim_format])
              (cases s2, simp add: xcpt_s2 error_free_s2) 
-	from conf_s3 fresh_vn 
-	have "s3\<Colon>\<preceq>(G,L)"
-	  by (blast intro: conforms_deallocL)
-	with wt error_free_s3
-	show ?thesis
-	  by simp
+        from conf_s3 fresh_vn 
+        have "s3\<Colon>\<preceq>(G,L)"
+          by (blast intro: conforms_deallocL)
+        with wt error_free_s3
+        show ?thesis
+          by simp
       qed
     qed
   next
@@ -2242,9 +2242,9 @@
         by (rule dom_locals_eval_mono_elim)
       hence "dom (locals (store ((Norm s0)::state))) 
               \<subseteq> dom (locals (store ((Norm s1)::state)))"
-	by simp
+        by simp
       with da_c2 show thesis
-	by (rule da_weakenE) (rule that)
+        by (rule da_weakenE) (rule that)
     qed
     ultimately
     obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
@@ -2261,19 +2261,19 @@
     next
       case (Some x) 
       from eval_c2 have 
-	"dom (locals (store ((Norm s1)::state))) \<subseteq> dom (locals (store s2))"
-	by (rule dom_locals_eval_mono_elim)
+        "dom (locals (store ((Norm s1)::state))) \<subseteq> dom (locals (store s2))"
+        by (rule dom_locals_eval_mono_elim)
       with Some eval_c2 wf conf_s1 conf_s2
       have conf: "(abrupt_if True (Some x) (abrupt s2), store s2)\<Colon>\<preceq>(G, L)"
-	by (cases s2) (auto dest: Fin_lemma)
+        by (cases s2) (auto dest: Fin_lemma)
       from Some error_free_s1
       have "\<not> (\<exists> err. x=Error err)"
-	by (simp add: error_free_def)
+        by (simp add: error_free_def)
       with error_free_s2
       have "error_free (abrupt_if True (Some x) (abrupt s2), store s2)"
-	by (cases s2) simp
+        by (cases s2) simp
       with Some wt conf s3' show ?thesis
-	by (cases s2) auto
+        by (cases s2) auto
     qed
   next
     case (Init C c s0 s3 s1 s2 L accC T A)
@@ -2288,9 +2288,9 @@
     proof (cases "inited C (globs s0)")
       case True
       with Init.hyps have "s3 = Norm s0"
-	by simp
+        by simp
       with conf_s0 wt show ?thesis 
-	by simp
+        by simp
     next
       case False
       with Init.hyps obtain 
@@ -2298,88 +2298,88 @@
            "G\<turnstile>Norm ((init_class_obj G C) s0) 
               \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1" and
         eval_init: "G\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<rightarrow> s2" and
-	s3: "s3 = (set_lvars (locals (store s1))) s2" 
-	by simp
+        s3: "s3 = (set_lvars (locals (store s1))) s2" 
+        by simp
       have "?TypeSafeObj (Norm ((init_class_obj G C) s0)) s1
-	              (In1r (if C = Object then Skip else Init (super c))) \<diamondsuit>"
-	using False Init.hyps by simp
+                      (In1r (if C = Object then Skip else Init (super c))) \<diamondsuit>"
+        using False Init.hyps by simp
       note hyp_init_super = this [rule_format] 
       have "?TypeSafeObj ((set_lvars empty) s1) s2 (In1r (init c)) \<diamondsuit>"
-	using False Init.hyps by simp
+        using False Init.hyps by simp
       note hyp_init_c = this [rule_format]
       from conf_s0 wf cls_C False
       have "(Norm ((init_class_obj G C) s0))\<Colon>\<preceq>(G, L)"
-	by (auto dest: conforms_init_class_obj)
+        by (auto dest: conforms_init_class_obj)
       moreover from wf cls_C have
-	wt_init_super: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>
+        wt_init_super: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>
                          \<turnstile>(if C = Object then Skip else Init (super c))\<Colon>\<surd>"
-	by (cases "C=Object")
+        by (cases "C=Object")
            (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
       moreover
       obtain S where
-	da_init_super:
-	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        da_init_super:
+        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
           \<turnstile> dom (locals (store ((Norm ((init_class_obj G C) s0))::state))) 
                \<guillemotright>In1r (if C = Object then Skip else Init (super c))\<guillemotright> S"
       proof (cases "C=Object")
-	case True 
-	with da_Skip show ?thesis
-	  using that by (auto intro: assigned.select_convs)
+        case True 
+        with da_Skip show ?thesis
+          using that by (auto intro: assigned.select_convs)
       next
-	case False 
-	with da_Init show ?thesis
-	  by - (rule that, auto intro: assigned.select_convs)
+        case False 
+        with da_Init show ?thesis
+          by - (rule that, auto intro: assigned.select_convs)
       qed
       ultimately 
       obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" and error_free_s1: "error_free s1"
-	by (rule hyp_init_super [elim_format]) simp
+        by (rule hyp_init_super [elim_format]) simp
       from eval_init_super wt_init_super wf
       have s1_no_ret: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
-	by - (rule eval_statement_no_jump [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>"],
+        by - (rule eval_statement_no_jump [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>"],
               auto)
       with conf_s1
       have "(set_lvars empty) s1\<Colon>\<preceq>(G, empty)"
-	by (cases s1) (auto intro: conforms_set_locals)
+        by (cases s1) (auto intro: conforms_set_locals)
       moreover 
       from error_free_s1
       have error_free_empty: "error_free ((set_lvars empty) s1)"
-	by simp
+        by simp
       from cls_C wf have wt_init_c: "\<lparr>prg=G, cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd>"
-	by (rule wf_prog_cdecl [THEN wf_cdecl_wt_init])
+        by (rule wf_prog_cdecl [THEN wf_cdecl_wt_init])
       moreover from cls_C wf obtain I
-	where "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>In1r (init c)\<guillemotright> I"
-	by (rule wf_prog_cdecl [THEN wf_cdeclE,simplified]) blast
+        where "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>In1r (init c)\<guillemotright> I"
+        by (rule wf_prog_cdecl [THEN wf_cdeclE,simplified]) blast
        (*  simplified: to rewrite \<langle>init c\<rangle> to In1r (init c) *) 
       then obtain I' where
-	"\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>dom (locals (store ((set_lvars empty) s1))) 
+        "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>dom (locals (store ((set_lvars empty) s1))) 
             \<guillemotright>In1r (init c)\<guillemotright> I'"
-	  by (rule da_weakenE) simp
+          by (rule da_weakenE) simp
       ultimately
       obtain conf_s2: "s2\<Colon>\<preceq>(G, empty)" and error_free_s2: "error_free s2"
-	by (rule hyp_init_c [elim_format]) (simp add: error_free_empty)
+        by (rule hyp_init_c [elim_format]) (simp add: error_free_empty)
       have "abrupt s2 \<noteq> Some (Jump Ret)"
       proof -
-	from s1_no_ret 
-	have "\<And> j. abrupt ((set_lvars empty) s1) \<noteq> Some (Jump j)"
-	  by simp
-	moreover
-	from cls_C wf have "jumpNestingOkS {} (init c)"
-	  by (rule wf_prog_cdecl [THEN wf_cdeclE])
-	ultimately 
-	show ?thesis
-	  using eval_init wt_init_c wf
-	  by - (rule eval_statement_no_jump 
+        from s1_no_ret 
+        have "\<And> j. abrupt ((set_lvars empty) s1) \<noteq> Some (Jump j)"
+          by simp
+        moreover
+        from cls_C wf have "jumpNestingOkS {} (init c)"
+          by (rule wf_prog_cdecl [THEN wf_cdeclE])
+        ultimately 
+        show ?thesis
+          using eval_init wt_init_c wf
+          by - (rule eval_statement_no_jump 
                      [where ?Env="\<lparr>prg=G,cls=C,lcl=empty\<rparr>"],simp+)
       qed
       with conf_s2 s3 conf_s1 eval_init
       have "s3\<Colon>\<preceq>(G, L)"
-	by (cases s2,cases s1) (force dest: conforms_return eval_gext')
+        by (cases s2,cases s1) (force dest: conforms_return eval_gext')
       moreover from error_free_s2 s3
       have "error_free s3"
-	by simp
+        by simp
       moreover note wt
       ultimately show ?thesis
-	by simp
+        by simp
     qed
   next
     case (NewC s0 C s1 a s2 L accC T A)
@@ -2436,39 +2436,39 @@
     proof -
       note conf_s0 wt_init
       moreover obtain I where
-	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
          \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1r (init_comp_ty elT)\<guillemotright> I"
       proof (cases "\<exists>C. elT = Class C")
-	case True
-	thus ?thesis
-	  by - (rule that, (auto intro: da_Init [simplified] 
+        case True
+        thus ?thesis
+          by - (rule that, (auto intro: da_Init [simplified] 
                                         assigned.select_convs
                               simp add: init_comp_ty_def))
-	 (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
+         (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
       next
-	case False
-	thus ?thesis
-	by - (rule that, (auto intro: da_Skip [simplified] 
+        case False
+        thus ?thesis
+        by - (rule that, (auto intro: da_Skip [simplified] 
                                       assigned.select_convs
                            simp add: init_comp_ty_def))
          (* simplified: to rewrite \<langle>Skip\<rangle> to In1r (Skip) *)
       qed
       ultimately show thesis
-	by (rule hyp_init [elim_format]) (auto intro: that)
+        by (rule hyp_init [elim_format]) (auto intro: that)
     qed 
     obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
     proof -
       from eval_init 
       have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
-	by (rule dom_locals_eval_mono_elim)
+        by (rule dom_locals_eval_mono_elim)
       with da_e 
       obtain A' where
        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
             \<turnstile> dom (locals (store s1)) \<guillemotright>In1l e\<guillemotright> A'"
-	by (rule da_weakenE)
+        by (rule da_weakenE)
       with conf_s1 wt_size
       show ?thesis
-	by (rule hyp_size [elim_format]) (simp add: that error_free_s1) 
+        by (rule hyp_size [elim_format]) (simp add: that error_free_s1) 
     qed
     from conf_s2 have "abupd (check_neg i) s2\<Colon>\<preceq>(G, L)"
       by (cases s2) auto
@@ -2515,14 +2515,14 @@
       assume norm_s2: "normal s2"
       have "G,L,store s2\<turnstile>In1l (Cast castT e)\<succ>In1 v\<Colon>\<preceq>T"
       proof -
-	from s2 norm_s2 have "normal s1"
-	  by (cases s1) simp
-	with v_ok 
-	have "G,store s1\<turnstile>v\<Colon>\<preceq>eT"
-	  by simp
-	with eT wf s2 T norm_s2
-	show ?thesis
-	  by (cases s1) (auto dest: fits_conf)
+        from s2 norm_s2 have "normal s1"
+          by (cases s1) simp
+        with v_ok 
+        have "G,store s1\<turnstile>v\<Colon>\<preceq>eT"
+          by simp
+        with eT wf s2 T norm_s2
+        show ?thesis
+          by (cases s1) (auto dest: fits_conf)
       qed
     }
     with conf_s2 error_free_s2
@@ -2644,32 +2644,32 @@
       case False
       with eval_e2 have "s2=s1" by auto
       with conf_s1 error_free_s1 False show ?thesis
-	by auto
+        by auto
     next
       case True
       note normal_s1 = this
       show ?thesis 
       proof (cases "need_second_arg binop v1")
-	case False
-	with normal_s1 eval_e2 have "s2=s1"
-	  by (cases s1) (simp, elim eval_elim_cases,simp)
-	with conf_s1 conf_v error_free_s1
-	show ?thesis by simp
+        case False
+        with normal_s1 eval_e2 have "s2=s1"
+          by (cases s1) (simp, elim eval_elim_cases,simp)
+        with conf_s1 conf_v error_free_s1
+        show ?thesis by simp
       next
-	case True
-	note need_second_arg = this
-	with hyp_e2 
-	have hyp_e2': "PROP ?TypeSafe s1 s2 (In1l e2) (In1 v2)" by simp
-	from da wt_e1 wt_e2 wt_binop conf_s0 normal_s1 eval_e1 
+        case True
+        note need_second_arg = this
+        with hyp_e2 
+        have hyp_e2': "PROP ?TypeSafe s1 s2 (In1l e2) (In1 v2)" by simp
+        from da wt_e1 wt_e2 wt_binop conf_s0 normal_s1 eval_e1 
           wt_v1 [rule_format,OF normal_s1] wf
-	obtain E2 where
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In1l e2\<guillemotright> E2"
-	  by (rule da_e2_BinOp [elim_format]) 
+        obtain E2 where
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In1l e2\<guillemotright> E2"
+          by (rule da_e2_BinOp [elim_format]) 
              (auto simp add: need_second_arg )
-	with conf_s1 wt_e2 
-	obtain "s2\<Colon>\<preceq>(G, L)" and "error_free s2"
-	  by (rule hyp_e2' [elim_format]) (simp add: error_free_s1)
-	with conf_v show ?thesis by simp
+        with conf_s1 wt_e2 
+        obtain "s2\<Colon>\<preceq>(G, L)" and "error_free s2"
+          by (rule hyp_e2' [elim_format]) (simp add: error_free_s1)
+        with conf_v show ?thesis by simp
       qed
     qed
   next
@@ -2713,18 +2713,18 @@
       fix n assume lvar: "v=LVar n"
       have "locals (store s1) n \<noteq> None"
       proof -
-	from Acc.prems lvar have 
-	  "n \<in> dom (locals s0)"
-	  by (cases "\<exists> n. v=LVar n") (auto elim!: da_elim_cases)
-	also
-	have "dom (locals s0) \<subseteq> dom (locals (store s1))"
-	proof -
-	  from `G\<turnstile>Norm s0 \<midarrow>v=\<succ>(w, upd)\<rightarrow> s1`
-	  show ?thesis
-	    by (rule dom_locals_eval_mono_elim) simp
-	qed
-	finally show ?thesis
-	  by blast
+        from Acc.prems lvar have 
+          "n \<in> dom (locals s0)"
+          by (cases "\<exists> n. v=LVar n") (auto elim!: da_elim_cases)
+        also
+        have "dom (locals s0) \<subseteq> dom (locals (store s1))"
+        proof -
+          from `G\<turnstile>Norm s0 \<midarrow>v=\<succ>(w, upd)\<rightarrow> s1`
+          show ?thesis
+            by (rule dom_locals_eval_mono_elim) simp
+        qed
+        finally show ?thesis
+          by blast
       qed
     } note lvar_in_locals = this 
     from conf_s0 wt_v da_v
@@ -2746,9 +2746,9 @@
     note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
     note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (var:=e)\<Colon>T`
     then obtain varT eT where
-	 wt_var: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>var\<Colon>=varT" and
-	   wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" and
-	  widen: "G\<turnstile>eT\<preceq>varT" and
+         wt_var: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>var\<Colon>=varT" and
+           wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" and
+          widen: "G\<turnstile>eT\<preceq>varT" and
               T: "T=Inl eT"
       by (rule wt_elim_cases) auto
     show "assign upd v s2\<Colon>\<preceq>(G, L) \<and>
@@ -2759,139 +2759,139 @@
       case False
       with Ass.prems
       obtain V E where
-	da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                    \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In2 var\<guillemotright> V" and
-	da_e:   "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>In1l e\<guillemotright> E"
-	by (elim da_elim_cases) simp+
+        da_e:   "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>In1l e\<guillemotright> E"
+        by (elim da_elim_cases) simp+
       from conf_s0 wt_var da_var 
       obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" 
-	and  conf_var: "normal s1 
+        and  conf_var: "normal s1 
                          \<longrightarrow> G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
-	and  error_free_s1: "error_free s1"
-	by (rule hyp_var [elim_format]) simp
+        and  error_free_s1: "error_free s1"
+        by (rule hyp_var [elim_format]) simp
       show ?thesis
       proof (cases "normal s1")
-	case False
-	with eval_e have "s2=s1" by auto
-	with False have "assign upd v s2=s1"
-	  by simp
-	with conf_s1 error_free_s1 False show ?thesis
-	  by auto
+        case False
+        with eval_e have "s2=s1" by auto
+        with False have "assign upd v s2=s1"
+          by simp
+        with conf_s1 error_free_s1 False show ?thesis
+          by auto
       next
-	case True
-	note normal_s1=this
-	obtain A' where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        case True
+        note normal_s1=this
+        obtain A' where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                          \<turnstile> dom (locals (store s1)) \<guillemotright>In1l e\<guillemotright> A'"
-	proof -
-	  from eval_var wt_var da_var wf normal_s1
-	  have "nrm V \<subseteq>  dom (locals (store s1))"
-	    by (cases rule: da_good_approxE') iprover
-	  with da_e show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	with conf_s1 wt_e 
-	obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and 
+        proof -
+          from eval_var wt_var da_var wf normal_s1
+          have "nrm V \<subseteq>  dom (locals (store s1))"
+            by (cases rule: da_good_approxE') iprover
+          with da_e show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        with conf_s1 wt_e 
+        obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and 
           conf_v: "normal s2 \<longrightarrow> G,store s2\<turnstile>v\<Colon>\<preceq>eT" and
           error_free_s2: "error_free s2"
-	  by (rule hyp_e [elim_format]) (simp add: error_free_s1)
-	show ?thesis 
-	proof (cases "normal s2")
-	  case False
-	  with conf_s2 error_free_s2 
-	  show ?thesis
-	    by auto
-	next
-	  case True
-	  from True conf_v
-	  have conf_v_eT: "G,store s2\<turnstile>v\<Colon>\<preceq>eT"
-	    by simp
-	  with widen wf
-	  have conf_v_varT: "G,store s2\<turnstile>v\<Colon>\<preceq>varT"
-	    by (auto intro: conf_widen)
-	  from normal_s1 conf_var
-	  have "G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
-	    by simp
-	  then 
-	  have conf_assign:  "store s1\<le>|upd\<preceq>varT\<Colon>\<preceq>(G, L)" 
-	    by (simp add: rconf_def)
-	  from conf_v_eT conf_v_varT conf_assign normal_s1 True wf eval_var 
-	    eval_e T conf_s2 error_free_s2
-	  show ?thesis
-	    by (cases s1, cases s2) 
-	       (auto dest!: Ass_lemma simp add: assign_conforms_def)
-	qed
+          by (rule hyp_e [elim_format]) (simp add: error_free_s1)
+        show ?thesis 
+        proof (cases "normal s2")
+          case False
+          with conf_s2 error_free_s2 
+          show ?thesis
+            by auto
+        next
+          case True
+          from True conf_v
+          have conf_v_eT: "G,store s2\<turnstile>v\<Colon>\<preceq>eT"
+            by simp
+          with widen wf
+          have conf_v_varT: "G,store s2\<turnstile>v\<Colon>\<preceq>varT"
+            by (auto intro: conf_widen)
+          from normal_s1 conf_var
+          have "G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
+            by simp
+          then 
+          have conf_assign:  "store s1\<le>|upd\<preceq>varT\<Colon>\<preceq>(G, L)" 
+            by (simp add: rconf_def)
+          from conf_v_eT conf_v_varT conf_assign normal_s1 True wf eval_var 
+            eval_e T conf_s2 error_free_s2
+          show ?thesis
+            by (cases s1, cases s2) 
+               (auto dest!: Ass_lemma simp add: assign_conforms_def)
+        qed
       qed
     next
       case True
       then obtain vn where vn: "var=LVar vn"
-	by blast
+        by blast
       with Ass.prems
       obtain E where
-	da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> 
-	           \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e\<guillemotright> E"
-	by (elim da_elim_cases) simp+
+        da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> 
+                   \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e\<guillemotright> E"
+        by (elim da_elim_cases) simp+
       from da.LVar vn obtain V where
-	da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                    \<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In2 var\<guillemotright> V"
-	by auto
+        by auto
       obtain E' where
-	da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                    \<turnstile> dom (locals (store s1)) \<guillemotright>In1l e\<guillemotright> E'"
       proof -
-	have "dom (locals (store ((Norm s0)::state))) 
+        have "dom (locals (store ((Norm s0)::state))) 
                 \<subseteq> dom (locals (store (s1)))"
-	  by (rule dom_locals_eval_mono_elim) (rule Ass.hyps)
-	with da_e show thesis
-	  by (rule da_weakenE) (rule that)
+          by (rule dom_locals_eval_mono_elim) (rule Ass.hyps)
+        with da_e show thesis
+          by (rule da_weakenE) (rule that)
       qed
       from conf_s0 wt_var da_var 
       obtain conf_s1: "s1\<Colon>\<preceq>(G, L)" 
-	and  conf_var: "normal s1 
+        and  conf_var: "normal s1 
                          \<longrightarrow> G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
-	and  error_free_s1: "error_free s1"
-	by (rule hyp_var [elim_format]) simp
+        and  error_free_s1: "error_free s1"
+        by (rule hyp_var [elim_format]) simp
       show ?thesis
       proof (cases "normal s1")
-	case False
-	with eval_e have "s2=s1" by auto
-	with False have "assign upd v s2=s1"
-	  by simp
-	with conf_s1 error_free_s1 False show ?thesis
-	  by auto
+        case False
+        with eval_e have "s2=s1" by auto
+        with False have "assign upd v s2=s1"
+          by simp
+        with conf_s1 error_free_s1 False show ?thesis
+          by auto
       next
-	case True
-	note normal_s1 = this
-	from conf_s1 wt_e da_e'
-	obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and 
+        case True
+        note normal_s1 = this
+        from conf_s1 wt_e da_e'
+        obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and 
           conf_v: "normal s2 \<longrightarrow> G,store s2\<turnstile>v\<Colon>\<preceq>eT" and
           error_free_s2: "error_free s2"
-	  by (rule hyp_e [elim_format]) (simp add: error_free_s1)
-	show ?thesis 
-	proof (cases "normal s2")
-	  case False
-	  with conf_s2 error_free_s2 
-	  show ?thesis
-	    by auto
-	next
-	  case True
-	  from True conf_v
-	  have conf_v_eT: "G,store s2\<turnstile>v\<Colon>\<preceq>eT"
-	    by simp
-	  with widen wf
-	  have conf_v_varT: "G,store s2\<turnstile>v\<Colon>\<preceq>varT"
-	    by (auto intro: conf_widen)
-	  from normal_s1 conf_var
-	  have "G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
-	    by simp
-	  then 
-	  have conf_assign:  "store s1\<le>|upd\<preceq>varT\<Colon>\<preceq>(G, L)" 
-	    by (simp add: rconf_def)
-	  from conf_v_eT conf_v_varT conf_assign normal_s1 True wf eval_var 
-	    eval_e T conf_s2 error_free_s2
-	  show ?thesis
-	    by (cases s1, cases s2) 
-	       (auto dest!: Ass_lemma simp add: assign_conforms_def)
-	qed
+          by (rule hyp_e [elim_format]) (simp add: error_free_s1)
+        show ?thesis 
+        proof (cases "normal s2")
+          case False
+          with conf_s2 error_free_s2 
+          show ?thesis
+            by auto
+        next
+          case True
+          from True conf_v
+          have conf_v_eT: "G,store s2\<turnstile>v\<Colon>\<preceq>eT"
+            by simp
+          with widen wf
+          have conf_v_varT: "G,store s2\<turnstile>v\<Colon>\<preceq>varT"
+            by (auto intro: conf_widen)
+          from normal_s1 conf_var
+          have "G,L,store s1\<turnstile>In2 var\<succ>In2 (w, upd)\<Colon>\<preceq>Inl varT"
+            by simp
+          then 
+          have conf_assign:  "store s1\<le>|upd\<preceq>varT\<Colon>\<preceq>(G, L)" 
+            by (simp add: rconf_def)
+          from conf_v_eT conf_v_varT conf_assign normal_s1 True wf eval_var 
+            eval_e T conf_s2 error_free_s2
+          show ?thesis
+            by (cases s1, cases s2) 
+               (auto dest!: Ass_lemma simp add: assign_conforms_def)
+        qed
       qed
     qed
   next
@@ -2931,60 +2931,60 @@
       case False
       with eval_e1_e2 have "s2=s1" by auto
       with conf_s1 error_free_s1 False show ?thesis
-	by auto
+        by auto
     next
       case True
       have s0_s1: "dom (locals (store ((Norm s0)::state))) 
                     \<union> assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
       proof -
-	from eval_e0 have 
-	  "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
-	  by (rule dom_locals_eval_mono_elim)
+        from eval_e0 have 
+          "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+          by (rule dom_locals_eval_mono_elim)
         moreover
-	from eval_e0 True wt_e0 
-	have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
-	  by (rule assigns_if_good_approx') 
-	ultimately show ?thesis by (rule Un_least)
+        from eval_e0 True wt_e0 
+        have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
+          by (rule assigns_if_good_approx') 
+        ultimately show ?thesis by (rule Un_least)
       qed 
       show ?thesis
       proof (cases "the_Bool b")
-	case True
-	with hyp_if have hyp_e1: "PROP ?TypeSafe s1 s2 (In1l e1) (In1 v)" 
-	  by simp
-	from da_e1 s0_s1 True obtain E1' where
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>In1l e1\<guillemotright> E1'"
-	  by - (rule da_weakenE, auto iff del: Un_subset_iff le_sup_iff)
-	with conf_s1 wt_e1
-	obtain 
-	  "s2\<Colon>\<preceq>(G, L)"
-	  "(normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1l e1\<succ>In1 v\<Colon>\<preceq>Inl T1)"
-	  "error_free s2"            
-	  by (rule hyp_e1 [elim_format]) (simp add: error_free_s1)
-	moreover
-	from statT  
-	have "G\<turnstile>T1\<preceq>statT"
-	  by auto
-	ultimately show ?thesis
-	  using T wf by auto
+        case True
+        with hyp_if have hyp_e1: "PROP ?TypeSafe s1 s2 (In1l e1) (In1 v)" 
+          by simp
+        from da_e1 s0_s1 True obtain E1' where
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>In1l e1\<guillemotright> E1'"
+          by - (rule da_weakenE, auto iff del: Un_subset_iff le_sup_iff)
+        with conf_s1 wt_e1
+        obtain 
+          "s2\<Colon>\<preceq>(G, L)"
+          "(normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1l e1\<succ>In1 v\<Colon>\<preceq>Inl T1)"
+          "error_free s2"            
+          by (rule hyp_e1 [elim_format]) (simp add: error_free_s1)
+        moreover
+        from statT  
+        have "G\<turnstile>T1\<preceq>statT"
+          by auto
+        ultimately show ?thesis
+          using T wf by auto
       next
-	case False
-	with hyp_if have hyp_e2: "PROP ?TypeSafe s1 s2 (In1l e2) (In1 v)" 
-	  by simp
-	from da_e2 s0_s1 False obtain E2' where
-	  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>In1l e2\<guillemotright> E2'"
-	  by - (rule da_weakenE, auto iff del: Un_subset_iff le_sup_iff)
-	with conf_s1 wt_e2
-	obtain 
-	  "s2\<Colon>\<preceq>(G, L)"
-	  "(normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1l e2\<succ>In1 v\<Colon>\<preceq>Inl T2)"
-	  "error_free s2"            
-	  by (rule hyp_e2 [elim_format]) (simp add: error_free_s1)
-	moreover
-	from statT  
-	have "G\<turnstile>T2\<preceq>statT"
-	  by auto
-	ultimately show ?thesis
-	  using T wf by auto
+        case False
+        with hyp_if have hyp_e2: "PROP ?TypeSafe s1 s2 (In1l e2) (In1 v)" 
+          by simp
+        from da_e2 s0_s1 False obtain E2' where
+          "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>In1l e2\<guillemotright> E2'"
+          by - (rule da_weakenE, auto iff del: Un_subset_iff le_sup_iff)
+        with conf_s1 wt_e2
+        obtain 
+          "s2\<Colon>\<preceq>(G, L)"
+          "(normal s2 \<longrightarrow> G,L,store s2\<turnstile>In1l e2\<succ>In1 v\<Colon>\<preceq>Inl T2)"
+          "error_free s2"            
+          by (rule hyp_e2 [elim_format]) (simp add: error_free_s1)
+        moreover
+        from statT  
+        have "G\<turnstile>T2\<preceq>statT"
+          by auto
+        ultimately show ?thesis
+          using T wf by auto
       qed
     qed
   next
@@ -3031,22 +3031,22 @@
       assume abnormal_s2: "\<not> normal s2"
       have "set_lvars (locals (store s2)) s4 = s2"
       proof -
-	from abnormal_s2 init_lvars 
-	obtain keep_abrupt: "abrupt s3 = abrupt s2" and
+        from abnormal_s2 init_lvars 
+        obtain keep_abrupt: "abrupt s3 = abrupt s2" and
              "store s3 = store (init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> 
                                             mode a vs s2)" 
-	  by (auto simp add: init_lvars_def2)
-	moreover
-	from keep_abrupt abnormal_s2 check
-	have eq_s3'_s3: "s3'=s3" 
-	  by (auto simp add: check_method_access_def Let_def)
-	moreover
-	from eq_s3'_s3 abnormal_s2 keep_abrupt eval_methd
-	have "s4=s3'"
-	  by auto
-	ultimately show
-	  "set_lvars (locals (store s2)) s4 = s2"
-	  by (cases s2,cases s3) (simp add: init_lvars_def2)
+          by (auto simp add: init_lvars_def2)
+        moreover
+        from keep_abrupt abnormal_s2 check
+        have eq_s3'_s3: "s3'=s3" 
+          by (auto simp add: check_method_access_def Let_def)
+        moreover
+        from eq_s3'_s3 abnormal_s2 keep_abrupt eval_methd
+        have "s4=s3'"
+          by auto
+        ultimately show
+          "set_lvars (locals (store s2)) s4 = s2"
+          by (cases s2,cases s3) (simp add: init_lvars_def2)
       qed
     } note propagate_abnormal_s2 = this
     show "(set_lvars (locals (store s2))) s4\<Colon>\<preceq>(G, L) \<and>
@@ -3060,113 +3060,113 @@
       with eval_args have "s2=s1" by auto
       with False propagate_abnormal_s2 conf_s1 error_free_s1 
       show ?thesis
-	by auto
+        by auto
     next
       case True
       note normal_s1 = this
       obtain A' where 
-	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In3 args\<guillemotright> A'"
+        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In3 args\<guillemotright> A'"
       proof -
-	from eval_e wt_e da_e wf normal_s1
-	have "nrm E \<subseteq>  dom (locals (store s1))"
-	  by (cases rule: da_good_approxE') iprover
-	with da_args show thesis
-	  by (rule da_weakenE) (rule that)
+        from eval_e wt_e da_e wf normal_s1
+        have "nrm E \<subseteq>  dom (locals (store s1))"
+          by (cases rule: da_good_approxE') iprover
+        with da_args show thesis
+          by (rule da_weakenE) (rule that)
       qed
       with conf_s1 wt_args 
       obtain    conf_s2: "s2\<Colon>\<preceq>(G, L)" and
               conf_args: "normal s2 
                          \<Longrightarrow>  list_all2 (conf G (store s2)) vs pTs" and
           error_free_s2: "error_free s2" 
-	by (rule hyp_args [elim_format]) (simp add: error_free_s1)
+        by (rule hyp_args [elim_format]) (simp add: error_free_s1)
       from error_free_s2 init_lvars
       have error_free_s3: "error_free s3"
-	by (auto simp add: init_lvars_def2)
+        by (auto simp add: init_lvars_def2)
       from statM 
       obtain
-	statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>" and
-	pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
-	by (blast dest: max_spec2mheads)
+        statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>" and
+        pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
+        by (blast dest: max_spec2mheads)
       from check
       have eq_store_s3'_s3: "store s3'=store s3"
-	by (cases s3) (simp add: check_method_access_def Let_def)
+        by (cases s3) (simp add: check_method_access_def Let_def)
       obtain invC
-	where invC: "invC = invocation_class mode (store s2) a statT"
-	by simp
+        where invC: "invC = invocation_class mode (store s2) a statT"
+        by simp
       with init_lvars
       have invC': "invC = (invocation_class mode (store s3) a statT)"
-	by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
+        by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
       show ?thesis
       proof (cases "normal s2")
-	case False
-	with propagate_abnormal_s2 conf_s2 error_free_s2
-	show ?thesis
-	  by auto
+        case False
+        with propagate_abnormal_s2 conf_s2 error_free_s2
+        show ?thesis
+          by auto
       next
-	case True
-	note normal_s2 = True
-	with normal_s1 conf_a eval_args 
-	have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
-	  by (auto dest: eval_gext intro: conf_gext)
-	show ?thesis
-	proof (cases "a=Null \<longrightarrow> is_static statM")
-	  case False
-	  then obtain not_static: "\<not> is_static statM" and Null: "a=Null" 
-	    by blast
-	  with normal_s2 init_lvars mode
-	  obtain np: "abrupt s3 = Some (Xcpt (Std NullPointer))" and
+        case True
+        note normal_s2 = True
+        with normal_s1 conf_a eval_args 
+        have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
+          by (auto dest: eval_gext intro: conf_gext)
+        show ?thesis
+        proof (cases "a=Null \<longrightarrow> is_static statM")
+          case False
+          then obtain not_static: "\<not> is_static statM" and Null: "a=Null" 
+            by blast
+          with normal_s2 init_lvars mode
+          obtain np: "abrupt s3 = Some (Xcpt (Std NullPointer))" and
                      "store s3 = store (init_lvars G invDeclC 
                                        \<lparr>name = mn, parTs = pTs'\<rparr> mode a vs s2)"
-	    by (auto simp add: init_lvars_def2)
-	  moreover
-	  from np check
-	  have eq_s3'_s3: "s3'=s3" 
-	    by (auto simp add: check_method_access_def Let_def)
-	  moreover
-	  from eq_s3'_s3 np eval_methd
-	  have "s4=s3'"
-	    by auto
-	  ultimately have
-	    "set_lvars (locals (store s2)) s4 
+            by (auto simp add: init_lvars_def2)
+          moreover
+          from np check
+          have eq_s3'_s3: "s3'=s3" 
+            by (auto simp add: check_method_access_def Let_def)
+          moreover
+          from eq_s3'_s3 np eval_methd
+          have "s4=s3'"
+            by auto
+          ultimately have
+            "set_lvars (locals (store s2)) s4 
             = (Some (Xcpt (Std NullPointer)),store s2)"
-	    by (cases s2,cases s3) (simp add: init_lvars_def2)
-	  with conf_s2 error_free_s2
-	  show ?thesis
-	    by (cases s2) (auto dest: conforms_NormI)
-	next
-	  case True
-	  with mode have notNull: "mode = IntVir \<longrightarrow> a \<noteq> Null"
-	    by (auto dest!: Null_staticD)
-	  with conf_s2 conf_a_s2 wf invC  
-	  have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
-	    by (cases s2) (auto intro: DynT_propI)
-	  with wt_e statM' invC mode wf 
-	  obtain dynM where 
+            by (cases s2,cases s3) (simp add: init_lvars_def2)
+          with conf_s2 error_free_s2
+          show ?thesis
+            by (cases s2) (auto dest: conforms_NormI)
+        next
+          case True
+          with mode have notNull: "mode = IntVir \<longrightarrow> a \<noteq> Null"
+            by (auto dest!: Null_staticD)
+          with conf_s2 conf_a_s2 wf invC  
+          have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
+            by (cases s2) (auto intro: DynT_propI)
+          with wt_e statM' invC mode wf 
+          obtain dynM where 
             dynM: "dynlookup G statT invC  \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
             acc_dynM: "G \<turnstile>Methd  \<lparr>name=mn,parTs=pTs'\<rparr> dynM 
                             in invC dyn_accessible_from accC"
-	    by (force dest!: call_access_ok)
-	  with invC' check eq_accC_accC'
-	  have eq_s3'_s3: "s3'=s3"
-	    by (auto simp add: check_method_access_def Let_def)
-	  from dynT_prop wf wt_e statM' mode invC invDeclC dynM 
-	  obtain 
-	    wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
-	      dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+            by (force dest!: call_access_ok)
+          with invC' check eq_accC_accC'
+          have eq_s3'_s3: "s3'=s3"
+            by (auto simp add: check_method_access_def Let_def)
+          from dynT_prop wf wt_e statM' mode invC invDeclC dynM 
+          obtain 
+            wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
+              dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
             iscls_invDeclC: "is_class G invDeclC" and
-	         invDeclC': "invDeclC = declclass dynM" and
-	      invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
-	     resTy_widen: "G\<turnstile>resTy dynM\<preceq>resTy statM" and
-	    is_static_eq: "is_static dynM = is_static statM" and
-	    involved_classes_prop:
+                 invDeclC': "invDeclC = declclass dynM" and
+              invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
+             resTy_widen: "G\<turnstile>resTy dynM\<preceq>resTy statM" and
+            is_static_eq: "is_static dynM = is_static statM" and
+            involved_classes_prop:
              "(if invmode statM e = IntVir
                then \<forall>statC. statT = ClassT statC \<longrightarrow> G\<turnstile>invC\<preceq>\<^sub>C statC
                else ((\<exists>statC. statT = ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C invDeclC) \<or>
                      (\<forall>statC. statT \<noteq> ClassT statC \<and> invDeclC = Object)) \<and>
                       statDeclT = ClassT invDeclC)"
-	    by (cases rule: DynT_mheadsE) simp
-	  obtain L' where 
-	   L':"L'=(\<lambda> k. 
+            by (cases rule: DynT_mheadsE) simp
+          obtain L' where 
+           L':"L'=(\<lambda> k. 
                  (case k of
                     EName e
                     \<Rightarrow> (case e of 
@@ -3176,158 +3176,158 @@
                         | Res \<Rightarrow> Some (resTy dynM))
                   | This \<Rightarrow> if is_static statM 
                             then None else Some (Class invDeclC)))"
-	    by simp
-	  from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
+            by simp
+          from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
             wf eval_args conf_a mode notNull wf_dynM involved_classes_prop
-	  have conf_s3: "s3\<Colon>\<preceq>(G,L')"
-	    apply - 
+          have conf_s3: "s3\<Colon>\<preceq>(G,L')"
+            apply - 
                (* FIXME confomrs_init_lvars should be 
                   adjusted to be more directy applicable *)
-	    apply (drule conforms_init_lvars [of G invDeclC 
+            apply (drule conforms_init_lvars [of G invDeclC 
                     "\<lparr>name=mn,parTs=pTs'\<rparr>" dynM "store s2" vs pTs "abrupt s2" 
                     L statT invC a "(statDeclT,statM)" e])
-	    apply (rule wf)
-	    apply (rule conf_args,assumption)
-	    apply (simp add: pTs_widen)
-	    apply (cases s2,simp)
-	    apply (rule dynM')
-	    apply (force dest: ty_expr_is_type)
-	    apply (rule invC_widen)
-	    apply (force intro: conf_gext dest: eval_gext)
-	    apply simp
-	    apply simp
-	    apply (simp add: invC)
-	    apply (simp add: invDeclC)
-	    apply (simp add: normal_s2)
-	    apply (cases s2, simp add: L' init_lvars
-	                     cong add: lname.case_cong ename.case_cong)
-	    done
-	  with eq_s3'_s3 
-	  have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
-	  moreover
-	  from  is_static_eq wf_dynM L'
-	  obtain mthdT where
-	    "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+            apply (rule wf)
+            apply (rule conf_args,assumption)
+            apply (simp add: pTs_widen)
+            apply (cases s2,simp)
+            apply (rule dynM')
+            apply (force dest: ty_expr_is_type)
+            apply (rule invC_widen)
+            apply (force intro: conf_gext dest: eval_gext)
+            apply simp
+            apply simp
+            apply (simp add: invC)
+            apply (simp add: invDeclC)
+            apply (simp add: normal_s2)
+            apply (cases s2, simp add: L' init_lvars
+                             cong add: lname.case_cong ename.case_cong)
+            done
+          with eq_s3'_s3 
+          have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
+          moreover
+          from  is_static_eq wf_dynM L'
+          obtain mthdT where
+            "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
                \<turnstile>Body invDeclC (stmt (mbody (mthd dynM)))\<Colon>-mthdT" and
-	    mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
-	    by - (drule wf_mdecl_bodyD,
+            mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
+            by - (drule wf_mdecl_bodyD,
                  auto simp add: callee_lcl_def  
                       cong add: lname.case_cong ename.case_cong)
-	  with dynM' iscls_invDeclC invDeclC'
-	  have
-	    "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+          with dynM' iscls_invDeclC invDeclC'
+          have
+            "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
                \<turnstile>(Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<Colon>-mthdT"
-	    by (auto intro: wt.Methd)
-	  moreover
-	  obtain M where 
-	    "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr> 
-	       \<turnstile> dom (locals (store s3')) 
-	          \<guillemotright>In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<guillemotright> M"
-	  proof -
-	    from wf_dynM
-	    obtain M' where
-	      da_body: 
-	      "\<lparr>prg=G, cls=invDeclC
+            by (auto intro: wt.Methd)
+          moreover
+          obtain M where 
+            "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr> 
+               \<turnstile> dom (locals (store s3')) 
+                  \<guillemotright>In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<guillemotright> M"
+          proof -
+            from wf_dynM
+            obtain M' where
+              da_body: 
+              "\<lparr>prg=G, cls=invDeclC
                ,lcl=callee_lcl invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> (mthd dynM)
                \<rparr> \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'" and
               res: "Result \<in> nrm M'"
-	      by (rule wf_mdeclE) iprover
-	    from da_body is_static_eq L' have
-	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
+              by (rule wf_mdeclE) iprover
+            from da_body is_static_eq L' have
+              "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
                  \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'"
-	      by (simp add: callee_lcl_def  
+              by (simp add: callee_lcl_def  
                   cong add: lname.case_cong ename.case_cong)
-	    moreover have "parameters (mthd dynM) \<subseteq>  dom (locals (store s3'))"
-	    proof -
-	      from is_static_eq 
-	      have "(invmode (mthd dynM) e) = (invmode statM e)"
-		by (simp add: invmode_def)
-	      moreover
-	      have "length (pars (mthd dynM)) = length vs" 
-	      proof -
-		from normal_s2 conf_args
-		have "length vs = length pTs"
-		  by (simp add: list_all2_def)
-		also from pTs_widen
-		have "\<dots> = length pTs'"
-		  by (simp add: widens_def list_all2_def)
-		also from wf_dynM
-		have "\<dots> = length (pars (mthd dynM))"
-		  by (simp add: wf_mdecl_def wf_mhead_def)
-		finally show ?thesis ..
-	      qed
-	      moreover note init_lvars dynM' is_static_eq normal_s2 mode 
-	      ultimately
-	      have "parameters (mthd dynM) = dom (locals (store s3))"
-		using dom_locals_init_lvars 
+            moreover have "parameters (mthd dynM) \<subseteq>  dom (locals (store s3'))"
+            proof -
+              from is_static_eq 
+              have "(invmode (mthd dynM) e) = (invmode statM e)"
+                by (simp add: invmode_def)
+              moreover
+              have "length (pars (mthd dynM)) = length vs" 
+              proof -
+                from normal_s2 conf_args
+                have "length vs = length pTs"
+                  by (simp add: list_all2_def)
+                also from pTs_widen
+                have "\<dots> = length pTs'"
+                  by (simp add: widens_def list_all2_def)
+                also from wf_dynM
+                have "\<dots> = length (pars (mthd dynM))"
+                  by (simp add: wf_mdecl_def wf_mhead_def)
+                finally show ?thesis ..
+              qed
+              moreover note init_lvars dynM' is_static_eq normal_s2 mode 
+              ultimately
+              have "parameters (mthd dynM) = dom (locals (store s3))"
+                using dom_locals_init_lvars 
                   [of "mthd dynM" G invDeclC "\<lparr>name=mn,parTs=pTs'\<rparr>" vs e a s2]
-		by simp
-	      also from check
-	      have "dom (locals (store s3)) \<subseteq>  dom (locals (store s3'))"
-		by (simp add:  eq_s3'_s3)
-	      finally show ?thesis .
-	    qed
-	    ultimately obtain M2 where
-	      da:
-	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
+                by simp
+              also from check
+              have "dom (locals (store s3)) \<subseteq>  dom (locals (store s3'))"
+                by (simp add:  eq_s3'_s3)
+              finally show ?thesis .
+            qed
+            ultimately obtain M2 where
+              da:
+              "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> 
                 \<turnstile> dom (locals (store s3')) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M2" and
               M2: "nrm M' \<subseteq> nrm M2"
-	      by (rule da_weakenE)
-	    from res M2 have "Result \<in> nrm M2"
-	      by blast
-	    moreover from wf_dynM
-	    have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
-	      by (rule wf_mdeclE)
-	    ultimately
-	    obtain M3 where
-	      "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3')) 
+              by (rule da_weakenE)
+            from res M2 have "Result \<in> nrm M2"
+              by blast
+            moreover from wf_dynM
+            have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
+              by (rule wf_mdeclE)
+            ultimately
+            obtain M3 where
+              "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3')) 
                      \<guillemotright>\<langle>Body (declclass dynM) (stmt (mbody (mthd dynM)))\<rangle>\<guillemotright> M3"
-	      using da
-	      by (iprover intro: da.Body assigned.select_convs)
-	    from _ this [simplified]
-	    show ?thesis
-	      by (rule da.Methd [simplified,elim_format]) (auto intro: dynM' that)
-	  qed
-	  ultimately obtain  
-	    conf_s4: "s4\<Colon>\<preceq>(G, L')" and 
-	    conf_Res: "normal s4 \<longrightarrow> G,store s4\<turnstile>v\<Colon>\<preceq>mthdT" and
-	    error_free_s4: "error_free s4"
-	    by (rule hyp_methd [elim_format]) 
+              using da
+              by (iprover intro: da.Body assigned.select_convs)
+            from _ this [simplified]
+            show ?thesis
+              by (rule da.Methd [simplified,elim_format]) (auto intro: dynM' that)
+          qed
+          ultimately obtain  
+            conf_s4: "s4\<Colon>\<preceq>(G, L')" and 
+            conf_Res: "normal s4 \<longrightarrow> G,store s4\<turnstile>v\<Colon>\<preceq>mthdT" and
+            error_free_s4: "error_free s4"
+            by (rule hyp_methd [elim_format]) 
                (simp add: error_free_s3 eq_s3'_s3)
-	  from init_lvars eval_methd eq_s3'_s3 
-	  have "store s2\<le>|store s4"
-	    by (cases s2) (auto dest!: eval_gext simp add: init_lvars_def2 )
-	  moreover
-	  have "abrupt s4 \<noteq> Some (Jump Ret)"
-	  proof -
-	    from normal_s2 init_lvars
-	    have "abrupt s3 \<noteq> Some (Jump Ret)"
-	      by (cases s2) (simp add: init_lvars_def2 abrupt_if_def)
-	    with check
-	    have "abrupt s3' \<noteq> Some (Jump Ret)"
-	      by (cases s3) (auto simp add: check_method_access_def Let_def)
-	    with eval_methd
-	    show ?thesis
-	      by (rule Methd_no_jump)
-	  qed
-	  ultimately 
-	  have "(set_lvars (locals (store s2))) s4\<Colon>\<preceq>(G, L)"
-	    using conf_s2 conf_s4
-	    by (cases s2,cases s4) (auto intro: conforms_return)
-	  moreover 
-	  from conf_Res mthdT_widen resTy_widen wf
-	  have "normal s4 
+          from init_lvars eval_methd eq_s3'_s3 
+          have "store s2\<le>|store s4"
+            by (cases s2) (auto dest!: eval_gext simp add: init_lvars_def2 )
+          moreover
+          have "abrupt s4 \<noteq> Some (Jump Ret)"
+          proof -
+            from normal_s2 init_lvars
+            have "abrupt s3 \<noteq> Some (Jump Ret)"
+              by (cases s2) (simp add: init_lvars_def2 abrupt_if_def)
+            with check
+            have "abrupt s3' \<noteq> Some (Jump Ret)"
+              by (cases s3) (auto simp add: check_method_access_def Let_def)
+            with eval_methd
+            show ?thesis
+              by (rule Methd_no_jump)
+          qed
+          ultimately 
+          have "(set_lvars (locals (store s2))) s4\<Colon>\<preceq>(G, L)"
+            using conf_s2 conf_s4
+            by (cases s2,cases s4) (auto intro: conforms_return)
+          moreover 
+          from conf_Res mthdT_widen resTy_widen wf
+          have "normal s4 
                   \<longrightarrow> G,store s4\<turnstile>v\<Colon>\<preceq>(resTy statM)"
-	    by (auto dest: widen_trans)
-	  then
-	  have "normal ((set_lvars (locals (store s2))) s4)
+            by (auto dest: widen_trans)
+          then
+          have "normal ((set_lvars (locals (store s2))) s4)
              \<longrightarrow> G,store((set_lvars (locals (store s2))) s4) \<turnstile>v\<Colon>\<preceq>(resTy statM)"
-	    by (cases s4) auto
-	  moreover note error_free_s4 T
-	  ultimately 
-	  show ?thesis
-	    by simp
-	qed
+            by (cases s4) auto
+          moreover note error_free_s4 T
+          ultimately 
+          show ?thesis
+            by simp
+        qed
       qed
     qed
   next
@@ -3393,7 +3393,7 @@
       from eval_init 
       have "(dom (locals (store ((Norm s0)::state)))) 
                      \<subseteq> (dom (locals (store s1)))"
-	by (rule dom_locals_eval_mono_elim)
+        by (rule dom_locals_eval_mono_elim)
       with da_c show thesis by (rule da_weakenE) (rule that)
     qed
     from conf_s1 wt_c da_C' 
@@ -3412,40 +3412,40 @@
     proof -
       from iscls_D
       have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init D)\<Colon>\<surd>"
-	by auto
+        by auto
       from eval_init wf
       have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
-	by - (rule eval_statement_no_jump [OF _ _ _ wt_init],auto)
+        by - (rule eval_statement_no_jump [OF _ _ _ wt_init],auto)
       from eval_c _ wt_c wf
       have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
-	by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
+        by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
       moreover 
       note `s3 =
                 (if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or> 
                         abrupt s2 = Some (Jump (Cont l))
                  then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)`
       ultimately show ?thesis
-	by force
+        by force
     qed
     moreover
     {
       assume normal_upd_s2:  "normal (abupd (absorb Ret) s2)"
       have "Result \<in> dom (locals (store s2))"
       proof -
-	from normal_upd_s2
-	have "normal s2 \<or> abrupt s2 = Some (Jump Ret)"
-	  by (cases s2) (simp add: absorb_def)
-	thus ?thesis
-	proof 
-	  assume "normal s2"
-	  with eval_c wt_c da_C' wf res nrm_C'
-	  show ?thesis
-	    by (cases rule: da_good_approxE') blast
-	next
-	  assume "abrupt s2 = Some (Jump Ret)"
-	  with conf_s2 show ?thesis
-	    by (cases s2) (auto dest: conforms_RetD simp add: dom_def)
-	qed 
+        from normal_upd_s2
+        have "normal s2 \<or> abrupt s2 = Some (Jump Ret)"
+          by (cases s2) (simp add: absorb_def)
+        thus ?thesis
+        proof 
+          assume "normal s2"
+          with eval_c wt_c da_C' wf res nrm_C'
+          show ?thesis
+            by (cases rule: da_good_approxE') blast
+        next
+          assume "abrupt s2 = Some (Jump Ret)"
+          with conf_s2 show ?thesis
+            by (cases s2) (auto dest: conforms_RetD simp add: dom_def)
+        qed 
       qed
     }
     moreover note T resultT
@@ -3494,7 +3494,7 @@
             accfield: "accfield G accC statC fn = Some (statDeclC,f)" and
        eq_accC_accC': "accC=accC'" and
                 stat: "stat=is_static f" and
-	           T: "T=(Inl (type f))"
+                   T: "T=(Inl (type f))"
       by (rule wt_elim_cases) (auto simp add: member_is_static_simp)
     from FVar.prems eq_accC_accC'
     have da_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>
@@ -3522,10 +3522,10 @@
     proof -
       from eval_init
       have "(dom (locals (store ((Norm s0)::state)))) 
-	       \<subseteq> (dom (locals (store s1)))"
-	by (rule dom_locals_eval_mono_elim)
+               \<subseteq> (dom (locals (store s1)))"
+        by (rule dom_locals_eval_mono_elim)
       with da_e show thesis
-	by (rule da_weakenE) (rule that)
+        by (rule da_weakenE) (rule that)
     qed
     with conf_s1 wt_e 
     obtain       conf_s2: "s2\<Colon>\<preceq>(G, L)" and
@@ -3549,41 +3549,41 @@
     proof - (*###FVar_lemma should be adjusted to be more directy applicable *)
       assume normal: "normal s2'"
       obtain vv vf x2 store2 store2'
-	where  v: "v=(vv,vf)" and
+        where  v: "v=(vv,vf)" and
               s2: "s2=(x2,store2)" and
          store2': "store s2' = store2'"
-	by (cases v,cases s2,cases s2') blast
+        by (cases v,cases s2,cases s2') blast
       from iscls_statDeclC obtain c
-	where c: "class G statDeclC = Some c"
-	by auto
+        where c: "class G statDeclC = Some c"
+        by auto
       have "G,store2'\<turnstile>vv\<Colon>\<preceq>type f \<and> store2'\<le>|vf\<preceq>type f\<Colon>\<preceq>(G, L)"
       proof (rule FVar_lemma [of vv vf store2' statDeclC f fn a x2 store2 
                                statC G c L "store s1"])
-	from v normal s2 fvar stat store2' 
-	show "((vv, vf), Norm store2') = 
+        from v normal s2 fvar stat store2' 
+        show "((vv, vf), Norm store2') = 
                fvar statDeclC (static f) fn a (x2, store2)"
-	  by (auto simp add: member_is_static_simp)
-	from accfield iscls_statC wf
-	show "G\<turnstile>statC\<preceq>\<^sub>C statDeclC"
-	  by (auto dest!: accfield_fields dest: fields_declC)
-	from accfield
-	show fld: "table_of (DeclConcepts.fields G statC) (fn, statDeclC) = Some f"
-	  by (auto dest!: accfield_fields)
-	from wf show "wf_prog G" .
-	from conf_a s2 show "x2 = None \<longrightarrow> G,store2\<turnstile>a\<Colon>\<preceq>Class statC"
-	  by auto
-	from fld wf iscls_statC
-	show "statDeclC \<noteq> Object "
-	  by (cases "statDeclC=Object") (drule fields_declC,simp+)+
-	from c show "class G statDeclC = Some c" .
-	from conf_s2 s2 show "(x2, store2)\<Colon>\<preceq>(G, L)" by simp
-	from eval_e s2 show "snd s1\<le>|store2" by (auto dest: eval_gext)
-	from initd_statDeclC_s1 show "inited statDeclC (globs (snd s1))" 
-	  by simp
+          by (auto simp add: member_is_static_simp)
+        from accfield iscls_statC wf
+        show "G\<turnstile>statC\<preceq>\<^sub>C statDeclC"
+          by (auto dest!: accfield_fields dest: fields_declC)
+        from accfield
+        show fld: "table_of (DeclConcepts.fields G statC) (fn, statDeclC) = Some f"
+          by (auto dest!: accfield_fields)
+        from wf show "wf_prog G" .
+        from conf_a s2 show "x2 = None \<longrightarrow> G,store2\<turnstile>a\<Colon>\<preceq>Class statC"
+          by auto
+        from fld wf iscls_statC
+        show "statDeclC \<noteq> Object "
+          by (cases "statDeclC=Object") (drule fields_declC,simp+)+
+        from c show "class G statDeclC = Some c" .
+        from conf_s2 s2 show "(x2, store2)\<Colon>\<preceq>(G, L)" by simp
+        from eval_e s2 show "snd s1\<le>|store2" by (auto dest: eval_gext)
+        from initd_statDeclC_s1 show "inited statDeclC (globs (snd s1))" 
+          by simp
       qed
       with v s2 store2'  
       show ?thesis
-	by simp
+        by simp
     qed
     from fvar error_free_s2
     have "error_free s2'"
@@ -3629,67 +3629,67 @@
       moreover
       from eq_s2_s1 False have  "\<not> normal s2" by simp
       then have "snd (avar G i a s2) = s2"
-	by (cases s2) (simp add: avar_def2)
+        by (cases s2) (simp add: avar_def2)
       with avar have "s2'=s2"
-	by (cases "(avar G i a s2)") simp
+        by (cases "(avar G i a s2)") simp
       ultimately show ?thesis
-	using conf_s1 error_free_s1
-	by auto
+        using conf_s1 error_free_s1
+        by auto
     next
       case True
       obtain A' where 
-	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In1l e2\<guillemotright> A'"
+        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In1l e2\<guillemotright> A'"
       proof -
-	from eval_e1 wt_e1 da_e1 wf True
-	have "nrm E1 \<subseteq> dom (locals (store s1))"
-	  by (cases rule: da_good_approxE') iprover
-	with da_e2 show thesis
-	  by (rule da_weakenE) (rule that)
+        from eval_e1 wt_e1 da_e1 wf True
+        have "nrm E1 \<subseteq> dom (locals (store s1))"
+          by (cases rule: da_good_approxE') iprover
+        with da_e2 show thesis
+          by (rule da_weakenE) (rule that)
       qed
       with conf_s1 wt_e2
       obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and error_free_s2: "error_free s2"
-	by (rule hyp_e2 [elim_format]) (simp add: error_free_s1)
+        by (rule hyp_e2 [elim_format]) (simp add: error_free_s1)
       from avar 
       have "store s2'=store s2"
-	by (cases s2) (simp add: avar_def2)
+        by (cases s2) (simp add: avar_def2)
       with avar conf_s2 
       have conf_s2': "s2'\<Colon>\<preceq>(G, L)"
-	by (cases s2) (auto simp add: avar_def2)
+        by (cases s2) (auto simp add: avar_def2)
       from avar error_free_s2
       have error_free_s2': "error_free s2'"
-	by (cases s2) (auto simp add: avar_def2 )
+        by (cases s2) (auto simp add: avar_def2 )
       have "normal s2' \<Longrightarrow> 
         G,store s2'\<turnstile>fst v\<Colon>\<preceq>elemT \<and> store s2'\<le>|snd v\<preceq>elemT\<Colon>\<preceq>(G, L)"
       proof -(*###AVar_lemma should be adjusted to be more directy applicable *)
-	assume normal: "normal s2'"
-	show ?thesis
-	proof -
-	  obtain vv vf x1 store1 x2 store2 store2'
-	    where  v: "v=(vv,vf)" and
+        assume normal: "normal s2'"
+        show ?thesis
+        proof -
+          obtain vv vf x1 store1 x2 store2 store2'
+            where  v: "v=(vv,vf)" and
                   s1: "s1=(x1,store1)" and
                   s2: "s2=(x2,store2)" and
-	     store2': "store2'=store s2'"
-	    by (cases v,cases s1, cases s2, cases s2') blast 
-	  have "G,store2'\<turnstile>vv\<Colon>\<preceq>elemT \<and> store2'\<le>|vf\<preceq>elemT\<Colon>\<preceq>(G, L)"
-	  proof (rule AVar_lemma [of G x1 store1 e2 i x2 store2 vv vf store2' a,
+             store2': "store2'=store s2'"
+            by (cases v,cases s1, cases s2, cases s2') blast 
+          have "G,store2'\<turnstile>vv\<Colon>\<preceq>elemT \<and> store2'\<le>|vf\<preceq>elemT\<Colon>\<preceq>(G, L)"
+          proof (rule AVar_lemma [of G x1 store1 e2 i x2 store2 vv vf store2' a,
                                   OF wf])
-	    from s1 s2 eval_e2 show "G\<turnstile>(x1, store1) \<midarrow>e2-\<succ>i\<rightarrow> (x2, store2)"
-	      by simp
-	    from v normal s2 store2' avar 
-	    show "((vv, vf), Norm store2') = avar G i a (x2, store2)"
-	      by auto
-	    from s2 conf_s2 show "(x2, store2)\<Colon>\<preceq>(G, L)" by simp
-	    from s1 conf_a show  "x1 = None \<longrightarrow> G,store1\<turnstile>a\<Colon>\<preceq>elemT.[]" by simp 
-	    from eval_e2 s1 s2 show "store1\<le>|store2" by (auto dest: eval_gext)
-	  qed
-	  with v s1 s2 store2' 
-	  show ?thesis
-	    by simp
-	qed
+            from s1 s2 eval_e2 show "G\<turnstile>(x1, store1) \<midarrow>e2-\<succ>i\<rightarrow> (x2, store2)"
+              by simp
+            from v normal s2 store2' avar 
+            show "((vv, vf), Norm store2') = avar G i a (x2, store2)"
+              by auto
+            from s2 conf_s2 show "(x2, store2)\<Colon>\<preceq>(G, L)" by simp
+            from s1 conf_a show  "x1 = None \<longrightarrow> G,store1\<turnstile>a\<Colon>\<preceq>elemT.[]" by simp 
+            from eval_e2 s1 s2 show "store1\<le>|store2" by (auto dest: eval_gext)
+          qed
+          with v s1 s2 store2' 
+          show ?thesis
+            by simp
+        qed
       qed
       with conf_s2' error_free_s2' T 
       show ?thesis 
-	by auto
+        by auto
     qed
   next
     case (Nil s0 L accC T)
@@ -3726,30 +3726,30 @@
       with eval_es have "s2=s1" by auto
       with False conf_s1 error_free_s1
       show ?thesis
-	by auto
+        by auto
     next
       case True
       obtain A' where 
-	"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In3 es\<guillemotright> A'"
+        "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>In3 es\<guillemotright> A'"
       proof -
-	from eval_e wt_e da_e wf True
-	have "nrm E \<subseteq> dom (locals (store s1))"
-	  by (cases rule: da_good_approxE') iprover
-	with da_es show thesis
-	  by (rule da_weakenE) (rule that)
+        from eval_e wt_e da_e wf True
+        have "nrm E \<subseteq> dom (locals (store s1))"
+          by (cases rule: da_good_approxE') iprover
+        with da_es show thesis
+          by (rule da_weakenE) (rule that)
       qed
       with conf_s1 wt_es
       obtain conf_s2: "s2\<Colon>\<preceq>(G, L)" and 
            error_free_s2: "error_free s2" and
            conf_vs: "normal s2 \<longrightarrow> list_all2 (conf G (store s2)) vs esT"
-	by (rule hyp_es [elim_format]) (simp add: error_free_s1)
+        by (rule hyp_es [elim_format]) (simp add: error_free_s1)
       moreover
       from True eval_es conf_v 
       have conf_v': "G,store s2\<turnstile>v\<Colon>\<preceq>eT"
-	apply clarify
-	apply (rule conf_gext)
-	apply (auto dest: eval_gext)
-	done
+        apply clarify
+        apply (rule conf_gext)
+        apply (auto dest: eval_gext)
+        done
       ultimately show ?thesis using T by simp
     qed
   qed
@@ -3956,19 +3956,19 @@
                        P L accC s1 \<langle>c2\<rangle>\<^sub>s \<diamondsuit> s2\<rbrakk> \<Longrightarrow> Q"
       have Q
       proof -
-	obtain C2' where 
-	  da: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
-	proof -
-	  from eval_c1 wt_c1 da_c1 wf normal_s1
-	  have "nrm C1 \<subseteq> dom (locals (store s1))"
-	    by (cases rule: da_good_approxE') iprover
-	  with da_c2 show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	with wt_c2 have "P L accC s1 \<langle>c2\<rangle>\<^sub>s \<diamondsuit> s2"
-	  by (rule Comp.hyps)
-	with da show ?thesis
-	  using elim by iprover
+        obtain C2' where 
+          da: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
+        proof -
+          from eval_c1 wt_c1 da_c1 wf normal_s1
+          have "nrm C1 \<subseteq> dom (locals (store s1))"
+            by (cases rule: da_good_approxE') iprover
+          with da_c2 show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        with wt_c2 have "P L accC s1 \<langle>c2\<rangle>\<^sub>s \<diamondsuit> s2"
+          by (rule Comp.hyps)
+        with da show ?thesis
+          using elim by iprover
       qed
     }
     with eval_c1 eval_c2 wt_c1 wt_c2 da_c1 P_c1 
@@ -4003,28 +4003,28 @@
                               \<rbrakk> \<Longrightarrow> Q"
       have Q
       proof -
-	obtain C' where
-	  da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
+        obtain C' where
+          da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
                 (dom (locals (store s1)))\<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s \<guillemotright> C'"
-	proof -
-	  from eval_e have 
-	    "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
-	    by (rule dom_locals_eval_mono_elim)
+        proof -
+          from eval_e have 
+            "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
+            by (rule dom_locals_eval_mono_elim)
           moreover
-	  from eval_e normal_s1 wt_e 
-	  have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
-	    by (rule assigns_if_good_approx')
-	  ultimately 
-	  have "dom (locals (store ((Norm s0)::state))) 
+          from eval_e normal_s1 wt_e 
+          have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+            by (rule assigns_if_good_approx')
+          ultimately 
+          have "dom (locals (store ((Norm s0)::state))) 
             \<union> assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
-	    by (rule Un_least)
-	  with da_then_else show thesis
-	    by (rule da_weakenE) (rule that)
-	qed
-	with wt_then_else
-	have "P L accC s1 \<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s \<diamondsuit> s2"
-	  by (rule If.hyps)
-	with da show ?thesis using elim by iprover
+            by (rule Un_least)
+          with da_then_else show thesis
+            by (rule da_weakenE) (rule that)
+        qed
+        with wt_then_else
+        have "P L accC s1 \<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s \<diamondsuit> s2"
+          by (rule If.hyps)
+        with da show ?thesis using elim by iprover
       qed
     }
     with eval_e eval_then_else wt_e wt_then_else da_e P_e
--- a/src/HOL/Bali/Value.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/Value.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -11,11 +11,11 @@
 typedecl loc            --{* locations, i.e. abstract references on objects *}
 
 datatype val
-	= Unit          --{* dummy result value of void methods *}
-	| Bool bool     --{* Boolean value *}
-	| Intg int      --{* integer value *}
-	| Null          --{* null reference *}
-	| Addr loc      --{* addresses, i.e. locations of objects *}
+        = Unit          --{* dummy result value of void methods *}
+        | Bool bool     --{* Boolean value *}
+        | Intg int      --{* integer value *}
+        | Null          --{* null reference *}
+        | Addr loc      --{* addresses, i.e. locations of objects *}
 
 
 translations "val" <= (type) "Term.val"
@@ -28,22 +28,22 @@
 consts   the_Addr   :: "val \<Rightarrow> loc"
 primrec "the_Addr (Addr a) = a"
 
-types	dyn_ty      = "loc \<Rightarrow> ty option"
+types   dyn_ty      = "loc \<Rightarrow> ty option"
 consts
   typeof        :: "dyn_ty \<Rightarrow> val \<Rightarrow> ty option"
   defpval       :: "prim_ty \<Rightarrow> val"  --{* default value for primitive types *}
   default_val   :: "     ty \<Rightarrow> val"  --{* default value for all types *}
 
 primrec "typeof dt  Unit    = Some (PrimT Void)"
-	"typeof dt (Bool b) = Some (PrimT Boolean)"
-	"typeof dt (Intg i) = Some (PrimT Integer)"
-	"typeof dt  Null    = Some NT"
-	"typeof dt (Addr a) = dt a"
+        "typeof dt (Bool b) = Some (PrimT Boolean)"
+        "typeof dt (Intg i) = Some (PrimT Integer)"
+        "typeof dt  Null    = Some NT"
+        "typeof dt (Addr a) = dt a"
 
-primrec	"defpval Void    = Unit"
-	"defpval Boolean = Bool False"
-	"defpval Integer = Intg 0"
-primrec	"default_val (PrimT pt) = defpval pt"
-	"default_val (RefT  r ) = Null"
+primrec "defpval Void    = Unit"
+        "defpval Boolean = Bool False"
+        "defpval Integer = Intg 0"
+primrec "default_val (PrimT pt) = defpval pt"
+        "default_val (RefT  r ) = Null"
 
 end
--- a/src/HOL/Bali/WellForm.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/WellForm.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -59,9 +59,9 @@
 constdefs
   wf_mhead :: "prog \<Rightarrow> pname \<Rightarrow> sig \<Rightarrow> mhead \<Rightarrow> bool"
  "wf_mhead G P \<equiv> \<lambda> sig mh. length (parTs sig) = length (pars mh) \<and>
-			    \<spacespace> ( \<forall>T\<in>set (parTs sig). is_acc_type G P T) \<and> 
+                            \<spacespace> ( \<forall>T\<in>set (parTs sig). is_acc_type G P T) \<and> 
                             is_acc_type G P (resTy mh) \<and>
-			    \<spacespace> distinct (pars mh)"
+                            \<spacespace> distinct (pars mh)"
 
 
 text {*
@@ -87,7 +87,7 @@
                   VNam v 
                   \<Rightarrow>(table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
                 | Res \<Rightarrow> Some (resTy m))
-	  | This \<Rightarrow> if is_static m then None else Some (Class C))"
+          | This \<Rightarrow> if is_static m then None else Some (Class C))"
 
 constdefs parameters :: "methd \<Rightarrow> lname set"
 "parameters m \<equiv>  set (map (EName \<circ> VNam) (pars m)) 
@@ -97,10 +97,10 @@
   wf_mdecl :: "prog \<Rightarrow> qtname \<Rightarrow> mdecl \<Rightarrow> bool"
  "wf_mdecl G C \<equiv> 
       \<lambda>(sig,m).
-	  wf_mhead G (pid C) sig (mhead m) \<and> 
+          wf_mhead G (pid C) sig (mhead m) \<and> 
           unique (lcls (mbody m)) \<and> 
           (\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T) \<and> 
-	  (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and>
+          (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and>
           jumpNestingOkS {Ret} (stmt (mbody m)) \<and> 
           is_class G C \<and>
           \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd> \<and>
@@ -225,11 +225,11 @@
  "wf_idecl G \<equiv> 
     \<lambda>(I,i). 
         ws_idecl G I (isuperIfs i) \<and> 
-	\<not>is_class G I \<and>
-	(\<forall>(sig,mh)\<in>set (imethods i). wf_mhead G (pid I) sig mh \<and> 
+        \<not>is_class G I \<and>
+        (\<forall>(sig,mh)\<in>set (imethods i). wf_mhead G (pid I) sig mh \<and> 
                                      \<not>is_static mh \<and>
                                       accmodi mh = Public) \<and>
-	unique (imethods i) \<and>
+        unique (imethods i) \<and>
         (\<forall> J\<in>set (isuperIfs i). is_acc_iface G (pid I) J) \<and>
         (table_of (imethods i)
           hiding (methd G Object)
@@ -238,7 +238,7 @@
                              is_static new = is_static old)) \<and> 
         (Option.set \<circ> table_of (imethods i) 
                hidings Un_tables((\<lambda>J.(imethds G J))`set (isuperIfs i))
-	       entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old))"
+               entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old))"
 
 lemma wf_idecl_mhead: "\<lbrakk>wf_idecl G (I,i); (sig,mh)\<in>set (imethods i)\<rbrakk> \<Longrightarrow>  
   wf_mhead G (pid I) sig mh \<and> \<not>is_static mh \<and> accmodi mh = Public"
@@ -340,8 +340,8 @@
       \<not>is_iface G C \<and>
       (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
         (\<forall>s. \<forall> im \<in> imethds G I s.
-      	    (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
-      	                             \<not> is_static cm \<and>
+            (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
+                                     \<not> is_static cm \<and>
                                      accmodi im \<le> accmodi cm))) \<and>
       (\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and> 
       (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and>
@@ -355,10 +355,10 @@
                        (G,sig\<turnstile>new overrides\<^sub>S old 
                         \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
                              accmodi old \<le> accmodi new \<and>
-      	                     \<not>is_static old)) \<and>
+                             \<not>is_static old)) \<and>
                        (G,sig\<turnstile>new hides old 
                          \<longrightarrow> (accmodi old \<le> accmodi new \<and>
-      	                      is_static old)))) 
+                              is_static old)))) 
             ))"
 
 (*
@@ -369,8 +369,8 @@
       \<not>is_iface G C \<and>
       (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
         (\<forall>s. \<forall> im \<in> imethds G I s.
-      	    (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy (mthd cm)\<preceq>resTy (mthd im) \<and>
-      	                             \<not> is_static cm \<and>
+            (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy (mthd cm)\<preceq>resTy (mthd im) \<and>
+                                     \<not> is_static cm \<and>
                                      accmodi im \<le> accmodi cm))) \<and>
       (\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and> 
       (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and> 
@@ -383,7 +383,7 @@
               entails (\<lambda> new old. 
                            (G\<turnstile>resTy (mthd new)\<preceq>resTy (mthd old) \<and>
                             accmodi old \<le> accmodi new \<and>
-      	                   \<not> is_static old)))  \<and>
+                           \<not> is_static old)))  \<and>
             (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c)) 
               hiding methd G (super c)
               under (\<lambda> new old. G\<turnstile>new hides old)
@@ -398,8 +398,8 @@
    \<lbrakk>\<not>is_iface G C;
     (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
         (\<forall>s. \<forall> im \<in> imethds G I s.
-      	    (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
-      	                             \<not> is_static cm \<and>
+            (\<exists> cm \<in> methd  G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
+                                     \<not> is_static cm \<and>
                                      accmodi im \<le> accmodi cm))); 
       \<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f; unique (cfields c); 
       \<forall>m\<in>set (methods c). wf_mdecl G C m; unique (methods c);
@@ -414,10 +414,10 @@
                        (G,sig\<turnstile>new overrides\<^sub>S old 
                         \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
                              accmodi old \<le> accmodi new \<and>
-      	                     \<not>is_static old)) \<and>
+                             \<not>is_static old)) \<and>
                        (G,sig\<turnstile>new hides old 
                          \<longrightarrow> (accmodi old \<le> accmodi new \<and>
-      	                      is_static old)))) 
+                              is_static old)))) 
             ))\<rbrakk> \<Longrightarrow> P
   \<rbrakk> \<Longrightarrow> P"
 by (unfold wf_cdecl_def) simp
@@ -522,11 +522,11 @@
 constdefs
   wf_prog  :: "prog \<Rightarrow> bool"
  "wf_prog G \<equiv> let is = ifaces G; cs = classes G in
-	         ObjectC \<in> set cs \<and> 
+                 ObjectC \<in> set cs \<and> 
                 (\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package) \<and>
                 (\<forall>xn. SXcptC xn \<in> set cs) \<and>
-		(\<forall>i\<in>set is. wf_idecl G i) \<and> unique is \<and>
-		(\<forall>c\<in>set cs. wf_cdecl G c) \<and> unique cs"
+                (\<forall>i\<in>set is. wf_idecl G i) \<and> unique is \<and>
+                (\<forall>c\<in>set cs. wf_cdecl G c) \<and> unique cs"
 
 lemma wf_prog_idecl: "\<lbrakk>iface G I = Some i; wf_prog G\<rbrakk> \<Longrightarrow> wf_idecl G (I,i)"
 apply (unfold wf_prog_def Let_def)
@@ -593,7 +593,7 @@
 done
 
 lemma wf_ObjectC [simp]: 
-	"wf_cdecl G ObjectC = (\<not>is_iface G Object \<and> Ball (set Object_mdecls)
+        "wf_cdecl G ObjectC = (\<not>is_iface G Object \<and> Ball (set Object_mdecls)
   (wf_mdecl G Object) \<and> unique Object_mdecls)"
 apply (unfold wf_cdecl_def ws_cdecl_def ObjectC_def)
 apply (auto intro: da.Skip)
@@ -753,26 +753,26 @@
       let ?inherited = "Un_tables (imethds G ` set (isuperIfs i))"
       let ?new = "(Option.set \<circ> table_of (map (\<lambda>(s, mh). (s, I, mh)) (imethods i)))"
       from if_I wf im have imethds:"im \<in> (?inherited \<oplus>\<oplus> ?new) sig"
-	by (simp add: imethds_rec)
+        by (simp add: imethds_rec)
       from wf if_I have 
-	wf_supI: "\<forall> J. J \<in> set (isuperIfs i) \<longrightarrow> (\<exists> j. iface G J = Some j)"
-	by (blast dest: wf_prog_idecl wf_idecl_supD is_acc_ifaceD)
+        wf_supI: "\<forall> J. J \<in> set (isuperIfs i) \<longrightarrow> (\<exists> j. iface G J = Some j)"
+        by (blast dest: wf_prog_idecl wf_idecl_supD is_acc_ifaceD)
       from wf if_I have
-	"\<forall> im \<in> set (imethods i). \<not> is_static im \<and> accmodi im = Public"
-	by (auto dest!: wf_prog_idecl wf_idecl_mhead)
+        "\<forall> im \<in> set (imethods i). \<not> is_static im \<and> accmodi im = Public"
+        by (auto dest!: wf_prog_idecl wf_idecl_mhead)
       then have new_ok: "\<forall> im. table_of (imethods i) sig = Some im 
                          \<longrightarrow>  \<not> is_static im \<and> accmodi im = Public"
-	by (auto dest!: table_of_Some_in_set)
+        by (auto dest!: table_of_Some_in_set)
       show ?thesis
-	proof (cases "?new sig = {}")
-	  case True
-	  from True wf wf_supI if_I imethds hyp 
-	  show ?thesis by (auto simp del:  split_paired_All)  
-	next
-	  case False
-	  from False wf wf_supI if_I imethds new_ok hyp 
-	  show ?thesis by (auto dest: wf_idecl_hidings hidings_entailsD)
-	qed
+        proof (cases "?new sig = {}")
+          case True
+          from True wf wf_supI if_I imethds hyp 
+          show ?thesis by (auto simp del:  split_paired_All)  
+        next
+          case False
+          from False wf wf_supI if_I imethds new_ok hyp 
+          show ?thesis by (auto dest: wf_idecl_hidings hidings_entailsD)
+        qed
       qed
     qed
   with asm show ?thesis by (auto simp del: split_paired_All)
@@ -870,14 +870,14 @@
     proof -
       from stat_override 
       have "accmodi old \<noteq> Private"
-	by (rule no_Private_stat_override)
+        by (rule no_Private_stat_override)
       moreover
       from stat_override wf
       have "accmodi old \<le> accmodi new"
-	by (auto dest: wf_prog_stat_overridesD)
+        by (auto dest: wf_prog_stat_overridesD)
       ultimately
       show ?thesis
-	by (auto dest: acc_modi_bottom)
+        by (auto dest: acc_modi_bottom)
     qed
     with Direct resTy_widen not_static_old 
     show "?Overrides new old" 
@@ -935,50 +935,50 @@
        by (cases old) (auto dest: member_of_declC)
       have "?P C old"
       proof (cases "G\<turnstile>mid (msig old) undeclared_in C")
-	case True
-	with inheritable super accessible_super member_super
-	have "G\<turnstile>Method old member_of C"
-	  by (cases old) (auto intro: members.Inherited)
-	then show ?thesis
-	  by auto
+        case True
+        with inheritable super accessible_super member_super
+        have "G\<turnstile>Method old member_of C"
+          by (cases old) (auto intro: members.Inherited)
+        then show ?thesis
+          by auto
       next
-	case False
-	then obtain new_member where
-	     "G\<turnstile>new_member declared_in C" and
+        case False
+        then obtain new_member where
+             "G\<turnstile>new_member declared_in C" and
              "mid (msig old) = memberid new_member"
           by (auto dest: not_undeclared_declared)
-	then obtain new where
-	          new: "G\<turnstile>Method new declared_in C" and
+        then obtain new where
+                  new: "G\<turnstile>Method new declared_in C" and
                eq_sig: "msig old = msig new" and
-	    declC_new: "declclass new = C" 
-	  by (cases new_member) auto
-	then have member_new: "G\<turnstile>Method new member_of C"
-	  by (cases new) (auto intro: members.Immediate)
-	from declC_new super member_super
-	have subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old"
-	  by (auto dest!: member_of_subclseq_declC
-	            dest: r_into_trancl intro: trancl_rtrancl_trancl)
-	show ?thesis
-	proof (cases "is_static new")
-	  case False
-	  with eq_sig declC_new new old_declared inheritable
-	       super member_super subcls_new_old
-	  have "G\<turnstile>new overrides\<^sub>S old"
-	    by (auto intro!: stat_overridesR.Direct)
-	  with member_new show ?thesis
-	    by blast
-	next
-	  case True
-	  with eq_sig declC_new subcls_new_old new old_declared inheritable
-	  have "G\<turnstile>new hides old"
-	    by (auto intro: hidesI)    
-	  with wf 
-	  have "is_static old"
-	    by (blast dest: wf_prog_hidesD)
-	  with instance_method
-	  show ?thesis
-	    by (contradiction)
-	qed
+            declC_new: "declclass new = C" 
+          by (cases new_member) auto
+        then have member_new: "G\<turnstile>Method new member_of C"
+          by (cases new) (auto intro: members.Immediate)
+        from declC_new super member_super
+        have subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old"
+          by (auto dest!: member_of_subclseq_declC
+                    dest: r_into_trancl intro: trancl_rtrancl_trancl)
+        show ?thesis
+        proof (cases "is_static new")
+          case False
+          with eq_sig declC_new new old_declared inheritable
+               super member_super subcls_new_old
+          have "G\<turnstile>new overrides\<^sub>S old"
+            by (auto intro!: stat_overridesR.Direct)
+          with member_new show ?thesis
+            by blast
+        next
+          case True
+          with eq_sig declC_new subcls_new_old new old_declared inheritable
+          have "G\<turnstile>new hides old"
+            by (auto intro: hidesI)    
+          with wf 
+          have "is_static old"
+            by (blast dest: wf_prog_hidesD)
+          with instance_method
+          show ?thesis
+            by (contradiction)
+        qed
       qed
     } note hyp_member_super = this
     from subclsC cls_C 
@@ -991,60 +991,60 @@
       assume "super c = declclass old"
       with old_declared 
       have "G\<turnstile>Method old member_of (super c)" 
-	by (cases old) (auto intro: members.Immediate)
+        by (cases old) (auto intro: members.Immediate)
       with inheritable instance_method 
       show ?thesis
-	by (blast dest: hyp_member_super)
+        by (blast dest: hyp_member_super)
     next
       case Subcls
       assume "G\<turnstile>super c\<prec>\<^sub>C declclass old"
       moreover
       from inheritable accmodi_old
       have "G \<turnstile>Method old inheritable_in pid (super c)"
-	by (cases "accmodi old") (auto simp add: inheritable_in_def)
+        by (cases "accmodi old") (auto simp add: inheritable_in_def)
       ultimately
       have "?P (super c) old"
-	by (blast dest: hyp)
+        by (blast dest: hyp)
       then show ?thesis
       proof
-	assume "G \<turnstile>Method old member_of super c"
-	with inheritable instance_method
-	show ?thesis
-	  by (blast dest: hyp_member_super)
+        assume "G \<turnstile>Method old member_of super c"
+        with inheritable instance_method
+        show ?thesis
+          by (blast dest: hyp_member_super)
       next
-	assume "\<exists>new. G \<turnstile> new overrides\<^sub>S old \<and> G \<turnstile>Method new member_of super c"
-	then obtain super_new where
-	  super_new_override:  "G \<turnstile> super_new overrides\<^sub>S old" and
+        assume "\<exists>new. G \<turnstile> new overrides\<^sub>S old \<and> G \<turnstile>Method new member_of super c"
+        then obtain super_new where
+          super_new_override:  "G \<turnstile> super_new overrides\<^sub>S old" and
             super_new_member:  "G \<turnstile>Method super_new member_of super c"
-	  by blast
-	from super_new_override wf
-	have "accmodi old \<le> accmodi super_new"
-	  by (auto dest: wf_prog_stat_overridesD)
-	with inheritable accmodi_old
-	have "G \<turnstile>Method super_new inheritable_in pid C"
-	  by (auto simp add: inheritable_in_def 
-	              split: acc_modi.splits
+          by blast
+        from super_new_override wf
+        have "accmodi old \<le> accmodi super_new"
+          by (auto dest: wf_prog_stat_overridesD)
+        with inheritable accmodi_old
+        have "G \<turnstile>Method super_new inheritable_in pid C"
+          by (auto simp add: inheritable_in_def 
+                      split: acc_modi.splits
                        dest: acc_modi_le_Dests)
-	moreover
-	from super_new_override 
-	have "\<not> is_static super_new"
-	  by (auto dest: stat_overrides_commonD)
-	moreover
-	note super_new_member
-	ultimately have "?P C super_new"
-	  by (auto dest: hyp_member_super)
-	then show ?thesis
-	proof 
-	  assume "G \<turnstile>Method super_new member_of C"
-	  with super_new_override
-	  show ?thesis
-	    by blast
-	next
-	  assume "\<exists>new. G \<turnstile> new overrides\<^sub>S super_new \<and> 
+        moreover
+        from super_new_override 
+        have "\<not> is_static super_new"
+          by (auto dest: stat_overrides_commonD)
+        moreover
+        note super_new_member
+        ultimately have "?P C super_new"
+          by (auto dest: hyp_member_super)
+        then show ?thesis
+        proof 
+          assume "G \<turnstile>Method super_new member_of C"
+          with super_new_override
+          show ?thesis
+            by blast
+        next
+          assume "\<exists>new. G \<turnstile> new overrides\<^sub>S super_new \<and> 
                   G \<turnstile>Method new member_of C"
-	  with super_new_override show ?thesis
-	    by (blast intro: stat_overridesR.Indirect) 
-	qed
+          with super_new_override show ?thesis
+            by (blast intro: stat_overridesR.Indirect) 
+        qed
       qed
     qed
   qed
@@ -1095,46 +1095,46 @@
       assume "G \<turnstile>Method old member_of declclass new"
       then have "G\<turnstile>mid (msig old) undeclared_in declclass new"
       proof cases
-	case Immediate 
-	with subcls_new_old wf show ?thesis 	
-	  by (auto dest: subcls_irrefl)
+        case Immediate 
+        with subcls_new_old wf show ?thesis     
+          by (auto dest: subcls_irrefl)
       next
-	case Inherited
-	then show ?thesis
-	  by (cases old) auto
+        case Inherited
+        then show ?thesis
+          by (cases old) auto
       qed
       with eq_sig_new_old new_declared
       show ?thesis
-	by (cases old,cases new) (auto dest!: declared_not_undeclared)
+        by (cases old,cases new) (auto dest!: declared_not_undeclared)
     next
       case (Overriding new') 
       assume stat_override_new': "G \<turnstile> new' overrides\<^sub>S old"
       then have "msig new' = msig old"
-	by (auto dest: stat_overrides_commonD)
+        by (auto dest: stat_overrides_commonD)
       with eq_sig_new_old have eq_sig_new_new': "msig new=msig new'"
-	by simp
+        by simp
       assume "G \<turnstile>Method new' member_of declclass new"
       then show ?thesis
       proof (cases)
-	case Immediate
-	then have declC_new: "declclass new' = declclass new" 
-	  by auto
-	from Immediate 
-	have "G\<turnstile>Method new' declared_in declclass new"
-	  by (cases new') auto
-	with new_declared eq_sig_new_new' declC_new 
-	have "new=new'"
-	  by (cases new, cases new') (auto dest: unique_declared_in) 
-	with stat_override_new'
-	show ?thesis
-	  by simp
+        case Immediate
+        then have declC_new: "declclass new' = declclass new" 
+          by auto
+        from Immediate 
+        have "G\<turnstile>Method new' declared_in declclass new"
+          by (cases new') auto
+        with new_declared eq_sig_new_new' declC_new 
+        have "new=new'"
+          by (cases new, cases new') (auto dest: unique_declared_in) 
+        with stat_override_new'
+        show ?thesis
+          by simp
       next
-	case Inherited
-	then have "G\<turnstile>mid (msig new') undeclared_in declclass new"
-	  by (cases new') (auto)
-	with eq_sig_new_new' new_declared
-	show ?thesis
-	  by (cases new,cases new') (auto dest!: declared_not_undeclared)
+        case Inherited
+        then have "G\<turnstile>mid (msig new') undeclared_in declclass new"
+          by (cases new') (auto)
+        with eq_sig_new_new' new_declared
+        show ?thesis
+          by (cases new,cases new') (auto dest!: declared_not_undeclared)
       qed
     qed
   next
@@ -1151,12 +1151,12 @@
     proof -
       from stat_override_inter_old wf 
       have "accmodi old \<le> accmodi inter"
-	by (auto dest: wf_prog_stat_overridesD)
+        by (auto dest: wf_prog_stat_overridesD)
       with stat_override_inter_old accmodi_old
       show ?thesis
-	by (auto dest!: no_Private_stat_override
+        by (auto dest!: no_Private_stat_override
                  split: acc_modi.splits 
-	         dest: acc_modi_le_Dests)
+                 dest: acc_modi_le_Dests)
     qed
     ultimately show "?Overrides new old"
       by (blast intro: stat_overridesR.Indirect)
@@ -1297,59 +1297,59 @@
       note same_pack_old_inter = this
       show ?thesis
       proof (cases "pid (declclass inter) = pid (declclass new)")
-	case True
-	with same_pack_old_inter outside_pack
-	show ?thesis
-	  by auto
+        case True
+        with same_pack_old_inter outside_pack
+        show ?thesis
+          by auto
       next
-	case False
-	note diff_pack_inter_new = this
-	show ?thesis
-	proof (cases "accmodi inter = Package")
-	  case True
-	  with diff_pack_inter_new hyp_new_inter  
-	  obtain newinter where
-	    over_new_newinter: "G \<turnstile> new overrides newinter" and
+        case False
+        note diff_pack_inter_new = this
+        show ?thesis
+        proof (cases "accmodi inter = Package")
+          case True
+          with diff_pack_inter_new hyp_new_inter  
+          obtain newinter where
+            over_new_newinter: "G \<turnstile> new overrides newinter" and
             over_newinter_inter: "G \<turnstile> newinter overrides inter" and 
             eq_pid: "pid (declclass inter) = pid (declclass newinter)" and
             accmodi_newinter: "Protected \<le> accmodi newinter"
-	    by auto
-	  from over_newinter_inter override_inter_old
-	  have "G\<turnstile>newinter overrides old"
-	    by (rule overridesR.Indirect)
-	  moreover
-	  from eq_pid same_pack_old_inter 
-	  have "pid (declclass old) = pid (declclass newinter)"
-	    by simp
-	  moreover
-	  note over_new_newinter accmodi_newinter
-	  ultimately show ?thesis
-	    by blast
-	next
-	  case False
-	  with override_new_inter
-	  have "Protected \<le> accmodi inter"
-	    by (cases "accmodi inter") (auto dest: no_Private_override)
-	  with override_new_inter override_inter_old same_pack_old_inter
-	  show ?thesis
-	    by blast
-	qed
+            by auto
+          from over_newinter_inter override_inter_old
+          have "G\<turnstile>newinter overrides old"
+            by (rule overridesR.Indirect)
+          moreover
+          from eq_pid same_pack_old_inter 
+          have "pid (declclass old) = pid (declclass newinter)"
+            by simp
+          moreover
+          note over_new_newinter accmodi_newinter
+          ultimately show ?thesis
+            by blast
+        next
+          case False
+          with override_new_inter
+          have "Protected \<le> accmodi inter"
+            by (cases "accmodi inter") (auto dest: no_Private_override)
+          with override_new_inter override_inter_old same_pack_old_inter
+          show ?thesis
+            by blast
+        qed
       qed
     next
       case False
       with accmodi_old hyp_inter_old
       obtain newinter where
-	over_inter_newinter: "G \<turnstile> inter overrides newinter" and
+        over_inter_newinter: "G \<turnstile> inter overrides newinter" and
           over_newinter_old: "G \<turnstile> newinter overrides old" and 
                 eq_pid: "pid (declclass old) = pid (declclass newinter)" and
-	accmodi_newinter: "Protected \<le> accmodi newinter"
-	by auto
+        accmodi_newinter: "Protected \<le> accmodi newinter"
+        by auto
       from override_new_inter over_inter_newinter 
       have "G \<turnstile> new overrides newinter"
-	by (rule overridesR.Indirect)
+        by (rule overridesR.Indirect)
       with eq_pid over_newinter_old accmodi_newinter
       show ?thesis
-	by blast
+        by blast
     qed
   qed
 qed
@@ -1379,13 +1379,13 @@
     proof (cases "?table sig")
       case None
       from this methd_C have "?filter (methd G (super c)) sig = Some m"
-	by simp
+        by simp
       moreover
       from wf cls_C False obtain sup where "class G (super c) = Some sup"
-	by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
+        by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
       moreover note wf False cls_C  
       ultimately have "G\<turnstile>super c \<preceq>\<^sub>C declclass m"  
-	by (auto intro: Hyp [rule_format])
+        by (auto intro: Hyp [rule_format])
       moreover from cls_C False have  "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c" by (rule subcls1I)
       ultimately show ?thesis by - (rule rtrancl_into_rtrancl2)
     next
@@ -1499,11 +1499,11 @@
     proof (cases rule: methd_rec_Some_cases)
       case NewMethod
       with clsC show ?thesis
-	by (auto dest: method_declared_inI intro!: members.Immediate)
+        by (auto dest: method_declared_inI intro!: members.Immediate)
     next
       case InheritedMethod
       then show "?thesis"
-	by (blast dest: inherits_member)
+        by (blast dest: inherits_member)
     qed
   qed
 qed
@@ -1611,7 +1611,7 @@
       by (auto dest: ws_prog_cdeclD)
     from clsC wf neq_C_Obj 
     have superAccessible: "G\<turnstile>(Class (super c)) accessible_in (pid C)" and
-	 subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
+         subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
       by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD
               intro: subcls1I)
     show "\<exists>new. ?Constraint C new old"
@@ -1619,87 +1619,87 @@
       case False
       from False Subcls 
       have eq_C_D: "C=D"
-	by (auto dest: subclseq_superD)
+        by (auto dest: subclseq_superD)
       with old 
       have "?Constraint C old old"
-	by auto
+        by auto
       with eq_C_D 
       show "\<exists> new. ?Constraint C new old" by auto
     next
       case True
       with hyp obtain super_method
-	where super: "?Constraint (super c) super_method old" by blast
+        where super: "?Constraint (super c) super_method old" by blast
       from super not_static_old
       have not_static_super: "\<not> is_static super_method"
-	by (auto dest!: stat_overrides_commonD)
+        by (auto dest!: stat_overrides_commonD)
       from super old wf accmodi_old
       have accmodi_super_method: "Protected \<le> accmodi super_method"
-	by (auto dest!: wf_prog_stat_overridesD)
+        by (auto dest!: wf_prog_stat_overridesD)
       from super accmodi_old wf
       have inheritable: "G\<turnstile>Methd sig super_method inheritable_in (pid C)"
-	by (auto dest!: wf_prog_stat_overridesD
+        by (auto dest!: wf_prog_stat_overridesD
                         acc_modi_le_Dests
-              simp add: inheritable_in_def)	           
+              simp add: inheritable_in_def)                
       from super wf is_cls_super
       have member: "G\<turnstile>Methd sig super_method member_of (super c)"
-	by (auto intro: methd_member_of) 
+        by (auto intro: methd_member_of) 
       from member
       have decl_super_method:
         "G\<turnstile>Methd sig super_method declared_in (declclass super_method)"
-	by (auto dest: member_of_declC)
+        by (auto dest: member_of_declC)
       from super subcls1_C_super ws is_cls_super 
       have subcls_C_super: "G\<turnstile>C \<prec>\<^sub>C (declclass super_method)"
-	by (auto intro: rtrancl_into_trancl2 dest: methd_declC) 
+        by (auto intro: rtrancl_into_trancl2 dest: methd_declC) 
       show "\<exists> new. ?Constraint C new old"
       proof (cases "methd G C sig")
-	case None
-	have "methd G (super c) sig = None"
-	proof -
-	  from clsC ws None 
-	  have no_new: "table_of (methods c) sig = None" 
-	    by (auto simp add: methd_rec)
-	  with clsC 
-	  have undeclared: "G\<turnstile>mid sig undeclared_in C"
-	    by (auto simp add: undeclared_in_def cdeclaredmethd_def)
-	  with inheritable member superAccessible subcls1_C_super
-	  have inherits: "G\<turnstile>C inherits (method sig super_method)"
-	    by (auto simp add: inherits_def)
-	  with clsC ws no_new super neq_C_Obj
-	  have "methd G C sig = Some super_method"
-	    by (auto simp add: methd_rec map_add_def intro: filter_tab_SomeI)
+        case None
+        have "methd G (super c) sig = None"
+        proof -
+          from clsC ws None 
+          have no_new: "table_of (methods c) sig = None" 
+            by (auto simp add: methd_rec)
+          with clsC 
+          have undeclared: "G\<turnstile>mid sig undeclared_in C"
+            by (auto simp add: undeclared_in_def cdeclaredmethd_def)
+          with inheritable member superAccessible subcls1_C_super
+          have inherits: "G\<turnstile>C inherits (method sig super_method)"
+            by (auto simp add: inherits_def)
+          with clsC ws no_new super neq_C_Obj
+          have "methd G C sig = Some super_method"
+            by (auto simp add: methd_rec map_add_def intro: filter_tab_SomeI)
           with None show ?thesis
-	    by simp
-	qed
-	with super show ?thesis by auto
+            by simp
+        qed
+        with super show ?thesis by auto
       next
-	case (Some new)
-	from this ws clsC neq_C_Obj
-	show ?thesis
-	proof (cases rule: methd_rec_Some_cases)
-	  case InheritedMethod
-	  with super Some show ?thesis 
-	    by auto
-	next
-	  case NewMethod
-	  assume new: "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig 
+        case (Some new)
+        from this ws clsC neq_C_Obj
+        show ?thesis
+        proof (cases rule: methd_rec_Some_cases)
+          case InheritedMethod
+          with super Some show ?thesis 
+            by auto
+        next
+          case NewMethod
+          assume new: "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig 
                        = Some new"
-	  from new 
-	  have declcls_new: "declclass new = C" 
-	    by auto
-	  from wf clsC neq_C_Obj super new not_static_super accmodi_super_method
-	  have not_static_new: "\<not> is_static new" 
-	    by (auto dest: wf_prog_staticD) 
-	  from clsC new
-	  have decl_new: "G\<turnstile>Methd sig new declared_in C"
-	    by (auto simp add: declared_in_def cdeclaredmethd_def)
-	  from not_static_new decl_new decl_super_method
-	       member subcls1_C_super inheritable declcls_new subcls_C_super 
-	  have "G,sig\<turnstile> new overrides\<^sub>S super_method"
-	    by (auto intro: stat_overridesR.Direct) 
-	  with super Some
-	  show ?thesis
-	    by (auto intro: stat_overridesR.Indirect)
-	qed
+          from new 
+          have declcls_new: "declclass new = C" 
+            by auto
+          from wf clsC neq_C_Obj super new not_static_super accmodi_super_method
+          have not_static_new: "\<not> is_static new" 
+            by (auto dest: wf_prog_staticD) 
+          from clsC new
+          have decl_new: "G\<turnstile>Methd sig new declared_in C"
+            by (auto simp add: declared_in_def cdeclaredmethd_def)
+          from not_static_new decl_new decl_super_method
+               member subcls1_C_super inheritable declcls_new subcls_C_super 
+          have "G,sig\<turnstile> new overrides\<^sub>S super_method"
+            by (auto intro: stat_overridesR.Direct) 
+          with super Some
+          show ?thesis
+            by (auto intro: stat_overridesR.Indirect)
+        qed
       qed
     qed
   qed
@@ -1749,7 +1749,7 @@
       by (auto dest!: wf_prog_stat_overridesD)
   qed
 qed
- 	      
+              
 (* local lemma *)
 lemma bexI': "x \<in> A \<Longrightarrow> P x \<Longrightarrow> \<exists>x\<in>A. P x" by blast
 lemma ballE': "\<forall>x\<in>A. P x \<Longrightarrow> (x \<notin> A \<Longrightarrow> Q) \<Longrightarrow> (P x \<Longrightarrow> Q) \<Longrightarrow> Q" by blast
@@ -1789,32 +1789,32 @@
       case True
       with ifI SI hyp wf jm 
       show "?thesis" 
-	by (auto simp add: imethds_rec) 
+        by (auto simp add: imethds_rec) 
     next
       case False
       from ifI wf False
       have imethds: "imethds G I sig = ?newMethods sig"
-	by (simp add: imethds_rec)
+        by (simp add: imethds_rec)
       from False
       obtain im where
         imdef: "im \<in> ?newMethods sig" 
-	by (blast)
+        by (blast)
       with imethds 
       have im: "im \<in> imethds G I sig"
-	by (blast)
+        by (blast)
       with im wf ifI 
       obtain
-	 imStatic: "\<not> is_static im" and
+         imStatic: "\<not> is_static im" and
          imPublic: "accmodi im = Public"
-	by (auto dest!: imethds_wf_mhead)	
+        by (auto dest!: imethds_wf_mhead)       
       from ifI wf 
       have wf_I: "wf_idecl G (I,i)" 
-	by (rule wf_prog_idecl)
+        by (rule wf_prog_idecl)
       with SI wf  
       obtain si where
-	 ifSI: "iface G SI = Some si" and
-	wf_SI: "wf_idecl G (SI,si)" 
-	by (auto dest!: wf_idecl_supD is_acc_ifaceD
+         ifSI: "iface G SI = Some si" and
+        wf_SI: "wf_idecl G (SI,si)" 
+        by (auto dest!: wf_idecl_supD is_acc_ifaceD
                   dest: wf_prog_idecl)
       from jm hyp 
       obtain sim::"qtname \<times> mhead"  where
@@ -1822,24 +1822,24 @@
          eq_static_sim_jm: "is_static sim = is_static jm" and 
          eq_access_sim_jm: "accmodi sim = accmodi jm" and 
         resTy_widen_sim_jm: "G\<turnstile>resTy sim\<preceq>resTy jm"
-	by blast
+        by blast
       with wf_I SI imdef sim 
       have "G\<turnstile>resTy im\<preceq>resTy sim"   
-	by (auto dest!: wf_idecl_hidings hidings_entailsD)
+        by (auto dest!: wf_idecl_hidings hidings_entailsD)
       with wf resTy_widen_sim_jm 
       have resTy_widen_im_jm: "G\<turnstile>resTy im\<preceq>resTy jm"
-	by (blast intro: widen_trans)
+        by (blast intro: widen_trans)
       from sim wf ifSI  
       obtain
-	simStatic: "\<not> is_static sim" and
+        simStatic: "\<not> is_static sim" and
         simPublic: "accmodi sim = Public"
-	by (auto dest!: imethds_wf_mhead)
+        by (auto dest!: imethds_wf_mhead)
       from im 
            imStatic simStatic eq_static_sim_jm
            imPublic simPublic eq_access_sim_jm
            resTy_widen_im_jm
       show ?thesis 
-	by auto 
+        by auto 
     qed
   qed
 qed
@@ -1999,7 +1999,7 @@
       case Some
       from Some clsC nObj ws m declC
       show "?thesis" 
-	by (auto simp add: methd_rec
+        by (auto simp add: methd_rec
                      dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
     qed
   qed
@@ -2085,14 +2085,14 @@
       case NewMethod
       with ifI hyp_newmethod
       show ?thesis
-	by auto
+        by auto
     next
       case (InheritedMethod J)
       assume "J \<in> set (isuperIfs i)" 
              "new \<in> imethds G J sig"
       with hyp 
       show "?thesis"
-	by auto
+        by auto
     qed
   qed
 qed
@@ -2250,46 +2250,46 @@
        case True
        with statI 
        have dynlookup: "dynlookup G statT dynC sig = dynmethd G Object dynC sig"
-	 by (simp add: dynlookup_def dynimethd_def)
+         by (simp add: dynlookup_def dynimethd_def)
        from wf dynC 
        have subclsObj: "G\<turnstile>dynC \<preceq>\<^sub>C Object"
-	 by (auto intro: subcls_ObjectI)
+         by (auto intro: subcls_ObjectI)
        from wf dynC dynC_Prop istype sm subclsObj 
        obtain dm where
-	 "dynmethd G Object dynC sig = Some dm"
-	 "is_static dm = is_static sm" 
-	 "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd sm)"  
-	 by (auto dest!: ws_dynmethd accmethd_SomeD 
+         "dynmethd G Object dynC sig = Some dm"
+         "is_static dm = is_static sm" 
+         "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd sm)"  
+         by (auto dest!: ws_dynmethd accmethd_SomeD 
                   intro: class_Object [OF wf] intro: that)
        with dynlookup eq_mheads
        show ?thesis 
-	 by (cases emh type: *) (auto)
+         by (cases emh type: *) (auto)
      next
        case False
        with statI
        have dynlookup: "dynlookup G statT dynC sig = methd G dynC sig"
-	 by (simp add: dynlookup_def dynimethd_def)
+         by (simp add: dynlookup_def dynimethd_def)
        from istype statI
        have "is_iface G I"
-	 by auto
+         by auto
        with wf sm nPriv False 
        obtain im where
-	      im: "im \<in> imethds G I sig" and
-	 eq_stat: "is_static im = is_static sm" and
+              im: "im \<in> imethds G I sig" and
+         eq_stat: "is_static im = is_static sm" and
          resProp: "G\<turnstile>resTy (mthd im)\<preceq>resTy (mthd sm)"
-	 by (auto dest: wf_imethds_hiding_objmethdsD accmethd_SomeD)
+         by (auto dest: wf_imethds_hiding_objmethdsD accmethd_SomeD)
        from im wf statI istype eq_stat eq_mheads
        have not_static_sm: "\<not> is_static emh"
-	 by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead)
+         by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead)
        from im wf dynC_Prop dynC istype statI not_static_sm
        obtain dm where
-	 "methd G dynC sig = Some dm"
-	 "is_static dm = is_static im" 
-	 "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd im)" 
-	 by (auto dest: implmt_methd)
+         "methd G dynC sig = Some dm"
+         "is_static dm = is_static im" 
+         "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd im)" 
+         by (auto dest: implmt_methd)
        with wf eq_stat resProp dynlookup eq_mheads
        show ?thesis 
-	 by (cases emh type: *) (auto intro: widen_trans)
+         by (cases emh type: *) (auto intro: widen_trans)
      qed
   next
     case Array_Object_methd
@@ -2382,7 +2382,7 @@
 proof -
   assume asm: "class G C = Some c" "wf_prog G" "methd G C sig = Some m"
   have "wf_prog G  \<longrightarrow> 
-	   (\<forall> c m. class G C = Some c \<longrightarrow>  methd G C sig = Some m 
+           (\<forall> c m. class G C = Some c \<longrightarrow>  methd G C sig = Some m 
                    \<longrightarrow>  methd G (declclass m) sig = Some m)"      (is "?P G C") 
   proof (rule class_rec.induct,intro allI impI)
     fix G C c m
@@ -2400,23 +2400,23 @@
       case False
       with cls_C wf m
       have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m "
-	by (simp add: methd_rec)
+        by (simp add: methd_rec)
       show ?thesis
       proof (cases "?table sig")
-	case None
-	from this methd_C have "?filter (methd G (super c)) sig = Some m"
-	  by simp
-	moreover
-	from wf cls_C False obtain sup where "class G (super c) = Some sup"
-	  by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
-	moreover note wf False cls_C 
-	ultimately show ?thesis by (auto intro: hyp [rule_format])
+        case None
+        from this methd_C have "?filter (methd G (super c)) sig = Some m"
+          by simp
+        moreover
+        from wf cls_C False obtain sup where "class G (super c) = Some sup"
+          by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
+        moreover note wf False cls_C 
+        ultimately show ?thesis by (auto intro: hyp [rule_format])
       next
-	case Some
-	from this methd_C m show ?thesis by auto 
+        case Some
+        from this methd_C m show ?thesis by auto 
       qed
     qed
-  qed	
+  qed   
   with asm show ?thesis by auto
 qed
 
@@ -2650,15 +2650,15 @@
       assume "dynC=statC"
       moreover
       from is_cls_statC obtain c
-	where "class G statC = Some c"
-	by auto
+        where "class G statC = Some c"
+        by auto
       moreover 
       note statM ws dynmethd 
       ultimately
       have "newM=statM" 
-	by (auto simp add: dynmethd_C_C)
+        by (auto simp add: dynmethd_C_C)
       with neq show ?thesis 
-	by (contradiction)
+        by (contradiction)
     next
       case Subcls then show ?thesis .
     qed 
@@ -2771,26 +2771,26 @@
       case False
       with statT dynC_prop 
       have widen_dynC: "G\<turnstile>Class dynC \<preceq> RefT statT"
-	by simp
+        by simp
       from statT widen_dynC
       have implmnt: "G\<turnstile>dynC\<leadsto>I"
-	by auto    
+        by auto    
       from eq_static False 
       have not_static_dynM: "\<not> is_static dynM" 
-	by simp
+        by simp
       from iface_methd implmnt isif_I wf dynM'
       show ?thesis
-	by - (drule implmt_dynimethd_access, auto)
+        by - (drule implmt_dynimethd_access, auto)
     next
       case True
       assume "is_static emh"
       moreover
       from wf isT_statT statT im 
       have "\<not> is_static im"
-	by (auto dest: accimethdsD wf_prog_idecl imethds_wf_mhead)
+        by (auto dest: accimethdsD wf_prog_idecl imethds_wf_mhead)
       moreover note eq_mhds
       ultimately show ?thesis
-	by (cases emh) auto
+        by (cases emh) auto
     qed
   next
     case (Iface_Object_methd I statM)
@@ -2809,43 +2809,43 @@
       case True
       from dynM statT True
       have dynM': "dynmethd G Object dynC sig = Some dynM"
-	by (simp add: dynlookup_def dynimethd_def)
+        by (simp add: dynlookup_def dynimethd_def)
       from statT  
       have "G\<turnstile>RefT statT \<preceq>Class Object"
-	by auto
+        by auto
       with statM statT_acc stat_acc widen_dynC_Obj statT isT_statT 
         wf dynM' eq_static dynC_prop  
       show ?thesis
-	by - (drule dynmethd_access_prop,force+) 
+        by - (drule dynmethd_access_prop,force+) 
     next
       case False
       then obtain im where
-	im: "im \<in>  imethds G I sig"
-	by auto
+        im: "im \<in>  imethds G I sig"
+        by auto
       have not_static_emh: "\<not> is_static emh"
       proof -
-	from im statM statT isT_statT wf not_Private_statM 
-	have "is_static im = is_static statM"
-	  by (fastsimp dest: wf_imethds_hiding_objmethdsD)
-	with wf isT_statT statT im 
-	have "\<not> is_static statM"
-	  by (auto dest: wf_prog_idecl imethds_wf_mhead)
-	with eq_mhds
-	show ?thesis  
-	  by (cases emh) auto
+        from im statM statT isT_statT wf not_Private_statM 
+        have "is_static im = is_static statM"
+          by (fastsimp dest: wf_imethds_hiding_objmethdsD)
+        with wf isT_statT statT im 
+        have "\<not> is_static statM"
+          by (auto dest: wf_prog_idecl imethds_wf_mhead)
+        with eq_mhds
+        show ?thesis  
+          by (cases emh) auto
       qed
       with statT dynC_prop
       have implmnt: "G\<turnstile>dynC\<leadsto>I"
-	by simp
+        by simp
       with isT_statT statT
       have isif_I: "is_iface G I"
-	by simp
+        by simp
       from dynM statT
       have dynM': "dynimethd G I dynC sig = Some dynM"
-	by (simp add: dynlookup_def) 
+        by (simp add: dynlookup_def) 
       from False implmnt isif_I wf dynM'
       show ?thesis
-	by - (drule implmt_dynimethd_access, auto)
+        by - (drule implmt_dynimethd_access, auto)
     qed
   next
     case (Array_Object_methd T statM)
--- a/src/HOL/Bali/WellType.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Bali/WellType.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -27,8 +27,8 @@
 \end{itemize}
 *}
 
-types	lenv
-	= "(lname, ty) table"  --{* local variables, including This and Result*}
+types lenv
+        = "(lname, ty) table"  --{* local variables, including This and Result*}
 
 record env = 
          prg:: "prog"    --{* program *}
@@ -121,7 +121,7 @@
    max_spec     :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
 
  "max_spec G S rt sig \<equiv>{m. m \<in>appl_methds G S rt sig \<and>
-		      (\<forall>m'\<in>appl_methds G S rt sig. more_spec G m' m \<longrightarrow> m'=m)}"
+                      (\<forall>m'\<in>appl_methds G S rt sig. more_spec G m' m \<longrightarrow> m'=m)}"
 
 
 lemma max_spec2appl_meths: 
@@ -262,47 +262,47 @@
 
 --{* well-typed statements *}
 
-| Skip:					"E,dt\<Turnstile>Skip\<Colon>\<surd>"
+| Skip:                                 "E,dt\<Turnstile>Skip\<Colon>\<surd>"
 
-| Expr:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Expr e\<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>"
+| 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>"
+| 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>"
+| 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>"
+          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>
+| 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>"
+                                         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>"
+| Fin:  "\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
+                                         E,dt\<Turnstile>c1 Finally c2\<Colon>\<surd>"
 
-| Init:	"\<lbrakk>is_class (prg E) C\<rbrakk> \<Longrightarrow>
-					 E,dt\<Turnstile>Init C\<Colon>\<surd>"
+| Init: "\<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 
@@ -312,69 +312,69 @@
 --{* 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"
+| 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.[]"
+| 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'"
+| 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"
+| 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"
+| 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"
+          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"
+           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)"
+          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"
+| 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'"
+| 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"
+| 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> 
+| 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)"
+                   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"
+          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 
@@ -386,11 +386,11 @@
     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;
+| 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"
+                                         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})
@@ -402,27 +402,27 @@
 --{* 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"
+| 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)"
+| 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"
+| 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>[]"
+| 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"
+| 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 *)
@@ -431,7 +431,7 @@
   "_ty_expr" :: "env \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_|-_:-_" [51,51,51] 50)
   "_ty_var"  :: "env \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_|-_:=_" [51,51,51] 50)
   "_ty_exprs":: "env \<Rightarrow> [expr list,
-		     \<spacespace> ty   list] \<Rightarrow> bool" ("_|-_:#_" [51,51,51] 50)
+                     \<spacespace> ty   list] \<Rightarrow> bool" ("_|-_:#_" [51,51,51] 50)
 
 syntax (xsymbols)
   "_wt"      :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>_"  [51,51,51] 50)
@@ -439,14 +439,14 @@
   "_ty_expr" :: "env \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>-_" [51,51,51] 50)
   "_ty_var"  :: "env \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>=_" [51,51,51] 50)
   "_ty_exprs" :: "env \<Rightarrow> [expr list,
-		    \<spacespace>  ty   list] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<doteq>_" [51,51,51] 50)
+                    \<spacespace>  ty   list] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<doteq>_" [51,51,51] 50)
 
 translations
-	"E\<turnstile>t\<Colon> T" == "E,empty_dt\<Turnstile>t\<Colon> T"
+        "E\<turnstile>t\<Colon> T" == "E,empty_dt\<Turnstile>t\<Colon> T"
         "E\<turnstile>s\<Colon>\<surd>"  == "E\<turnstile>In1r s\<Colon>Inl (PrimT Void)"
-	"E\<turnstile>e\<Colon>-T" == "E\<turnstile>In1l e\<Colon>Inl T"
-	"E\<turnstile>e\<Colon>=T" == "E\<turnstile>In2  e\<Colon>Inl T"
-	"E\<turnstile>e\<Colon>\<doteq>T" == "E\<turnstile>In3  e\<Colon>Inr T"
+        "E\<turnstile>e\<Colon>-T" == "E\<turnstile>In1l e\<Colon>Inl T"
+        "E\<turnstile>e\<Colon>=T" == "E\<turnstile>In2  e\<Colon>Inl T"
+        "E\<turnstile>e\<Colon>\<doteq>T" == "E\<turnstile>In3  e\<Colon>Inr T"
 
 
 declare not_None_eq [simp del] 
@@ -455,36 +455,36 @@
 declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *}
 
 inductive_cases wt_elim_cases [cases set]:
-	"E,dt\<Turnstile>In2  (LVar vn)               \<Colon>T"
-	"E,dt\<Turnstile>In2  ({accC,statDeclC,s}e..fn)\<Colon>T"
-	"E,dt\<Turnstile>In2  (e.[i])                 \<Colon>T"
-	"E,dt\<Turnstile>In1l (NewC C)                \<Colon>T"
-	"E,dt\<Turnstile>In1l (New T'[i])             \<Colon>T"
-	"E,dt\<Turnstile>In1l (Cast T' e)             \<Colon>T"
-	"E,dt\<Turnstile>In1l (e InstOf T')           \<Colon>T"
-	"E,dt\<Turnstile>In1l (Lit x)                 \<Colon>T"
+        "E,dt\<Turnstile>In2  (LVar vn)               \<Colon>T"
+        "E,dt\<Turnstile>In2  ({accC,statDeclC,s}e..fn)\<Colon>T"
+        "E,dt\<Turnstile>In2  (e.[i])                 \<Colon>T"
+        "E,dt\<Turnstile>In1l (NewC C)                \<Colon>T"
+        "E,dt\<Turnstile>In1l (New T'[i])             \<Colon>T"
+        "E,dt\<Turnstile>In1l (Cast T' e)             \<Colon>T"
+        "E,dt\<Turnstile>In1l (e InstOf T')           \<Colon>T"
+        "E,dt\<Turnstile>In1l (Lit x)                 \<Colon>T"
         "E,dt\<Turnstile>In1l (UnOp unop e)           \<Colon>T"
         "E,dt\<Turnstile>In1l (BinOp binop e1 e2)     \<Colon>T"
-	"E,dt\<Turnstile>In1l (Super)                 \<Colon>T"
-	"E,dt\<Turnstile>In1l (Acc va)                \<Colon>T"
-	"E,dt\<Turnstile>In1l (Ass va v)              \<Colon>T"
-	"E,dt\<Turnstile>In1l (e0 ? e1 : e2)          \<Colon>T"
-	"E,dt\<Turnstile>In1l ({accC,statT,mode}e\<cdot>mn({pT'}p))\<Colon>T"
-	"E,dt\<Turnstile>In1l (Methd C sig)           \<Colon>T"
-	"E,dt\<Turnstile>In1l (Body D blk)            \<Colon>T"
-	"E,dt\<Turnstile>In3  ([])                    \<Colon>Ts"
-	"E,dt\<Turnstile>In3  (e#es)                  \<Colon>Ts"
-	"E,dt\<Turnstile>In1r  Skip                   \<Colon>x"
-	"E,dt\<Turnstile>In1r (Expr e)                \<Colon>x"
-	"E,dt\<Turnstile>In1r (c1;; c2)               \<Colon>x"
+        "E,dt\<Turnstile>In1l (Super)                 \<Colon>T"
+        "E,dt\<Turnstile>In1l (Acc va)                \<Colon>T"
+        "E,dt\<Turnstile>In1l (Ass va v)              \<Colon>T"
+        "E,dt\<Turnstile>In1l (e0 ? e1 : e2)          \<Colon>T"
+        "E,dt\<Turnstile>In1l ({accC,statT,mode}e\<cdot>mn({pT'}p))\<Colon>T"
+        "E,dt\<Turnstile>In1l (Methd C sig)           \<Colon>T"
+        "E,dt\<Turnstile>In1l (Body D blk)            \<Colon>T"
+        "E,dt\<Turnstile>In3  ([])                    \<Colon>Ts"
+        "E,dt\<Turnstile>In3  (e#es)                  \<Colon>Ts"
+        "E,dt\<Turnstile>In1r  Skip                   \<Colon>x"
+        "E,dt\<Turnstile>In1r (Expr e)                \<Colon>x"
+        "E,dt\<Turnstile>In1r (c1;; c2)               \<Colon>x"
         "E,dt\<Turnstile>In1r (l\<bullet> c)                  \<Colon>x" 
-	"E,dt\<Turnstile>In1r (If(e) c1 Else c2)      \<Colon>x"
-	"E,dt\<Turnstile>In1r (l\<bullet> While(e) c)         \<Colon>x"
+        "E,dt\<Turnstile>In1r (If(e) c1 Else c2)      \<Colon>x"
+        "E,dt\<Turnstile>In1r (l\<bullet> While(e) c)         \<Colon>x"
         "E,dt\<Turnstile>In1r (Jmp jump)              \<Colon>x"
-	"E,dt\<Turnstile>In1r (Throw e)               \<Colon>x"
-	"E,dt\<Turnstile>In1r (Try c1 Catch(tn vn) c2)\<Colon>x"
-	"E,dt\<Turnstile>In1r (c1 Finally c2)         \<Colon>x"
-	"E,dt\<Turnstile>In1r (Init C)                \<Colon>x"
+        "E,dt\<Turnstile>In1r (Throw e)               \<Colon>x"
+        "E,dt\<Turnstile>In1r (Try c1 Catch(tn vn) c2)\<Colon>x"
+        "E,dt\<Turnstile>In1r (c1 Finally c2)         \<Colon>x"
+        "E,dt\<Turnstile>In1r (Init C)                \<Colon>x"
 declare not_None_eq [simp] 
 declare split_if [split] split_if_asm [split]
 declare split_paired_All [simp] split_paired_Ex [simp]
@@ -549,7 +549,7 @@
 \<Longrightarrow> E,dt\<Turnstile>Super\<Colon>-Class D"
 by (auto elim: wt.Super)
 
-lemma wt_FVar:	
+lemma wt_FVar:  
 "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; accfield (prg E) (cls E) C fn = Some (statDeclC,f);
   sf=is_static f; fT=(type f); accC=cls E\<rbrakk> 
 \<Longrightarrow> E,dt\<Turnstile>{accC,statDeclC,sf}e..fn\<Colon>=fT"
@@ -677,7 +677,7 @@
 lemma typeof_empty_is_type [rule_format (no_asm)]: 
   "typeof (\<lambda>a. None) v = Some T \<longrightarrow> is_type G T"
 apply (rule val.induct)
-apply    	auto
+apply           auto
 done
 
 (* unused *)
--- a/src/HOL/Decision_Procs/Approximation.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -175,16 +175,16 @@
     proof (cases "u < 0")
       case True hence "0 \<le> - real u" and "- real u \<le> - x" and "0 \<le> - x" and "-x \<le> - real l" using assms unfolding less_float_def by auto
       hence "real u ^ n \<le> x ^ n" and "x ^ n \<le> real l ^ n" using power_mono[of  "-x" "-real l" n] power_mono[of "-real u" "-x" n]
-	unfolding power_minus_even[OF `even n`] by auto
+        unfolding power_minus_even[OF `even n`] by auto
       moreover have u1: "u1 = l ^ n" and l1: "l1 = u ^ n" using assms unfolding float_power_bnds_def if_not_P[OF P] if_P[OF True] by auto
       ultimately show ?thesis using float_power by auto
     next
       case False
       have "\<bar>x\<bar> \<le> real (max (-l) u)"
       proof (cases "-l \<le> u")
-	case True thus ?thesis unfolding max_def if_P[OF True] using assms unfolding le_float_def by auto
+        case True thus ?thesis unfolding max_def if_P[OF True] using assms unfolding le_float_def by auto
       next
-	case False thus ?thesis unfolding max_def if_not_P[OF False] using assms unfolding le_float_def by auto
+        case False thus ?thesis unfolding max_def if_not_P[OF False] using assms unfolding le_float_def by auto
       qed
       hence x_abs: "\<bar>x\<bar> \<le> \<bar>real (max (-l) u)\<bar>" by auto
       have u1: "u1 = (max (-l) u) ^ n" and l1: "l1 = 0" using assms unfolding float_power_bnds_def if_not_P[OF P] if_not_P[OF False] by auto
@@ -258,7 +258,7 @@
     also have "\<dots> < 1 * pow2 (e + int (nat (bitlen m)))"
     proof (rule mult_strict_right_mono, auto)
       show "real m < 2^nat (bitlen m)" using bitlen_bounds[OF `0 < m`, THEN conjunct2]
-	unfolding real_of_int_less_iff[of m, symmetric] by auto
+        unfolding real_of_int_less_iff[of m, symmetric] by auto
     qed
     finally have "sqrt (real x) < sqrt (pow2 (e + bitlen m))" unfolding int_nat_bl by auto
     also have "\<dots> \<le> pow2 ((e + bitlen m) div 2 + 1)"
@@ -266,26 +266,26 @@
       let ?E = "e + bitlen m"
       have E_mod_pow: "pow2 (?E mod 2) < 4"
       proof (cases "?E mod 2 = 1")
-	case True thus ?thesis by auto
+        case True thus ?thesis by auto
       next
-	case False
-	have "0 \<le> ?E mod 2" by auto
-	have "?E mod 2 < 2" by auto
-	from this[THEN zless_imp_add1_zle]
-	have "?E mod 2 \<le> 0" using False by auto
-	from xt1(5)[OF `0 \<le> ?E mod 2` this]
-	show ?thesis by auto
+        case False
+        have "0 \<le> ?E mod 2" by auto
+        have "?E mod 2 < 2" by auto
+        from this[THEN zless_imp_add1_zle]
+        have "?E mod 2 \<le> 0" using False by auto
+        from xt1(5)[OF `0 \<le> ?E mod 2` this]
+        show ?thesis by auto
       qed
       hence "sqrt (pow2 (?E mod 2)) < sqrt (2 * 2)" by auto
       hence E_mod_pow: "sqrt (pow2 (?E mod 2)) < 2" unfolding real_sqrt_abs2 by auto
 
       have E_eq: "pow2 ?E = pow2 (?E div 2 + ?E div 2 + ?E mod 2)" by auto
       have "sqrt (pow2 ?E) = sqrt (pow2 (?E div 2) * pow2 (?E div 2) * pow2 (?E mod 2))"
-	unfolding E_eq unfolding pow2_add ..
+        unfolding E_eq unfolding pow2_add ..
       also have "\<dots> = pow2 (?E div 2) * sqrt (pow2 (?E mod 2))"
-	unfolding real_sqrt_mult[of _ "pow2 (?E mod 2)"] real_sqrt_abs2 by auto
+        unfolding real_sqrt_mult[of _ "pow2 (?E mod 2)"] real_sqrt_abs2 by auto
       also have "\<dots> < pow2 (?E div 2) * 2"
-	by (rule mult_strict_left_mono, auto intro: E_mod_pow)
+        by (rule mult_strict_left_mono, auto intro: E_mod_pow)
       also have "\<dots> = pow2 (?E div 2 + 1)" unfolding zadd_commute[of _ 1] pow2_add1 by auto
       finally show ?thesis by auto
     qed
@@ -338,7 +338,7 @@
                mult_pos_pos[OF order_less_trans[OF sqrt_gt0 sqrt_ub] sqrt_gt0]])
     also have "\<dots> = sqrt (real x)"
       unfolding inverse_eq_iff_eq[of _ "sqrt (real x)", symmetric]
-	        sqrt_divide_self_eq[OF `0 \<le> real x`, symmetric] by auto
+                sqrt_divide_self_eq[OF `0 \<le> real x`, symmetric] by auto
     finally have "real (lb_sqrt prec x) \<le> sqrt (real x)"
       unfolding lb_sqrt.simps if_P[OF `0 < x`] by auto }
   note lb = this
@@ -605,7 +605,7 @@
 
       have "real x \<le> sqrt (real (1 + x * x))" using real_sqrt_sum_squares_ge2[where x=1, unfolded numeral_2_eq_2] by auto
       also have "\<dots> \<le> real (ub_sqrt prec (1 + x * x))"
-	using bnds_sqrt'[of "1 + x * x"] by auto
+        using bnds_sqrt'[of "1 + x * x"] by auto
       finally have "real x \<le> real ?fR" by auto
       moreover have "real ?DIV \<le> real x / real ?fR" by (rule float_divl)
       ultimately have "real ?DIV \<le> 1" unfolding divide_le_eq_1_pos[OF `0 < real ?fR`, symmetric] by auto
@@ -613,9 +613,9 @@
       have "0 \<le> real ?DIV" using float_divl_lower_bound[OF `0 \<le> x` `0 < ?fR`] unfolding le_float_def by auto
 
       have "real (Float 1 1 * ?lb_horner ?DIV) \<le> 2 * arctan (real (float_divl prec x ?fR))" unfolding real_of_float_mult[of "Float 1 1"] Float_num
-	using arctan_0_1_bounds[OF `0 \<le> real ?DIV` `real ?DIV \<le> 1`] by auto
+        using arctan_0_1_bounds[OF `0 \<le> real ?DIV` `real ?DIV \<le> 1`] by auto
       also have "\<dots> \<le> 2 * arctan (real x / ?R)"
-	using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono)
+        using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono)
       also have "2 * arctan (real x / ?R) = arctan (real x)" using arctan_half[symmetric] unfolding numeral_2_eq_2 power_Suc2 power_0 real_mult_1 .
       finally show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_P[OF True] .
     next
@@ -628,26 +628,26 @@
 
       show ?thesis
       proof (cases "1 < ?invx")
-	case True
-	show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_not_P[OF False] if_P[OF True]
-	  using `0 \<le> arctan (real x)` by auto
+        case True
+        show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_not_P[OF False] if_P[OF True]
+          using `0 \<le> arctan (real x)` by auto
       next
-	case False
-	hence "real ?invx \<le> 1" unfolding less_float_def by auto
-	have "0 \<le> real ?invx" by (rule order_trans[OF _ float_divr], auto simp add: `0 \<le> real x`)
-
-	have "1 / real x \<noteq> 0" and "0 < 1 / real x" using `0 < real x` by auto
-
-	have "arctan (1 / real x) \<le> arctan (real ?invx)" unfolding real_of_float_1[symmetric] by (rule arctan_monotone', rule float_divr)
-	also have "\<dots> \<le> real (?ub_horner ?invx)" using arctan_0_1_bounds[OF `0 \<le> real ?invx` `real ?invx \<le> 1`] by auto
-	finally have "pi / 2 - real (?ub_horner ?invx) \<le> arctan (real x)"
-	  using `0 \<le> arctan (real x)` arctan_inverse[OF `1 / real x \<noteq> 0`]
-	  unfolding real_sgn_pos[OF `0 < 1 / real x`] le_diff_eq by auto
-	moreover
-	have "real (lb_pi prec * Float 1 -1) \<le> pi / 2" unfolding real_of_float_mult Float_num times_divide_eq_right real_mult_1 using pi_boundaries by auto
-	ultimately
-	show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_not_P[OF `\<not> x \<le> Float 1 1`] if_not_P[OF False]
-	  by auto
+        case False
+        hence "real ?invx \<le> 1" unfolding less_float_def by auto
+        have "0 \<le> real ?invx" by (rule order_trans[OF _ float_divr], auto simp add: `0 \<le> real x`)
+
+        have "1 / real x \<noteq> 0" and "0 < 1 / real x" using `0 < real x` by auto
+
+        have "arctan (1 / real x) \<le> arctan (real ?invx)" unfolding real_of_float_1[symmetric] by (rule arctan_monotone', rule float_divr)
+        also have "\<dots> \<le> real (?ub_horner ?invx)" using arctan_0_1_bounds[OF `0 \<le> real ?invx` `real ?invx \<le> 1`] by auto
+        finally have "pi / 2 - real (?ub_horner ?invx) \<le> arctan (real x)"
+          using `0 \<le> arctan (real x)` arctan_inverse[OF `1 / real x \<noteq> 0`]
+          unfolding real_sgn_pos[OF `0 < 1 / real x`] le_diff_eq by auto
+        moreover
+        have "real (lb_pi prec * Float 1 -1) \<le> pi / 2" unfolding real_of_float_mult Float_num times_divide_eq_right real_mult_1 using pi_boundaries by auto
+        ultimately
+        show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_not_P[OF `\<not> x \<le> Float 1 1`] if_not_P[OF False]
+          by auto
       qed
     qed
   qed
@@ -695,23 +695,23 @@
       case True
       show ?thesis
       proof (cases "?DIV > 1")
-	case True
-	have "pi / 2 \<le> real (ub_pi prec * Float 1 -1)" unfolding real_of_float_mult Float_num times_divide_eq_right real_mult_1 using pi_boundaries by auto
-	from order_less_le_trans[OF arctan_ubound this, THEN less_imp_le]
-	show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_P[OF `x \<le> Float 1 1`] if_P[OF True] .
+        case True
+        have "pi / 2 \<le> real (ub_pi prec * Float 1 -1)" unfolding real_of_float_mult Float_num times_divide_eq_right real_mult_1 using pi_boundaries by auto
+        from order_less_le_trans[OF arctan_ubound this, THEN less_imp_le]
+        show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_P[OF `x \<le> Float 1 1`] if_P[OF True] .
       next
-	case False
-	hence "real ?DIV \<le> 1" unfolding less_float_def by auto
-
-	have "0 \<le> real x / ?R" using `0 \<le> real x` `0 < ?R` unfolding real_0_le_divide_iff by auto
-	hence "0 \<le> real ?DIV" using monotone by (rule order_trans)
-
-	have "arctan (real x) = 2 * arctan (real x / ?R)" using arctan_half unfolding numeral_2_eq_2 power_Suc2 power_0 real_mult_1 .
-	also have "\<dots> \<le> 2 * arctan (real ?DIV)"
-	  using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono)
-	also have "\<dots> \<le> real (Float 1 1 * ?ub_horner ?DIV)" unfolding real_of_float_mult[of "Float 1 1"] Float_num
-	  using arctan_0_1_bounds[OF `0 \<le> real ?DIV` `real ?DIV \<le> 1`] by auto
-	finally show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_P[OF `x \<le> Float 1 1`] if_not_P[OF False] .
+        case False
+        hence "real ?DIV \<le> 1" unfolding less_float_def by auto
+
+        have "0 \<le> real x / ?R" using `0 \<le> real x` `0 < ?R` unfolding real_0_le_divide_iff by auto
+        hence "0 \<le> real ?DIV" using monotone by (rule order_trans)
+
+        have "arctan (real x) = 2 * arctan (real x / ?R)" using arctan_half unfolding numeral_2_eq_2 power_Suc2 power_0 real_mult_1 .
+        also have "\<dots> \<le> 2 * arctan (real ?DIV)"
+          using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono)
+        also have "\<dots> \<le> real (Float 1 1 * ?ub_horner ?DIV)" unfolding real_of_float_mult[of "Float 1 1"] Float_num
+          using arctan_0_1_bounds[OF `0 \<le> real ?DIV` `real ?DIV \<le> 1`] by auto
+        finally show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_P[OF `x \<le> Float 1 1`] if_not_P[OF False] .
       qed
     next
       case False
@@ -731,13 +731,13 @@
       have "real (?lb_horner ?invx) \<le> arctan (real ?invx)" using arctan_0_1_bounds[OF `0 \<le> real ?invx` `real ?invx \<le> 1`] by auto
       also have "\<dots> \<le> arctan (1 / real x)" unfolding real_of_float_1[symmetric] by (rule arctan_monotone', rule float_divl)
       finally have "arctan (real x) \<le> pi / 2 - real (?lb_horner ?invx)"
-	using `0 \<le> arctan (real x)` arctan_inverse[OF `1 / real x \<noteq> 0`]
-	unfolding real_sgn_pos[OF `0 < 1 / real x`] le_diff_eq by auto
+        using `0 \<le> arctan (real x)` arctan_inverse[OF `1 / real x \<noteq> 0`]
+        unfolding real_sgn_pos[OF `0 < 1 / real x`] le_diff_eq by auto
       moreover
       have "pi / 2 \<le> real (ub_pi prec * Float 1 -1)" unfolding real_of_float_mult Float_num times_divide_eq_right mult_1_right using pi_boundaries by auto
       ultimately
       show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_not_P[OF `\<not> x \<le> Float 1 1`] if_not_P[OF False]
-	by auto
+        by auto
     qed
   qed
 qed
@@ -848,12 +848,12 @@
     {
       assume "even n"
       have "real (lb_sin_cos_aux prec n 1 1 (x * x)) \<le> ?SUM"
-	unfolding morph_to_if_power[symmetric] using cos_aux by auto
+        unfolding morph_to_if_power[symmetric] using cos_aux by auto
       also have "\<dots> \<le> cos (real x)"
       proof -
-	from even[OF `even n`] `0 < ?fact` `0 < ?pow`
-	have "0 \<le> (?rest / ?fact) * ?pow" by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
-	thus ?thesis unfolding cos_eq by auto
+        from even[OF `even n`] `0 < ?fact` `0 < ?pow`
+        have "0 \<le> (?rest / ?fact) * ?pow" by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
+        thus ?thesis unfolding cos_eq by auto
       qed
       finally have "real (lb_sin_cos_aux prec n 1 1 (x * x)) \<le> cos (real x)" .
     } note lb = this
@@ -862,13 +862,13 @@
       assume "odd n"
       have "cos (real x) \<le> ?SUM"
       proof -
-	from `0 < ?fact` and `0 < ?pow` and odd[OF `odd n`]
-	have "0 \<le> (- ?rest) / ?fact * ?pow"
-	  by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
-	thus ?thesis unfolding cos_eq by auto
+        from `0 < ?fact` and `0 < ?pow` and odd[OF `odd n`]
+        have "0 \<le> (- ?rest) / ?fact * ?pow"
+          by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
+        thus ?thesis unfolding cos_eq by auto
       qed
       also have "\<dots> \<le> real (ub_sin_cos_aux prec n 1 1 (x * x))"
-	unfolding morph_to_if_power[symmetric] using cos_aux by auto
+        unfolding morph_to_if_power[symmetric] using cos_aux by auto
       finally have "cos (real x) \<le> real (ub_sin_cos_aux prec n 1 1 (x * x))" .
     } note ub = this and lb
   } note ub = this(1) and lb = this(2)
@@ -931,9 +931,9 @@
       have pow: "!!i. x ^ (2 * i + 1) = x * x ^ (2 * i)" by auto
       have "?SUM = (\<Sum> j = 0 ..< n. 0) + ?SUM" by auto
       also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. if even i then 0 else -1 ^ ((i - Suc 0) div 2) / (real (fact i)) * x ^ i)"
-	unfolding sum_split_even_odd ..
+        unfolding sum_split_even_odd ..
       also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. (if even i then 0 else -1 ^ ((i - Suc 0) div 2) / (real (fact i))) * x ^ i)"
-	by (rule setsum_cong2) auto
+        by (rule setsum_cong2) auto
       finally show ?thesis by assumption
     qed } note setsum_morph = this
 
@@ -958,13 +958,13 @@
       assume "even n"
       have "real (x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le>
             (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else (-1 ^ ((i - Suc 0) div 2))/(real (fact i))) * (real x) ^ i)"
-	using sin_aux[OF `0 \<le> real x`] unfolding setsum_morph[symmetric] by auto
+        using sin_aux[OF `0 \<le> real x`] unfolding setsum_morph[symmetric] by auto
       also have "\<dots> \<le> ?SUM" by auto
       also have "\<dots> \<le> sin (real x)"
       proof -
-	from even[OF `even n`] `0 < ?fact` `0 < ?pow`
-	have "0 \<le> (?rest / ?fact) * ?pow" by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
-	thus ?thesis unfolding sin_eq by auto
+        from even[OF `even n`] `0 < ?fact` `0 < ?pow`
+        have "0 \<le> (?rest / ?fact) * ?pow" by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
+        thus ?thesis unfolding sin_eq by auto
       qed
       finally have "real (x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le> sin (real x)" .
     } note lb = this
@@ -973,15 +973,15 @@
       assume "odd n"
       have "sin (real x) \<le> ?SUM"
       proof -
-	from `0 < ?fact` and `0 < ?pow` and odd[OF `odd n`]
-	have "0 \<le> (- ?rest) / ?fact * ?pow"
-	  by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
-	thus ?thesis unfolding sin_eq by auto
+        from `0 < ?fact` and `0 < ?pow` and odd[OF `odd n`]
+        have "0 \<le> (- ?rest) / ?fact * ?pow"
+          by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
+        thus ?thesis unfolding sin_eq by auto
       qed
       also have "\<dots> \<le> (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else (-1 ^ ((i - Suc 0) div 2))/(real (fact i))) * (real x) ^ i)"
-	 by auto
+         by auto
       also have "\<dots> \<le> real (x * ub_sin_cos_aux prec n 2 1 (x * x))"
-	using sin_aux[OF `0 \<le> real x`] unfolding setsum_morph[symmetric] by auto
+        using sin_aux[OF `0 \<le> real x`] unfolding setsum_morph[symmetric] by auto
       finally have "sin (real x) \<le> real (x * ub_sin_cos_aux prec n 2 1 (x * x))" .
     } note ub = this and lb
   } note ub = this(1) and lb = this(2)
@@ -1058,15 +1058,15 @@
 
       have "real (?lb_half y) \<le> cos (real x)"
       proof (cases "y < 0")
-	case True show ?thesis using cos_ge_minus_one unfolding if_P[OF True] by auto
+        case True show ?thesis using cos_ge_minus_one unfolding if_P[OF True] by auto
       next
-	case False
-	hence "0 \<le> real y" unfolding less_float_def by auto
-	from mult_mono[OF `real y \<le> cos ?x2` `real y \<le> cos ?x2` `0 \<le> cos ?x2` this]
-	have "real y * real y \<le> cos ?x2 * cos ?x2" .
-	hence "2 * real y * real y \<le> 2 * cos ?x2 * cos ?x2" by auto
-	hence "2 * real y * real y - 1 \<le> 2 * cos (real x / 2) * cos (real x / 2) - 1" unfolding Float_num real_of_float_mult by auto
-	thus ?thesis unfolding if_not_P[OF False] x_half Float_num real_of_float_mult real_of_float_sub by auto
+        case False
+        hence "0 \<le> real y" unfolding less_float_def by auto
+        from mult_mono[OF `real y \<le> cos ?x2` `real y \<le> cos ?x2` `0 \<le> cos ?x2` this]
+        have "real y * real y \<le> cos ?x2 * cos ?x2" .
+        hence "2 * real y * real y \<le> 2 * cos ?x2 * cos ?x2" by auto
+        hence "2 * real y * real y - 1 \<le> 2 * cos (real x / 2) * cos (real x / 2) - 1" unfolding Float_num real_of_float_mult by auto
+        thus ?thesis unfolding if_not_P[OF False] x_half Float_num real_of_float_mult real_of_float_sub by auto
       qed
     } note lb_half = this
 
@@ -1077,12 +1077,12 @@
 
       have "cos (real x) \<le> real (?ub_half y)"
       proof -
-	have "0 \<le> real y" using `0 \<le> cos ?x2` ub by (rule order_trans)
-	from mult_mono[OF ub ub this `0 \<le> cos ?x2`]
-	have "cos ?x2 * cos ?x2 \<le> real y * real y" .
-	hence "2 * cos ?x2 * cos ?x2 \<le> 2 * real y * real y" by auto
-	hence "2 * cos (real x / 2) * cos (real x / 2) - 1 \<le> 2 * real y * real y - 1" unfolding Float_num real_of_float_mult by auto
-	thus ?thesis unfolding x_half real_of_float_mult Float_num real_of_float_sub by auto
+        have "0 \<le> real y" using `0 \<le> cos ?x2` ub by (rule order_trans)
+        from mult_mono[OF ub ub this `0 \<le> cos ?x2`]
+        have "cos ?x2 * cos ?x2 \<le> real y * real y" .
+        hence "2 * cos ?x2 * cos ?x2 \<le> 2 * real y * real y" by auto
+        hence "2 * cos (real x / 2) * cos (real x / 2) - 1 \<le> 2 * real y * real y - 1" unfolding Float_num real_of_float_mult by auto
+        thus ?thesis unfolding x_half real_of_float_mult Float_num real_of_float_sub by auto
       qed
     } note ub_half = this
 
@@ -1100,13 +1100,13 @@
 
       have "real (?lb x) \<le> ?cos x"
       proof -
-	from lb_half[OF lb `-pi \<le> real x` `real x \<le> pi`]
-	show ?thesis unfolding lb_cos_def[where x=x] Let_def using `\<not> x < 0` `\<not> x < Float 1 -1` `x < 1` by auto
+        from lb_half[OF lb `-pi \<le> real x` `real x \<le> pi`]
+        show ?thesis unfolding lb_cos_def[where x=x] Let_def using `\<not> x < 0` `\<not> x < Float 1 -1` `x < 1` by auto
       qed
       moreover have "?cos x \<le> real (?ub x)"
       proof -
-	from ub_half[OF ub `-pi \<le> real x` `real x \<le> pi`]
-	show ?thesis unfolding ub_cos_def[where x=x] Let_def using `\<not> x < 0` `\<not> x < Float 1 -1` `x < 1` by auto
+        from ub_half[OF ub `-pi \<le> real x` `real x \<le> pi`]
+        show ?thesis unfolding ub_cos_def[where x=x] Let_def using `\<not> x < 0` `\<not> x < Float 1 -1` `x < 1` by auto
       qed
       ultimately show ?thesis by auto
     next
@@ -1119,15 +1119,15 @@
 
       have "real (?lb x) \<le> ?cos x"
       proof -
-	have "-pi \<le> real ?x2" and "real ?x2 \<le> pi" unfolding real_of_float_mult Float_num using pi_ge_two `0 \<le> real x` `real x \<le> pi` by auto
-	from lb_half[OF lb_half[OF lb this] `-pi \<le> real x` `real x \<le> pi`, unfolded eq_4]
-	show ?thesis unfolding lb_cos_def[where x=x] if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x < Float 1 -1`] if_not_P[OF `\<not> x < 1`] Let_def .
+        have "-pi \<le> real ?x2" and "real ?x2 \<le> pi" unfolding real_of_float_mult Float_num using pi_ge_two `0 \<le> real x` `real x \<le> pi` by auto
+        from lb_half[OF lb_half[OF lb this] `-pi \<le> real x` `real x \<le> pi`, unfolded eq_4]
+        show ?thesis unfolding lb_cos_def[where x=x] if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x < Float 1 -1`] if_not_P[OF `\<not> x < 1`] Let_def .
       qed
       moreover have "?cos x \<le> real (?ub x)"
       proof -
-	have "-pi \<le> real ?x2" and "real ?x2 \<le> pi" unfolding real_of_float_mult Float_num using pi_ge_two `0 \<le> real x` `real x \<le> pi` by auto
-	from ub_half[OF ub_half[OF ub this] `-pi \<le> real x` `real x \<le> pi`, unfolded eq_4]
-	show ?thesis unfolding ub_cos_def[where x=x] if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x < Float 1 -1`] if_not_P[OF `\<not> x < 1`] Let_def .
+        have "-pi \<le> real ?x2" and "real ?x2 \<le> pi" unfolding real_of_float_mult Float_num using pi_ge_two `0 \<le> real x` `real x \<le> pi` by auto
+        from ub_half[OF ub_half[OF ub this] `-pi \<le> real x` `real x \<le> pi`, unfolded eq_4]
+        show ?thesis unfolding ub_cos_def[where x=x] if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x < Float 1 -1`] if_not_P[OF `\<not> x < 1`] Let_def .
       qed
       ultimately show ?thesis by auto
     qed
@@ -1227,7 +1227,7 @@
     also have "\<dots> \<le> cos (x + real (-k) * 2 * pi)"
       using cos_monotone_minus_pi_0'[OF pi_lx lx x_le_0]
       by (simp only: real_of_float_minus real_of_int_minus
-	cos_minus real_diff_def mult_minus_left)
+        cos_minus real_diff_def mult_minus_left)
     finally have "real (lb_cos prec (- ?lx)) \<le> cos x"
       unfolding cos_periodic_int . }
   note negative_lx = this
@@ -1240,7 +1240,7 @@
     have "cos (x + real (-k) * 2 * pi) \<le> cos (real ?lx)"
       using cos_monotone_0_pi'[OF lx_0 lx pi_x]
       by (simp only: real_of_float_minus real_of_int_minus
-	cos_minus real_diff_def mult_minus_left)
+        cos_minus real_diff_def mult_minus_left)
     also have "\<dots> \<le> real (ub_cos prec ?lx)"
       using lb_cos[OF lx_0 pi_lx] by simp
     finally have "cos x \<le> real (ub_cos prec ?lx)"
@@ -1255,7 +1255,7 @@
     have "cos (x + real (-k) * 2 * pi) \<le> cos (real (- ?ux))"
       using cos_monotone_minus_pi_0'[OF pi_x ux ux_0]
       by (simp only: real_of_float_minus real_of_int_minus
-	  cos_minus real_diff_def mult_minus_left)
+          cos_minus real_diff_def mult_minus_left)
     also have "\<dots> \<le> real (ub_cos prec (- ?ux))"
       using lb_cos_minus[OF pi_ux ux_0, of prec] by simp
     finally have "cos x \<le> real (ub_cos prec (- ?ux))"
@@ -1272,7 +1272,7 @@
     also have "\<dots> \<le> cos (x + real (-k) * 2 * pi)"
       using cos_monotone_0_pi'[OF x_ge_0 ux pi_ux]
       by (simp only: real_of_float_minus real_of_int_minus
-	cos_minus real_diff_def mult_minus_left)
+        cos_minus real_diff_def mult_minus_left)
     finally have "real (lb_cos prec ?ux) \<le> cos x"
       unfolding cos_periodic_int . }
   note positive_ux = this
@@ -1332,26 +1332,26 @@
       hence "x - real k * 2 * pi - 2 * pi \<le> 0" using ux by simp
 
       have ux_0: "real (?ux - 2 * ?lpi) \<le> 0"
-	using Cond by (auto simp add: le_float_def)
+        using Cond by (auto simp add: le_float_def)
 
       from 2 and Cond have "\<not> ?ux \<le> ?lpi" by auto
       hence "- ?lpi \<le> ?ux - 2 * ?lpi" by (auto simp add: le_float_def)
       hence pi_ux: "- pi \<le> real (?ux - 2 * ?lpi)"
-	using lpi[THEN le_imp_neg_le] by (auto simp add: le_float_def)
+        using lpi[THEN le_imp_neg_le] by (auto simp add: le_float_def)
 
       have x_le_ux: "x - real k * 2 * pi - 2 * pi \<le> real (?ux - 2 * ?lpi)"
-	using ux lpi by auto
+        using ux lpi by auto
 
       have "cos x = cos (x + real (-k) * 2 * pi + real (-1 :: int) * 2 * pi)"
-	unfolding cos_periodic_int ..
+        unfolding cos_periodic_int ..
       also have "\<dots> \<le> cos (real (?ux - 2 * ?lpi))"
-	using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0]
-	by (simp only: real_of_float_minus real_of_int_minus real_of_one
-	    number_of_Min real_diff_def mult_minus_left real_mult_1)
+        using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0]
+        by (simp only: real_of_float_minus real_of_int_minus real_of_one
+            number_of_Min real_diff_def mult_minus_left real_mult_1)
       also have "\<dots> = cos (real (- (?ux - 2 * ?lpi)))"
-	unfolding real_of_float_minus cos_minus ..
+        unfolding real_of_float_minus cos_minus ..
       also have "\<dots> \<le> real (ub_cos prec (- (?ux - 2 * ?lpi)))"
-	using lb_cos_minus[OF pi_ux ux_0] by simp
+        using lb_cos_minus[OF pi_ux ux_0] by simp
       finally show ?thesis unfolding u by (simp add: real_of_float_max)
     qed
     thus ?thesis unfolding l by auto
@@ -1376,24 +1376,24 @@
       hence "0 \<le> x - real k * 2 * pi + 2 * pi" using lx by simp
 
       have lx_0: "0 \<le> real (?lx + 2 * ?lpi)"
-	using Cond lpi by (auto simp add: le_float_def)
+        using Cond lpi by (auto simp add: le_float_def)
 
       from 1 and Cond have "\<not> -?lpi \<le> ?lx" by auto
       hence "?lx + 2 * ?lpi \<le> ?lpi" by (auto simp add: le_float_def)
       hence pi_lx: "real (?lx + 2 * ?lpi) \<le> pi"
-	using lpi[THEN le_imp_neg_le] by (auto simp add: le_float_def)
+        using lpi[THEN le_imp_neg_le] by (auto simp add: le_float_def)
 
       have lx_le_x: "real (?lx + 2 * ?lpi) \<le> x - real k * 2 * pi + 2 * pi"
-	using lx lpi by auto
+        using lx lpi by auto
 
       have "cos x = cos (x + real (-k) * 2 * pi + real (1 :: int) * 2 * pi)"
-	unfolding cos_periodic_int ..
+        unfolding cos_periodic_int ..
       also have "\<dots> \<le> cos (real (?lx + 2 * ?lpi))"
-	using cos_monotone_0_pi'[OF lx_0 lx_le_x pi_x]
-	by (simp only: real_of_float_minus real_of_int_minus real_of_one
-	  number_of_Min real_diff_def mult_minus_left real_mult_1)
+        using cos_monotone_0_pi'[OF lx_0 lx_le_x pi_x]
+        by (simp only: real_of_float_minus real_of_int_minus real_of_one
+          number_of_Min real_diff_def mult_minus_left real_mult_1)
       also have "\<dots> \<le> real (ub_cos prec (?lx + 2 * ?lpi))"
-	using lb_cos[OF lx_0 pi_lx] by simp
+        using lb_cos[OF lx_0 pi_lx] by simp
       finally show ?thesis unfolding u by (simp add: real_of_float_max)
     qed
     thus ?thesis unfolding l by auto
@@ -1427,11 +1427,11 @@
     also have "\<dots> \<le> exp (real x)"
     proof -
       obtain t where "\<bar>t\<bar> \<le> \<bar>real x\<bar>" and "exp (real x) = (\<Sum>m = 0..<get_even n. (real x) ^ m / real (fact m)) + exp t / real (fact (get_even n)) * (real x) ^ (get_even n)"
-	using Maclaurin_exp_le by blast
+        using Maclaurin_exp_le by blast
       moreover have "0 \<le> exp t / real (fact (get_even n)) * (real x) ^ (get_even n)"
-	by (auto intro!: mult_nonneg_nonneg divide_nonneg_pos simp add: get_even zero_le_even_power exp_gt_zero)
+        by (auto intro!: mult_nonneg_nonneg divide_nonneg_pos simp add: get_even zero_le_even_power exp_gt_zero)
       ultimately show ?thesis
-	using get_odd exp_gt_zero by (auto intro!: pordered_cancel_semiring_class.mult_nonneg_nonneg)
+        using get_odd exp_gt_zero by (auto intro!: pordered_cancel_semiring_class.mult_nonneg_nonneg)
     qed
     finally have "real (lb_exp_horner prec (get_even n) 1 1 x) \<le> exp (real x)" .
   } moreover
@@ -1548,14 +1548,14 @@
     have "exp (real x) \<le> real (ub_exp prec x)"
     proof -
       have div_less_zero: "real (float_divr prec x (- floor_fl x)) \<le> 0"
-	using float_divr_nonpos_pos_upper_bound[OF `x \<le> 0` `0 < - floor_fl x`] unfolding le_float_def real_of_float_0 .
+        using float_divr_nonpos_pos_upper_bound[OF `x \<le> 0` `0 < - floor_fl x`] unfolding le_float_def real_of_float_0 .
 
       have "exp (real x) = exp (real ?num * (real x / real ?num))" using `real ?num \<noteq> 0` by auto
       also have "\<dots> = exp (real x / real ?num) ^ ?num" unfolding exp_real_of_nat_mult ..
       also have "\<dots> \<le> exp (real (float_divr prec x (- floor_fl x))) ^ ?num" unfolding num_eq
-	by (rule power_mono, rule exp_le_cancel_iff[THEN iffD2], rule float_divr) auto
+        by (rule power_mono, rule exp_le_cancel_iff[THEN iffD2], rule float_divr) auto
       also have "\<dots> \<le> real ((?ub_exp_horner (float_divr prec x (- floor_fl x))) ^ ?num)" unfolding float_power
-	by (rule power_mono, rule bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct2], auto)
+        by (rule power_mono, rule bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct2], auto)
       finally show ?thesis unfolding ub_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def .
     qed
     moreover
@@ -1566,32 +1566,32 @@
 
       show ?thesis
       proof (cases "?horner \<le> 0")
-	case False hence "0 \<le> real ?horner" unfolding le_float_def by auto
-
-	have div_less_zero: "real (float_divl prec x (- floor_fl x)) \<le> 0"
-	  using `real (floor_fl x) < 0` `real x \<le> 0` by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg)
-
-	have "real ((?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num) \<le>
+        case False hence "0 \<le> real ?horner" unfolding le_float_def by auto
+
+        have div_less_zero: "real (float_divl prec x (- floor_fl x)) \<le> 0"
+          using `real (floor_fl x) < 0` `real x \<le> 0` by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg)
+
+        have "real ((?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num) \<le>
           exp (real (float_divl prec x (- floor_fl x))) ^ ?num" unfolding float_power
-	  using `0 \<le> real ?horner`[unfolded Float_floor[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono)
-	also have "\<dots> \<le> exp (real x / real ?num) ^ ?num" unfolding num_eq
-	  using float_divl by (auto intro!: power_mono simp del: real_of_float_minus)
-	also have "\<dots> = exp (real ?num * (real x / real ?num))" unfolding exp_real_of_nat_mult ..
-	also have "\<dots> = exp (real x)" using `real ?num \<noteq> 0` by auto
-	finally show ?thesis
-	  unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def if_not_P[OF False] by auto
+          using `0 \<le> real ?horner`[unfolded Float_floor[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono)
+        also have "\<dots> \<le> exp (real x / real ?num) ^ ?num" unfolding num_eq
+          using float_divl by (auto intro!: power_mono simp del: real_of_float_minus)
+        also have "\<dots> = exp (real ?num * (real x / real ?num))" unfolding exp_real_of_nat_mult ..
+        also have "\<dots> = exp (real x)" using `real ?num \<noteq> 0` by auto
+        finally show ?thesis
+          unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def if_not_P[OF False] by auto
       next
-	case True
-	have "real (floor_fl x) \<noteq> 0" and "real (floor_fl x) \<le> 0" using `real (floor_fl x) < 0` by auto
-	from divide_right_mono_neg[OF floor_fl[of x] `real (floor_fl x) \<le> 0`, unfolded divide_self[OF `real (floor_fl x) \<noteq> 0`]]
-	have "- 1 \<le> real x / real (- floor_fl x)" unfolding real_of_float_minus by auto
-	from order_trans[OF exp_m1_ge_quarter this[unfolded exp_le_cancel_iff[where x="- 1", symmetric]]]
-	have "real (Float 1 -2) \<le> exp (real x / real (- floor_fl x))" unfolding Float_num .
-	hence "real (Float 1 -2) ^ ?num \<le> exp (real x / real (- floor_fl x)) ^ ?num"
-	  by (auto intro!: power_mono simp add: Float_num)
-	also have "\<dots> = exp (real x)" unfolding num_eq exp_real_of_nat_mult[symmetric] using `real (floor_fl x) \<noteq> 0` by auto
-	finally show ?thesis
-	  unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def if_P[OF True] float_power .
+        case True
+        have "real (floor_fl x) \<noteq> 0" and "real (floor_fl x) \<le> 0" using `real (floor_fl x) < 0` by auto
+        from divide_right_mono_neg[OF floor_fl[of x] `real (floor_fl x) \<le> 0`, unfolded divide_self[OF `real (floor_fl x) \<noteq> 0`]]
+        have "- 1 \<le> real x / real (- floor_fl x)" unfolding real_of_float_minus by auto
+        from order_trans[OF exp_m1_ge_quarter this[unfolded exp_le_cancel_iff[where x="- 1", symmetric]]]
+        have "real (Float 1 -2) \<le> exp (real x / real (- floor_fl x))" unfolding Float_num .
+        hence "real (Float 1 -2) ^ ?num \<le> exp (real x / real (- floor_fl x)) ^ ?num"
+          by (auto intro!: power_mono simp add: Float_num)
+        also have "\<dots> = exp (real x)" unfolding num_eq exp_real_of_nat_mult[symmetric] using `real (floor_fl x) \<noteq> 0` by auto
+        finally show ?thesis
+          unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def if_P[OF True] float_power .
       qed
     qed
     ultimately show ?thesis by auto
@@ -1614,8 +1614,8 @@
 
       have "real (float_divl prec 1 (ub_exp prec (-x))) \<le> 1 / real (ub_exp prec (-x))" using float_divl[where x=1] by auto
       also have "\<dots> \<le> exp (real x)"
-	using ub_exp[unfolded inverse_le_iff_le[OF order_less_le_trans[OF exp_gt_zero ub_exp] exp_gt_zero, symmetric]]
-	unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide by auto
+        using ub_exp[unfolded inverse_le_iff_le[OF order_less_le_trans[OF exp_gt_zero ub_exp] exp_gt_zero, symmetric]]
+        unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide by auto
       finally show ?thesis unfolding lb_exp.simps if_P[OF True] .
     qed
     moreover
@@ -1627,9 +1627,9 @@
       have lb_exp: "real (lb_exp prec (-x)) \<le> exp (- real x)" unfolding atLeastAtMost_iff real_of_float_minus by auto
 
       have "exp (real x) \<le> real (1 :: float) / real (lb_exp prec (-x))"
-	using lb_exp[unfolded inverse_le_iff_le[OF exp_gt_zero lb_exp_pos[OF `\<not> 0 < -x`, unfolded less_float_def real_of_float_0],
-	                                        symmetric]]
-	unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide real_of_float_1 by auto
+        using lb_exp[unfolded inverse_le_iff_le[OF exp_gt_zero lb_exp_pos[OF `\<not> 0 < -x`, unfolded less_float_def real_of_float_0],
+                                                symmetric]]
+        unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide real_of_float_1 by auto
       also have "\<dots> \<le> real (float_divr prec 1 (lb_exp prec (-x)))" using float_divr .
       finally show ?thesis unfolding ub_exp.simps if_P[OF True] .
     qed
@@ -1683,7 +1683,7 @@
     proof (rule mult_mono)
       show "0 \<le> x ^ Suc (Suc n)" by (auto intro!: mult_nonneg_nonneg simp add: `0 \<le> x`)
       have "x ^ Suc (Suc n) \<le> x ^ Suc n * 1" unfolding power_Suc2 real_mult_assoc[symmetric]
-	by (rule mult_left_mono, fact less_imp_le[OF `x < 1`], auto intro!: mult_nonneg_nonneg simp add: `0 \<le> x`)
+        by (rule mult_left_mono, fact less_imp_le[OF `x < 1`], auto intro!: mult_nonneg_nonneg simp add: `0 \<le> x`)
       thus "x ^ Suc (Suc n) \<le> x ^ Suc n" by auto
     qed auto }
   from summable_Leibniz'(2,4)[OF `?a ----> 0` `\<And>n. 0 \<le> ?a n`, OF `\<And>n. ?a (Suc n) \<le> ?a n`, unfolded ln_eq]
@@ -1857,58 +1857,58 @@
     let "?lb_horner x" = "x * lb_ln_horner prec (get_even prec) 1 x"
 
     { have up: "real (rapprox_rat prec 2 3) \<le> 1"
-	by (rule rapprox_rat_le1) simp_all
+        by (rule rapprox_rat_le1) simp_all
       have low: "2 / 3 \<le> real (rapprox_rat prec 2 3)"
-	by (rule order_trans[OF _ rapprox_rat]) simp
+        by (rule order_trans[OF _ rapprox_rat]) simp
       from mult_less_le_imp_less[OF * low] *
       have pos: "0 < real (x * rapprox_rat prec 2 3 - 1)" by auto
 
       have "ln (real x * 2/3)
-	\<le> ln (real (x * rapprox_rat prec 2 3 - 1) + 1)"
+        \<le> ln (real (x * rapprox_rat prec 2 3 - 1) + 1)"
       proof (rule ln_le_cancel_iff[symmetric, THEN iffD1])
-	show "real x * 2 / 3 \<le> real (x * rapprox_rat prec 2 3 - 1) + 1"
-	  using * low by auto
-	show "0 < real x * 2 / 3" using * by simp
-	show "0 < real (x * rapprox_rat prec 2 3 - 1) + 1" using pos by auto
+        show "real x * 2 / 3 \<le> real (x * rapprox_rat prec 2 3 - 1) + 1"
+          using * low by auto
+        show "0 < real x * 2 / 3" using * by simp
+        show "0 < real (x * rapprox_rat prec 2 3 - 1) + 1" using pos by auto
       qed
       also have "\<dots> \<le> real (?ub_horner (x * rapprox_rat prec 2 3 - 1))"
       proof (rule ln_float_bounds(2))
-	from mult_less_le_imp_less[OF `real x < 2` up] low *
-	show "real (x * rapprox_rat prec 2 3 - 1) < 1" by auto
-	show "0 \<le> real (x * rapprox_rat prec 2 3 - 1)" using pos by auto
+        from mult_less_le_imp_less[OF `real x < 2` up] low *
+        show "real (x * rapprox_rat prec 2 3 - 1) < 1" by auto
+        show "0 \<le> real (x * rapprox_rat prec 2 3 - 1)" using pos by auto
       qed
       finally have "ln (real x)
-	\<le> real (?ub_horner (Float 1 -1))
-	  + real (?ub_horner (x * rapprox_rat prec 2 3 - 1))"
-	using ln_float_bounds(2)[of "Float 1 -1" prec prec] add by auto }
+        \<le> real (?ub_horner (Float 1 -1))
+          + real (?ub_horner (x * rapprox_rat prec 2 3 - 1))"
+        using ln_float_bounds(2)[of "Float 1 -1" prec prec] add by auto }
     moreover
     { let ?max = "max (x * lapprox_rat prec 2 3 - 1) 0"
 
       have up: "real (lapprox_rat prec 2 3) \<le> 2/3"
-	by (rule order_trans[OF lapprox_rat], simp)
+        by (rule order_trans[OF lapprox_rat], simp)
 
       have low: "0 \<le> real (lapprox_rat prec 2 3)"
-	using lapprox_rat_bottom[of 2 3 prec] by simp
+        using lapprox_rat_bottom[of 2 3 prec] by simp
 
       have "real (?lb_horner ?max)
-	\<le> ln (real ?max + 1)"
+        \<le> ln (real ?max + 1)"
       proof (rule ln_float_bounds(1))
-	from mult_less_le_imp_less[OF `real x < 2` up] * low
-	show "real ?max < 1" by (cases "real (lapprox_rat prec 2 3) = 0",
-	  auto simp add: real_of_float_max)
-	show "0 \<le> real ?max" by (auto simp add: real_of_float_max)
+        from mult_less_le_imp_less[OF `real x < 2` up] * low
+        show "real ?max < 1" by (cases "real (lapprox_rat prec 2 3) = 0",
+          auto simp add: real_of_float_max)
+        show "0 \<le> real ?max" by (auto simp add: real_of_float_max)
       qed
       also have "\<dots> \<le> ln (real x * 2/3)"
       proof (rule ln_le_cancel_iff[symmetric, THEN iffD1])
-	show "0 < real ?max + 1" by (auto simp add: real_of_float_max)
-	show "0 < real x * 2/3" using * by auto
-	show "real ?max + 1 \<le> real x * 2/3" using * up
-	  by (cases "0 < real x * real (lapprox_posrat prec 2 3) - 1",
-	      auto simp add: real_of_float_max min_max.sup_absorb1)
+        show "0 < real ?max + 1" by (auto simp add: real_of_float_max)
+        show "0 < real x * 2/3" using * by auto
+        show "real ?max + 1 \<le> real x * 2/3" using * up
+          by (cases "0 < real x * real (lapprox_posrat prec 2 3) - 1",
+              auto simp add: real_of_float_max min_max.sup_absorb1)
       qed
       finally have "real (?lb_horner (Float 1 -1)) + real (?lb_horner ?max)
-	\<le> ln (real x)"
-	using ln_float_bounds(1)[of "Float 1 -1" prec prec] add by auto }
+        \<le> ln (real x)"
+        using ln_float_bounds(1)[of "Float 1 -1" prec prec] add by auto }
     ultimately
     show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps Let_def
       using `\<not> x \<le> 0` `\<not> x < 1` True False by auto
@@ -1928,12 +1928,12 @@
 
     {
       have "real (lb_ln2 prec * ?s) \<le> ln 2 * real (e + (bitlen m - 1))" (is "?lb2 \<le> _")
-	unfolding real_of_float_mult real_of_float_ge0_exp[OF order_refl] nat_0 power_0 mult_1_right
-	using lb_ln2[of prec]
+        unfolding real_of_float_mult real_of_float_ge0_exp[OF order_refl] nat_0 power_0 mult_1_right
+        using lb_ln2[of prec]
       proof (rule mult_right_mono)
-	have "1 \<le> Float m e" using `1 \<le> x` Float unfolding le_float_def by auto
-	from float_gt1_scale[OF this]
-	show "0 \<le> real (e + (bitlen m - 1))" by auto
+        have "1 \<le> Float m e" using `1 \<le> x` Float unfolding le_float_def by auto
+        from float_gt1_scale[OF this]
+        show "0 \<le> real (e + (bitlen m - 1))" by auto
       qed
       moreover
       from bitlen_div[OF `0 < m`, unfolded normalized_float[OF `m \<noteq> 0`, symmetric]]
@@ -1941,7 +1941,7 @@
       from ln_float_bounds(1)[OF this]
       have "real ((?x - 1) * lb_ln_horner prec (get_even prec) 1 (?x - 1)) \<le> ln (real ?x)" (is "?lb_horner \<le> _") by auto
       ultimately have "?lb2 + ?lb_horner \<le> ln (real x)"
-	unfolding Float ln_shifted_float[OF `0 < m`, of e] by auto
+        unfolding Float ln_shifted_float[OF `0 < m`, of e] by auto
     }
     moreover
     {
@@ -1951,15 +1951,15 @@
       have "ln (real ?x) \<le> real ((?x - 1) * ub_ln_horner prec (get_odd prec) 1 (?x - 1))" (is "_ \<le> ?ub_horner") by auto
       moreover
       have "ln 2 * real (e + (bitlen m - 1)) \<le> real (ub_ln2 prec * ?s)" (is "_ \<le> ?ub2")
-	unfolding real_of_float_mult real_of_float_ge0_exp[OF order_refl] nat_0 power_0 mult_1_right
-	using ub_ln2[of prec]
+        unfolding real_of_float_mult real_of_float_ge0_exp[OF order_refl] nat_0 power_0 mult_1_right
+        using ub_ln2[of prec]
       proof (rule mult_right_mono)
-	have "1 \<le> Float m e" using `1 \<le> x` Float unfolding le_float_def by auto
-	from float_gt1_scale[OF this]
-	show "0 \<le> real (e + (bitlen m - 1))" by auto
+        have "1 \<le> Float m e" using `1 \<le> x` Float unfolding le_float_def by auto
+        from float_gt1_scale[OF this]
+        show "0 \<le> real (e + (bitlen m - 1))" by auto
       qed
       ultimately have "ln (real x) \<le> ?ub2 + ?ub_horner"
-	unfolding Float ln_shifted_float[OF `0 < m`, of e] by auto
+        unfolding Float ln_shifted_float[OF `0 < m`, of e] by auto
     }
     ultimately show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps
       unfolding if_not_P[OF `\<not> x \<le> 0`] if_not_P[OF `\<not> x < 1`] if_not_P[OF False] if_not_P[OF `\<not> x \<le> Float 3 -1`] Let_def
@@ -2413,7 +2413,7 @@
       unfolding less_float_def using l1_le_u1 l1 by auto
     show ?thesis
       unfolding inverse_le_iff_le[OF `0 < real u1` `0 < interpret_floatarith a xs`]
-	inverse_le_iff_le[OF `0 < interpret_floatarith a xs` `0 < real l1`]
+        inverse_le_iff_le[OF `0 < interpret_floatarith a xs` `0 < real l1`]
       using l1 u1 by auto
   next
     case False hence "u1 < 0" using either by blast
@@ -2421,7 +2421,7 @@
       unfolding less_float_def using l1_le_u1 u1 by auto
     show ?thesis
       unfolding inverse_le_iff_le_neg[OF `real u1 < 0` `interpret_floatarith a xs < 0`]
-	inverse_le_iff_le_neg[OF `interpret_floatarith a xs < 0` `real l1 < 0`]
+        inverse_le_iff_le_neg[OF `interpret_floatarith a xs < 0` `real l1 < 0`]
       using l1 u1 by auto
   qed
 
@@ -2767,11 +2767,11 @@
     proof (cases "i = j")
       case True
       thus ?thesis using `vs ! i = Some (l, u)` Some and bnd `i < length ?xs`
-	by auto
+        by auto
     next
       case False
       thus ?thesis using `bounded_by xs vs`[THEN bounded_byE, OF `j < length vs`] Some
-	by auto
+        by auto
     qed
   qed auto }
   thus ?thesis unfolding bounded_by_def by auto
@@ -2878,7 +2878,7 @@
       note bounded_by_update_var[OF `bounded_by xs vs` bnd_x this]
       from approx[OF this ap]
       have "(interpret_floatarith f (xs[x := t])) \<in> {real l .. real u}"
-	by (auto simp add: algebra_simps)
+        by (auto simp add: algebra_simps)
     } thus ?thesis by (auto intro!: exI[of _ 0])
   next
     case True
@@ -2904,7 +2904,7 @@
 
     { fix f :: "'a \<Rightarrow> 'a" fix n :: nat fix x :: 'a
       have "(f ^^ Suc n) x = (f ^^ n) (f x)"
-	by (induct n, auto) }
+        by (induct n, auto) }
     note funpow_Suc = this[symmetric]
     from Suc.hyps[OF ate, unfolded this]
     obtain n
@@ -2918,14 +2918,14 @@
       assume "m < Suc n" and bnd_z: "z \<in> { real lx .. real ux }"
       have "DERIV (?f m) z :> ?f (Suc m) z"
       proof (cases m)
-	case 0
-	with DERIV_floatarith[OF `x < length xs` isDERIV_approx'[OF `bounded_by xs vs` bnd_x bnd_z True]]
-	show ?thesis by simp
+        case 0
+        with DERIV_floatarith[OF `x < length xs` isDERIV_approx'[OF `bounded_by xs vs` bnd_x bnd_z True]]
+        show ?thesis by simp
       next
-	case (Suc m')
-	hence "m' < n" using `m < Suc n` by auto
-	from DERIV_hyp[OF this bnd_z]
-	show ?thesis using Suc by simp
+        case (Suc m')
+        hence "m' < n" using `m < Suc n` by auto
+        from DERIV_hyp[OF this bnd_z]
+        show ?thesis using Suc by simp
       qed } note DERIV = this
 
     have "\<And> k i. k < i \<Longrightarrow> {k ..< i} = insert k {Suc k ..< i}" by auto
@@ -2937,20 +2937,20 @@
 
     { fix t assume t: "t \<in> {real lx .. real ux}"
       hence "bounded_by [xs!x] [vs!x]"
-	using `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`]
-	by (cases "vs!x", auto simp add: bounded_by_def)
+        using `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`]
+        by (cases "vs!x", auto simp add: bounded_by_def)
 
       with hyp[THEN bspec, OF t] f_c
       have "bounded_by [?f 0 (real c), ?X (Suc k) f n t, xs!x] [Some (l1, u1), Some (l2, u2), vs!x]"
-	by (auto intro!: bounded_by_Cons)
+        by (auto intro!: bounded_by_Cons)
       from approx[OF this final, unfolded atLeastAtMost_iff[symmetric]]
       have "?X (Suc k) f n t * (xs!x - real c) * inverse (real k) + ?f 0 (real c) \<in> {real l .. real u}"
-	by (auto simp add: algebra_simps)
+        by (auto simp add: algebra_simps)
       also have "?X (Suc k) f n t * (xs!x - real c) * inverse (real k) + ?f 0 (real c) =
                (\<Sum> i = 0..<Suc n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) * ?f i (real c) * (xs!x - real c)^i) +
                inverse (real (\<Prod> j \<in> {k..<k+Suc n}. j)) * ?f (Suc n) t * (xs!x - real c)^Suc n" (is "_ = ?T")
-	unfolding funpow_Suc C_def[symmetric] setsum_move0 setprod_head_Suc
-	by (auto simp add: algebra_simps setsum_right_distrib[symmetric])
+        unfolding funpow_Suc C_def[symmetric] setsum_move0 setprod_head_Suc
+        by (auto simp add: algebra_simps setsum_right_distrib[symmetric])
       finally have "?T \<in> {real l .. real u}" . }
     thus ?thesis using DERIV by blast
   qed
@@ -2999,24 +2999,24 @@
     proof (cases "xs ! x = real c")
       case True
       from True[symmetric] hyp[OF bnd_xs] Suc show ?thesis
-	unfolding F_def Suc setsum_head_upt_Suc[OF zero_less_Suc] setsum_shift_bounds_Suc_ivl by auto
+        unfolding F_def Suc setsum_head_upt_Suc[OF zero_less_Suc] setsum_shift_bounds_Suc_ivl by auto
     next
       case False
 
       have "real lx \<le> real c" "real c \<le> real ux" "real lx \<le> xs!x" "xs!x \<le> real ux"
-	using Suc bnd_c `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`] bnd_x by auto
+        using Suc bnd_c `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`] bnd_x by auto
       from Taylor.taylor[OF zero_less_Suc, of F, OF F0 DERIV[unfolded Suc] this False]
       obtain t where t_bnd: "if xs ! x < real c then xs ! x < t \<and> t < real c else real c < t \<and> t < xs ! x"
-	and fl_eq: "interpret_floatarith f (xs[x := xs ! x]) =
-	   (\<Sum>m = 0..<Suc n'. F m (real c) / real (fact m) * (xs ! x - real c) ^ m) +
+        and fl_eq: "interpret_floatarith f (xs[x := xs ! x]) =
+           (\<Sum>m = 0..<Suc n'. F m (real c) / real (fact m) * (xs ! x - real c) ^ m) +
            F (Suc n') t / real (fact (Suc n')) * (xs ! x - real c) ^ Suc n'"
-	by blast
+        by blast
 
       from t_bnd bnd_xs bnd_c have *: "t \<in> {real lx .. real ux}"
-	by (cases "xs ! x < real c", auto)
+        by (cases "xs ! x < real c", auto)
 
       have "interpret_floatarith f (xs[x := xs ! x]) = ?taylor t"
-	unfolding fl_eq Suc by (auto simp add: algebra_simps divide_inverse)
+        unfolding fl_eq Suc by (auto simp add: algebra_simps divide_inverse)
       also have "\<dots> \<in> {real l .. real u}" using * by (rule hyp)
       finally show ?thesis by simp
     qed
@@ -3154,22 +3154,22 @@
       case (Less lf rt)
       with Bound a b assms
       have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 < l)"
-	unfolding approx_tse_form_def by auto
+        unfolding approx_tse_form_def by auto
       from approx_tse_form'_less[OF this bnd]
       show ?thesis using Less by auto
     next
       case (LessEqual lf rt)
       with Bound a b assms
       have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)"
-	unfolding approx_tse_form_def by auto
+        unfolding approx_tse_form_def by auto
       from approx_tse_form'_le[OF this bnd]
       show ?thesis using LessEqual by auto
     next
       case (AtLeastAtMost x lf rt)
       with Bound a b assms
       have "approx_tse_form' prec t (Add rt (Minus x)) s l u' (\<lambda> l u. 0 \<le> l)"
-	and "approx_tse_form' prec t (Add x (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)"
-	unfolding approx_tse_form_def lazy_conj by auto
+        and "approx_tse_form' prec t (Add x (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)"
+        unfolding approx_tse_form_def lazy_conj by auto
       from approx_tse_form'_le[OF this(1) bnd] approx_tse_form'_le[OF this(2) bnd]
       show ?thesis using AtLeastAtMost by auto
     next
--- a/src/HOL/Decision_Procs/Cooper.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -63,8 +63,8 @@
   "fmsize (NDvd i t) = 2"
   "fmsize p = 1"
   (* several lemmas about fmsize *)
-lemma fmsize_pos: "fmsize p > 0"	
-by (induct p rule: fmsize.induct) simp_all
+lemma fmsize_pos: "fmsize p > 0"
+  by (induct p rule: fmsize.induct) simp_all
 
   (* Semantics of formulae (fm) *)
 consts Ifm ::"bool list \<Rightarrow> int list \<Rightarrow> fm \<Rightarrow> bool"
@@ -622,14 +622,14 @@
   {assume i1: "abs i = 1"
       from one_dvd[of "Inum bs a"] uminus_dvd_conv[where d="1" and t="Inum bs a"]
       have ?case using i1 apply (cases "i=0", simp_all add: Let_def) 
-	by (cases "i > 0", simp_all)}
+        by (cases "i > 0", simp_all)}
   moreover   
   {assume inz: "i\<noteq>0" and cond: "abs i \<noteq> 1"
     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
-	by (cases "abs i = 1", auto) }
+        by (cases "abs i = 1", auto) }
     moreover {assume "\<not> (\<exists> v. ?sa = C v)" 
       hence "simpfm (Dvd i a) = Dvd i ?sa" using inz cond 
-	by (cases ?sa, auto simp add: Let_def)
+        by (cases ?sa, auto simp add: Let_def)
       hence ?case using sa by simp}
     ultimately have ?case by blast}
   ultimately show ?case by blast
@@ -646,10 +646,10 @@
   moreover   
   {assume inz: "i\<noteq>0" and cond: "abs i \<noteq> 1"
     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
-	by (cases "abs i = 1", auto) }
+        by (cases "abs i = 1", auto) }
     moreover {assume "\<not> (\<exists> v. ?sa = C v)" 
       hence "simpfm (NDvd i a) = NDvd i ?sa" using inz cond 
-	by (cases ?sa, auto simp add: Let_def)
+        by (cases ?sa, auto simp add: Let_def)
       hence ?case using sa by simp}
     ultimately have ?case by blast}
   ultimately show ?case by blast
@@ -1291,24 +1291,24 @@
     show ?case 
     proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib, rule iffI)
       assume 
-	"i dvd c * x - c*(k*d) + Inum (x # bs) e"
+        "i dvd c * x - c*(k*d) + Inum (x # bs) e"
       (is "?ri dvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri dvd ?rt")
       hence "\<exists> (l::int). ?rt = i * l" by (simp add: dvd_def)
       hence "\<exists> (l::int). c*x+ ?I x e = i*l+c*(k * i*di)" 
-	by (simp add: algebra_simps di_def)
+        by (simp add: algebra_simps di_def)
       hence "\<exists> (l::int). c*x+ ?I x e = i*(l + c*k*di)"
-	by (simp add: algebra_simps)
+        by (simp add: algebra_simps)
       hence "\<exists> (l::int). c*x+ ?I x e = i*l" by blast
       thus "i dvd c*x + Inum (x # bs) e" by (simp add: dvd_def) 
     next
       assume 
-	"i dvd c*x + Inum (x # bs) e" (is "?ri dvd ?rc*?rx+?e")
+        "i dvd c*x + Inum (x # bs) e" (is "?ri dvd ?rc*?rx+?e")
       hence "\<exists> (l::int). c*x+?e = i*l" by (simp add: dvd_def)
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*d)" by simp
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*i*di)" by (simp add: di_def)
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: algebra_simps)
       hence "\<exists> (l::int). c*x - c * (k*d) +?e = i*l"
-	by blast
+        by blast
       thus "i dvd c*x - c*(k*d) + Inum (x # bs) e" by (simp add: dvd_def)
     qed
 next
@@ -1318,24 +1318,24 @@
     show ?case 
     proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib, rule iffI)
       assume 
-	"i dvd c * x - c*(k*d) + Inum (x # bs) e"
+        "i dvd c * x - c*(k*d) + Inum (x # bs) e"
       (is "?ri dvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri dvd ?rt")
       hence "\<exists> (l::int). ?rt = i * l" by (simp add: dvd_def)
       hence "\<exists> (l::int). c*x+ ?I x e = i*l+c*(k * i*di)" 
-	by (simp add: algebra_simps di_def)
+        by (simp add: algebra_simps di_def)
       hence "\<exists> (l::int). c*x+ ?I x e = i*(l + c*k*di)"
-	by (simp add: algebra_simps)
+        by (simp add: algebra_simps)
       hence "\<exists> (l::int). c*x+ ?I x e = i*l" by blast
       thus "i dvd c*x + Inum (x # bs) e" by (simp add: dvd_def) 
     next
       assume 
-	"i dvd c*x + Inum (x # bs) e" (is "?ri dvd ?rc*?rx+?e")
+        "i dvd c*x + Inum (x # bs) e" (is "?ri dvd ?rc*?rx+?e")
       hence "\<exists> (l::int). c*x+?e = i*l" by (simp add: dvd_def)
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*d)" by simp
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*i*di)" by (simp add: di_def)
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: algebra_simps)
       hence "\<exists> (l::int). c*x - c * (k*d) +?e = i*l"
-	by blast
+        by blast
       thus "i dvd c*x - c*(k*d) + Inum (x # bs) e" by (simp add: dvd_def)
     qed
 qed (auto simp add: gr0_conv_Suc numbound0_I[where bs="bs" and b="x - k*d" and b'="x"])
@@ -1611,7 +1611,7 @@
       hence "x + ?e \<ge> 1 \<and> x + ?e \<le> d"  by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + ?e" by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. x = (- ?e + j)" 
-	by (simp add: algebra_simps)
+        by (simp add: algebra_simps)
       with nob have ?case by auto}
     ultimately show ?case by blast
 next
@@ -1620,7 +1620,7 @@
     let ?e = "Inum (x # bs) e"
     {assume "(x-d) +?e \<ge> 0" hence ?case using  c1 
       numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"]
-	by simp}
+        by simp}
     moreover
     {assume H: "\<not> (x-d) + ?e \<ge> 0" 
       let ?v="Sub (C -1) e"
@@ -1640,7 +1640,7 @@
     have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp
     from p have "x= - ?e" by (simp add: c1) with prems(11) show ?case using dp
       by simp (erule ballE[where x="1"],
-	simp_all add:algebra_simps numbound0_I[OF bn,where b="x"and b'="a"and bs="bs"])
+        simp_all add:algebra_simps numbound0_I[OF bn,where b="x"and b'="a"and bs="bs"])
 next
   case (4 c e)hence p: "Ifm bbs (x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" by simp+
     let ?e = "Inum (x # bs) e"
@@ -1652,7 +1652,7 @@
     {assume H: "x - d + Inum (((x -d)) # bs) e = 0"
       hence "x = - Inum (((x -d)) # bs) e + d" by simp
       hence "x = - Inum (a # bs) e + d"
-	by (simp add: numbound0_I[OF bn,where b="x - d"and b'="a"and bs="bs"])
+        by (simp add: numbound0_I[OF bn,where b="x - d"and b'="a"and bs="bs"])
        with prems(11) have ?case using dp by simp}
   ultimately show ?case by blast
 next 
--- a/src/HOL/Decision_Procs/Ferrack.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -601,7 +601,7 @@
 proof-
   let ?g = "numgcd t"
   have "?g \<ge> 0"  by (simp add: numgcd_pos)
-  hence	"?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto
+  hence "?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto
   moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} 
   moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} 
   moreover { assume g1:"?g > 1"
@@ -750,20 +750,20 @@
       hence "?g'= 1 \<or> ?g' > 1" by arith
       moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)}
       moreover {assume g'1:"?g'>1"
-	from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
-	let ?tt = "reducecoeffh ?t' ?g'"
-	let ?t = "Inum bs ?tt"
-	have gpdg: "?g' dvd ?g" by simp
-	have gpdd: "?g' dvd n" by simp 
-	have gpdgp: "?g' dvd ?g'" by simp
-	from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
-	have th2:"real ?g' * ?t = Inum bs ?t'" by simp
-	from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def)
-	also have "\<dots> = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp
-	also have "\<dots> = (Inum bs ?t' / real n)"
-	  using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp
-	finally have "?lhs = Inum bs t / real n" by (simp add: simpnum_ci)
-	then have ?thesis using prems by (simp add: simp_num_pair_def)}
+        from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
+        let ?tt = "reducecoeffh ?t' ?g'"
+        let ?t = "Inum bs ?tt"
+        have gpdg: "?g' dvd ?g" by simp
+        have gpdd: "?g' dvd n" by simp 
+        have gpdgp: "?g' dvd ?g'" by simp
+        from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
+        have th2:"real ?g' * ?t = Inum bs ?t'" by simp
+        from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def)
+        also have "\<dots> = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp
+        also have "\<dots> = (Inum bs ?t' / real n)"
+          using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp
+        finally have "?lhs = Inum bs t / real n" by (simp add: simpnum_ci)
+        then have ?thesis using prems by (simp add: simp_num_pair_def)}
       ultimately have ?thesis by blast}
     ultimately have ?thesis by blast} 
   ultimately show ?thesis by blast
@@ -785,16 +785,16 @@
       hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
       hence "?g'= 1 \<or> ?g' > 1" by arith
       moreover {assume "?g'=1" hence ?thesis using prems 
-	  by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)}
+          by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)}
       moreover {assume g'1:"?g'>1"
-	have gpdg: "?g' dvd ?g" by simp
-	have gpdd: "?g' dvd n" by simp 
-	have gpdgp: "?g' dvd ?g'" by simp
-	from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
-	from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]]
-	have "n div ?g' >0" by simp
-	hence ?thesis using prems 
-	  by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0 simpnum_numbound0)}
+        have gpdg: "?g' dvd ?g" by simp
+        have gpdd: "?g' dvd n" by simp 
+        have gpdgp: "?g' dvd ?g'" by simp
+        from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
+        from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]]
+        have "n div ?g' >0" by simp
+        hence ?thesis using prems 
+          by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0 simpnum_numbound0)}
       ultimately have ?thesis by blast}
     ultimately have ?thesis by blast} 
   ultimately show ?thesis by blast
@@ -1510,7 +1510,7 @@
     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
     moreover {assume y: "y < (-?N x e)/ real c"
       hence "y * real c < - ?N x e"
-	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
+        by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
       hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y > (- ?N x e) / real c" 
@@ -1529,7 +1529,7 @@
     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
     moreover {assume y: "y < (-?N x e)/ real c"
       hence "y * real c < - ?N x e"
-	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
+        by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
       hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y > (- ?N x e) / real c" 
@@ -1548,7 +1548,7 @@
     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
     moreover {assume y: "y > (-?N x e)/ real c"
       hence "y * real c > - ?N x e"
-	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
+        by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
       hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y < (- ?N x e) / real c" 
@@ -1567,7 +1567,7 @@
     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
     moreover {assume y: "y > (-?N x e)/ real c"
       hence "y * real c > - ?N x e"
-	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
+        by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
       hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y < (- ?N x e) / real c" 
@@ -1694,8 +1694,8 @@
     moreover{
       assume "\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p"
       then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M" 
-	and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
-	by blast
+        and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
+        by blast
       from t1M have "\<exists> (t1u,t1n) \<in> ?U. t1 = ?N a t1u / real t1n" by auto
       then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \<in> ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast
       from t2M have "\<exists> (t2u,t2n) \<in> ?U. t2 = ?N a t2u / real t2n" by auto
@@ -1751,13 +1751,13 @@
     let ?N = "\<lambda> t. Inum (x#bs) t"
     {fix t n s m assume "(t,n)\<in> set (uset p)" and "(s,m) \<in> set (uset p)"
       with uset_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0"
-	by auto
+        by auto
       let ?st = "Add (Mul m t) (Mul n s)"
       from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
-	by (simp add: mult_commute)
+        by (simp add: mult_commute)
       from tnb snb have st_nb: "numbound0 ?st" by simp
       have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-	using mnp mp np by (simp add: algebra_simps add_divide_distrib)
+        using mnp mp np by (simp add: algebra_simps add_divide_distrib)
       from usubst_I[OF lp mnp st_nb, where x="x" and bs="bs"] 
       have "?I x (usubst p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])}
     with rinf_uset[OF lp nmi npi px] have "?F" by blast hence "?D" by blast}
@@ -1928,7 +1928,7 @@
     { fix t n assume tnY: "(t,n) \<in> set ?Y" 
       hence "(t,n) \<in> set ?SS" by simp
       hence "\<exists> (t',n') \<in> set ?S. simp_num_pair (t',n') = (t,n)"
-	by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all)
+        by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all)
       then obtain t' n' where tn'S: "(t',n') \<in> set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast
       from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto
       from simp_num_pair_l[OF tnb np tns]
--- a/src/HOL/Decision_Procs/MIR.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -100,7 +100,7 @@
   then show ?thesis by simp
 qed
 
-  (* The Divisibility relation between reals *)	
+(* The Divisibility relation between reals *)
 definition
   rdvd:: "real \<Rightarrow> real \<Rightarrow> bool" (infixl "rdvd" 50)
 where
@@ -268,7 +268,7 @@
 | "fmsize (NDvd i t) = 2"
 | "fmsize p = 1"
   (* several lemmas about fmsize *)
-lemma fmsize_pos: "fmsize p > 0"	
+lemma fmsize_pos: "fmsize p > 0"
 by (induct p rule: fmsize.induct) simp_all
 
   (* Semantics of formulae (fm) *)
@@ -807,7 +807,7 @@
 proof-
   let ?g = "numgcd t"
   have "?g \<ge> 0"  by (simp add: numgcd_pos)
-  hence	"?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto
+  hence "?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto
   moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} 
   moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} 
   moreover { assume g1:"?g > 1"
@@ -1064,20 +1064,20 @@
       hence "?g'= 1 \<or> ?g' > 1" by arith
       moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
       moreover {assume g'1:"?g'>1"
-	from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
-	let ?tt = "reducecoeffh ?t' ?g'"
-	let ?t = "Inum bs ?tt"
-	have gpdg: "?g' dvd ?g" by simp
-	have gpdd: "?g' dvd n" by simp
-	have gpdgp: "?g' dvd ?g'" by simp
-	from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
-	have th2:"real ?g' * ?t = Inum bs ?t'" by simp
-	from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def)
-	also have "\<dots> = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp
-	also have "\<dots> = (Inum bs ?t' / real n)"
-	  using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp
-	finally have "?lhs = Inum bs t / real n" by simp
-	then have ?thesis using prems by (simp add: simp_num_pair_def)}
+        from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
+        let ?tt = "reducecoeffh ?t' ?g'"
+        let ?t = "Inum bs ?tt"
+        have gpdg: "?g' dvd ?g" by simp
+        have gpdd: "?g' dvd n" by simp
+        have gpdgp: "?g' dvd ?g'" by simp
+        from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
+        have th2:"real ?g' * ?t = Inum bs ?t'" by simp
+        from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def)
+        also have "\<dots> = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp
+        also have "\<dots> = (Inum bs ?t' / real n)"
+          using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp
+        finally have "?lhs = Inum bs t / real n" by simp
+        then have ?thesis using prems by (simp add: simp_num_pair_def)}
       ultimately have ?thesis by blast}
     ultimately have ?thesis by blast} 
   ultimately show ?thesis by blast
@@ -1099,16 +1099,16 @@
       hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
       hence "?g'= 1 \<or> ?g' > 1" by arith
       moreover {assume "?g'=1" hence ?thesis using prems 
-	  by (auto simp add: Let_def simp_num_pair_def)}
+          by (auto simp add: Let_def simp_num_pair_def)}
       moreover {assume g'1:"?g'>1"
-	have gpdg: "?g' dvd ?g" by simp
-	have gpdd: "?g' dvd n" by simp
-	have gpdgp: "?g' dvd ?g'" by simp
-	from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
-	from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]]
-	have "n div ?g' >0" by simp
-	hence ?thesis using prems 
-	  by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)}
+        have gpdg: "?g' dvd ?g" by simp
+        have gpdd: "?g' dvd n" by simp
+        have gpdgp: "?g' dvd ?g'" by simp
+        from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
+        from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]]
+        have "n div ?g' >0" by simp
+        hence ?thesis using prems 
+          by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)}
       ultimately have ?thesis by blast}
     ultimately have ?thesis by blast} 
   ultimately show ?thesis by blast
@@ -1246,10 +1246,10 @@
       from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
       have th2:"real ?g' * ?t = Inum bs t" by simp
       from prems have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)"
-	by (simp add: simpdvd_def Let_def)
+        by (simp add: simpdvd_def Let_def)
       also have "\<dots> = (real d rdvd (Inum bs t))"
-	using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified zdiv_self[OF gp0]] 
-	  th2[symmetric] by simp
+        using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified zdiv_self[OF gp0]] 
+          th2[symmetric] by simp
       finally have ?thesis by simp  }
     ultimately have ?thesis by blast
   }
@@ -1393,13 +1393,13 @@
       have ?case using i1 ai by simp }
     moreover {assume i1: "i = - 1" 
       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
-	rdvd_abs1[where d="- 1" and t="Inum bs a"]
+        rdvd_abs1[where d="- 1" and t="Inum bs a"]
       have ?case using i1 ai by simp }
     ultimately have ?case by blast}
   moreover   
   {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
-	by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
+        by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
     moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" 
       hence th: "simpfm (Dvd i a) = Dvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond by (cases ?sa, auto simp add: Let_def split_def)
       from simpnum_nz have nz:"nozerocoeff ?sa" by simp
@@ -1418,16 +1418,16 @@
       have ?case using i1 ai by simp }
     moreover {assume i1: "i = - 1" 
       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
-	rdvd_abs1[where d="- 1" and t="Inum bs a"]
+        rdvd_abs1[where d="- 1" and t="Inum bs a"]
       have ?case using i1 ai by simp }
     ultimately have ?case by blast}
   moreover   
   {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
-	by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
+        by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
     moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" 
       hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond 
-	by (cases ?sa, auto simp add: Let_def split_def)
+        by (cases ?sa, auto simp add: Let_def split_def)
       from simpnum_nz have nz:"nozerocoeff ?sa" by simp
       from simpdvd [OF nz inz] th have ?case using sa by simp}
     ultimately have ?case by blast}
@@ -1990,7 +1990,7 @@
       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
     also have "\<dots> = (?I (?l (Dvd j a)))" using cp cnz jnz  
       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
-	del: real_of_int_mult) (auto simp add: add_ac)
+        del: real_of_int_mult) (auto simp add: add_ac)
     finally have ?case using l jnz  by simp }
   moreover
   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
@@ -2005,7 +2005,7 @@
     also have "\<dots> = (?I (?l (Dvd j a)))" using cn cnz jnz
       using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
-	del: real_of_int_mult) (auto simp add: add_ac)
+        del: real_of_int_mult) (auto simp add: add_ac)
     finally have ?case using l jnz by blast }
   ultimately show ?case by blast
 next
@@ -2036,7 +2036,7 @@
       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
     also have "\<dots> = (?I (?l (NDvd j a)))" using cp cnz jnz  
       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
-	del: real_of_int_mult) (auto simp add: add_ac)
+        del: real_of_int_mult) (auto simp add: add_ac)
     finally have ?case using l jnz  by simp }
   moreover
   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))" 
@@ -2051,7 +2051,7 @@
     also have "\<dots> = (?I (?l (NDvd j a)))" using cn cnz jnz
       using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
-	del: real_of_int_mult) (auto simp add: add_ac)
+        del: real_of_int_mult) (auto simp add: add_ac)
     finally have ?case using l jnz by blast }
   ultimately show ?case by blast
 qed auto
@@ -2277,24 +2277,24 @@
     show ?case 
     proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI)
       assume 
-	"real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
+        "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
       (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
       hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
-	by (simp add: algebra_simps di_def)
+        by (simp add: algebra_simps di_def)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
-	by (simp add: algebra_simps)
+        by (simp add: algebra_simps)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
       thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
     next
       assume 
-	"real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
+        "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
       hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
-	by blast
+        by blast
       thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
     qed
 next
@@ -2304,24 +2304,24 @@
     show ?case 
     proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI)
       assume 
-	"real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
+        "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
       (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
       hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
-	by (simp add: algebra_simps di_def)
+        by (simp add: algebra_simps di_def)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
-	by (simp add: algebra_simps)
+        by (simp add: algebra_simps)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
       thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
     next
       assume 
-	"real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
+        "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
       hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
-	by blast
+        by blast
       thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
     qed
 qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="real(x - k*d)" and b'="real x"] simp del: real_of_int_mult real_of_int_diff)
@@ -2706,7 +2706,7 @@
       by (simp add: isint_iff)
     {assume "real (x-d) +?e > 0" hence ?case using c1 
       numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
-	by (simp del: real_of_int_minus)}
+        by (simp del: real_of_int_minus)}
     moreover
     {assume H: "\<not> real (x-d) + ?e > 0" 
       let ?v="Neg e"
@@ -2715,13 +2715,13 @@
       have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e + real j)" by auto 
       from H p have "real x + ?e > 0 \<and> real x + ?e \<le> real d" by (simp add: c1)
       hence "real (x + floor ?e) > real (0::int) \<and> real (x + floor ?e) \<le> real d"
-	using ie by simp
+        using ie by simp
       hence "x + floor ?e \<ge> 1 \<and> x + floor ?e \<le> d"  by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e" by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. real x = real (- floor ?e + j)" 
-	by (simp only: real_of_int_inject) (simp add: algebra_simps)
+        by (simp only: real_of_int_inject) (simp add: algebra_simps)
       hence "\<exists> (j::int) \<in> {1 .. d}. real x = - ?e + real j" 
-	by (simp add: ie[simplified isint_iff])
+        by (simp add: ie[simplified isint_iff])
       with nob have ?case by auto}
     ultimately show ?case by blast
 next
@@ -2732,7 +2732,7 @@
       by (simp add: isint_iff)
     {assume "real (x-d) +?e \<ge> 0" hence ?case using  c1 
       numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
-	by (simp del: real_of_int_minus)}
+        by (simp del: real_of_int_minus)}
     moreover
     {assume H: "\<not> real (x-d) + ?e \<ge> 0" 
       let ?v="Sub (C -1) e"
@@ -2741,14 +2741,14 @@
       have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e - 1 + real j)" by auto 
       from H p have "real x + ?e \<ge> 0 \<and> real x + ?e < real d" by (simp add: c1)
       hence "real (x + floor ?e) \<ge> real (0::int) \<and> real (x + floor ?e) < real d"
-	using ie by simp
+        using ie by simp
       hence "x + floor ?e +1 \<ge> 1 \<and> x + floor ?e + 1 \<le> d"  by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e + 1" by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. x= - floor ?e - 1 + j" by (simp add: algebra_simps)
       hence "\<exists> (j::int) \<in> {1 .. d}. real x= real (- floor ?e - 1 + j)" 
-	by (simp only: real_of_int_inject)
+        by (simp only: real_of_int_inject)
       hence "\<exists> (j::int) \<in> {1 .. d}. real x= - ?e - 1 + real j" 
-	by (simp add: ie[simplified isint_iff])
+        by (simp add: ie[simplified isint_iff])
       with nob have ?case by simp }
     ultimately show ?case by blast
 next
@@ -2759,7 +2759,7 @@
     have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp
     from p have "real x= - ?e" by (simp add: c1) with prems(11) show ?case using dp
       by simp (erule ballE[where x="1"],
-	simp_all add:algebra_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"])
+        simp_all add:algebra_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"])
 next
   case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" 
     and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
@@ -2772,7 +2772,7 @@
     {assume H: "real x - real d + Inum ((real (x -d)) # bs) e = 0"
       hence "real x = - Inum ((real (x -d)) # bs) e + real d" by simp
       hence "real x = - Inum (a # bs) e + real d"
-	by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"])
+        by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"])
        with prems(11) have ?case using dp by simp}
   ultimately show ?case by blast
 next 
@@ -2974,19 +2974,19 @@
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
-	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
     moreover 
     {assume "\<not> k dvd c"
       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have "?I (real x) (?s (Eq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k = 0)"
-	using real_of_int_div[OF knz kdt]
-	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
+        using real_of_int_div[OF knz kdt]
+          numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Eq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
-	by (simp add: ti)
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
+        by (simp add: ti)
       finally have ?case . }
     ultimately show ?case by blast 
 next
@@ -2996,19 +2996,19 @@
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
-	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
     moreover 
     {assume "\<not> k dvd c"
       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have "?I (real x) (?s (NEq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<noteq> 0)"
-	using real_of_int_div[OF knz kdt]
-	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
+        using real_of_int_div[OF knz kdt]
+          numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (NEq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
-	by (simp add: ti)
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
+        by (simp add: ti)
       finally have ?case . }
     ultimately show ?case by blast 
 next
@@ -3018,19 +3018,19 @@
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
-	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
     moreover 
     {assume "\<not> k dvd c"
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have "?I (real x) (?s (Lt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k < 0)"
-	using real_of_int_div[OF knz kdt]
-	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
+        using real_of_int_div[OF knz kdt]
+          numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Lt (CN 0 c e)))" using pos_less_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
-	by (simp add: ti)
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
+        by (simp add: ti)
       finally have ?case . }
     ultimately show ?case by blast 
 next
@@ -3040,19 +3040,19 @@
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
-	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
     moreover 
     {assume "\<not> k dvd c"
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have "?I (real x) (?s (Le (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<le> 0)"
-	using real_of_int_div[OF knz kdt]
-	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
+        using real_of_int_div[OF knz kdt]
+          numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Le (CN 0 c e)))" using pos_le_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
-	by (simp add: ti)
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
+        by (simp add: ti)
       finally have ?case . }
     ultimately show ?case by blast 
 next
@@ -3062,19 +3062,19 @@
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
-	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
     moreover 
     {assume "\<not> k dvd c"
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have "?I (real x) (?s (Gt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k > 0)"
-	using real_of_int_div[OF knz kdt]
-	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
+        using real_of_int_div[OF knz kdt]
+          numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Gt (CN 0 c e)))" using pos_divide_less_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
-	by (simp add: ti)
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
+        by (simp add: ti)
       finally have ?case . }
     ultimately show ?case by blast 
 next
@@ -3084,19 +3084,19 @@
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
-	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
     moreover 
     {assume "\<not> k dvd c"
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have "?I (real x) (?s (Ge (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<ge> 0)"
-	using real_of_int_div[OF knz kdt]
-	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
+        using real_of_int_div[OF knz kdt]
+          numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Ge (CN 0 c e)))" using pos_divide_le_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
-	by (simp add: ti)
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
+        by (simp add: ti)
       finally have ?case . }
     ultimately show ?case by blast 
 next
@@ -3105,19 +3105,19 @@
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
-	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
     moreover 
     {assume "\<not> k dvd c"
       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have "?I (real x) (?s (Dvd i (CN 0 c e))) = (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k)"
-	using real_of_int_div[OF knz kdt]
-	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
+        using real_of_int_div[OF knz kdt]
+          numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Dvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
-	by (simp add: ti)
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
+        by (simp add: ti)
       finally have ?case . }
     ultimately show ?case by blast 
 next
@@ -3126,19 +3126,19 @@
       from kpos have knz: "k\<noteq>0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
-	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
     moreover 
     {assume "\<not> k dvd c"
       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
       from prems have "?I (real x) (?s (NDvd i (CN 0 c e))) = (\<not> (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k))"
-	using real_of_int_div[OF knz kdt]
-	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
+        using real_of_int_div[OF knz kdt]
+          numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (NDvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
-	by (simp add: ti)
+          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
+        by (simp add: ti)
       finally have ?case . }
     ultimately show ?case by blast 
 qed (simp_all add: nth_pos2 bound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"] numbound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"])
@@ -3330,7 +3330,7 @@
   moreover
   {assume "real (c*i) + ?N i e > real (c*d)" hence ?case
       by (simp add: algebra_simps 
-	numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
+        numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
   moreover 
   {assume H:"real (c*i) + ?N i e \<le> real (c*d)"
     with ei[simplified isint_iff] have "real (c*i + ?fe) \<le> real (c*d)" by simp
@@ -3353,7 +3353,7 @@
   moreover
   {assume "real (c*i) + ?N i e \<ge> real (c*d)" hence ?case
       by (simp add: algebra_simps 
-	numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
+        numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
   moreover 
   {assume H:"real (c*i) + ?N i e < real (c*d)"
     with ei[simplified isint_iff] have "real (c*i + ?fe) < real (c*d)" by simp
@@ -3364,7 +3364,7 @@
     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
       by (simp only: algebra_simps diff_def[symmetric])
         hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
-	  by (simp only: add_ac diff_def)
+          by (simp only: add_ac diff_def)
     with nob  have ?case by blast }
   ultimately show ?case by blast
 next
@@ -3629,60 +3629,60 @@
                 n = 0 \<and> s = Add (Floor ba) (C j) \<and> ac \<le> j \<and> j \<le> 0))"
       moreover 
       {fix s'
-	assume "(p, 0, s') \<in> ?SS a" and "n = 0" and "s = Floor s'"
-	hence ?ths using prems by auto}
+        assume "(p, 0, s') \<in> ?SS a" and "n = 0" and "s = Floor s'"
+        hence ?ths using prems by auto}
       moreover
-      {	fix p' n' s' j
-	assume pns: "(p', n', s') \<in> ?SS a" 
-	  and np: "0 < n'" 
-	  and p_def: "p = ?p (p',n',s') j" 
-	  and n0: "n = 0" 
-	  and s_def: "s = (Add (Floor s') (C j))" 
-	  and jp: "0 \<le> j" and jn: "j \<le> n'"
-	from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
+      { fix p' n' s' j
+        assume pns: "(p', n', s') \<in> ?SS a" 
+          and np: "0 < n'" 
+          and p_def: "p = ?p (p',n',s') j" 
+          and n0: "n = 0" 
+          and s_def: "s = (Add (Floor s') (C j))" 
+          and jp: "0 \<le> j" and jn: "j \<le> n'"
+        from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
           Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
           numbound0 s' \<and> isrlfm p'" by blast
-	hence nb: "numbound0 s'" by simp
-	from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numsub_nb)
-	let ?nxs = "CN 0 n' s'"
-	let ?l = "floor (?N s') + j"
-	from H 
-	have "?I (?p (p',n',s') j) \<longrightarrow> 
-	  (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
-	  by (simp add: fp_def np algebra_simps numsub numadd numfloor)
-	also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
-	  using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
-	moreover
-	have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp
-	ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
-	  by blast
-	with s_def n0 p_def nb nf have ?ths by auto}
+        hence nb: "numbound0 s'" by simp
+        from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numsub_nb)
+        let ?nxs = "CN 0 n' s'"
+        let ?l = "floor (?N s') + j"
+        from H 
+        have "?I (?p (p',n',s') j) \<longrightarrow> 
+          (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
+          by (simp add: fp_def np algebra_simps numsub numadd numfloor)
+        also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
+          using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
+        moreover
+        have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp
+        ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
+          by blast
+        with s_def n0 p_def nb nf have ?ths by auto}
       moreover
       {fix p' n' s' j
-	assume pns: "(p', n', s') \<in> ?SS a" 
-	  and np: "n' < 0" 
-	  and p_def: "p = ?p (p',n',s') j" 
-	  and n0: "n = 0" 
-	  and s_def: "s = (Add (Floor s') (C j))" 
-	  and jp: "n' \<le> j" and jn: "j \<le> 0"
-	from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
+        assume pns: "(p', n', s') \<in> ?SS a" 
+          and np: "n' < 0" 
+          and p_def: "p = ?p (p',n',s') j" 
+          and n0: "n = 0" 
+          and s_def: "s = (Add (Floor s') (C j))" 
+          and jp: "n' \<le> j" and jn: "j \<le> 0"
+        from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
           Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
           numbound0 s' \<and> isrlfm p'" by blast
-	hence nb: "numbound0 s'" by simp
-	from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numneg_nb)
-	let ?nxs = "CN 0 n' s'"
-	let ?l = "floor (?N s') + j"
-	from H 
-	have "?I (?p (p',n',s') j) \<longrightarrow> 
-	  (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
-	  by (simp add: np fp_def algebra_simps numneg numfloor numadd numsub)
-	also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
-	  using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
-	moreover
-	have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"  by simp
-	ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
-	  by blast
-	with s_def n0 p_def nb nf have ?ths by auto}
+        hence nb: "numbound0 s'" by simp
+        from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numneg_nb)
+        let ?nxs = "CN 0 n' s'"
+        let ?l = "floor (?N s') + j"
+        from H 
+        have "?I (?p (p',n',s') j) \<longrightarrow> 
+          (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
+          by (simp add: np fp_def algebra_simps numneg numfloor numadd numsub)
+        also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
+          using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
+        moreover
+        have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"  by simp
+        ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
+          by blast
+        with s_def n0 p_def nb nf have ?ths by auto}
       ultimately show ?ths by auto
     qed
 next
@@ -3811,7 +3811,7 @@
     {from x1 nn have "real n *x + ?N s \<ge> real n + ?N s" by simp
       moreover from real_of_int_floor_le[where r="?N s"]  have "real n + ?N s \<ge> real n + ?N (Floor s)" by simp
       ultimately have "real n *x + ?N s \<ge> ?N (Floor s) + real n" 
-	by (simp only: algebra_simps)}
+        by (simp only: algebra_simps)}
     ultimately have "?N (Floor s) + real n \<le> real n *x + ?N s\<and> real n *x + ?N s < ?N (Floor s) + real (1::int)" by simp
     hence th: "real n \<le> real n *x + ?N s - ?N (Floor s) \<and> real n *x + ?N s - ?N (Floor s) < real (1::int)" by simp
     have th1: "\<forall> (a::real). (- a > 0) = (a < 0)" by auto
@@ -3823,7 +3823,7 @@
     hence "\<exists> j\<in> {n .. 0}. 0 \<ge> - (real n *x + ?N s - ?N (Floor s) - real j) \<and> - (real n *x + ?N s - ?N (Floor s) - real (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format])
     hence "\<exists> j\<in> {n.. 0}. ?I (?p (p,n,s) j)"
       using pns by (simp add: fp_def nn diff_def add_ac mult_ac numfloor numadd numneg
-	del: diff_less_0_iff_less diff_le_0_iff_le) 
+        del: diff_less_0_iff_less diff_le_0_iff_le) 
     then obtain "j" where j_def: "j\<in> {n .. 0} \<and> ?I (?p (p,n,s) j)" by blast
     hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto
     hence ?case using pns 
@@ -4108,7 +4108,7 @@
   moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H DVD_def) } 
   moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th 
       by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 
-	rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
+        rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
   moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)}
   ultimately show ?th by blast
 qed
@@ -4126,7 +4126,7 @@
   moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H NDVD_def) } 
   moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th 
       by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 
-	rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
+        rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
   moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th 
       by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)}
   ultimately show ?th by blast
@@ -4197,11 +4197,11 @@
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
-	by (simp add: numgcd_def zgcd_le1)
+        by (simp add: numgcd_def zgcd_le1)
       from prems have th': "c\<noteq>0" by auto
       from prems have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
-	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
+        have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
     with prems have ?case
       by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
@@ -4222,11 +4222,11 @@
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
-	by (simp add: numgcd_def zgcd_le1)
+        by (simp add: numgcd_def zgcd_le1)
       from prems have th': "c\<noteq>0" by auto
       from prems have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
-	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
+        have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
     with prems have ?case
       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
@@ -4247,11 +4247,11 @@
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
-	by (simp add: numgcd_def zgcd_le1)
+        by (simp add: numgcd_def zgcd_le1)
       from prems have th': "c\<noteq>0" by auto
       from prems have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
-	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
+        have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
     with prems have ?case
       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
@@ -4272,11 +4272,11 @@
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
-	by (simp add: numgcd_def zgcd_le1)
+        by (simp add: numgcd_def zgcd_le1)
       from prems have th': "c\<noteq>0" by auto
       from prems have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
-	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
+        have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
     with prems have ?case
       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
@@ -4297,11 +4297,11 @@
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
-	by (simp add: numgcd_def zgcd_le1)
+        by (simp add: numgcd_def zgcd_le1)
       from prems have th': "c\<noteq>0" by auto
       from prems have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
-	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
+        have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
     with prems have ?case
       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
@@ -4322,11 +4322,11 @@
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
-	by (simp add: numgcd_def zgcd_le1)
+        by (simp add: numgcd_def zgcd_le1)
       from prems have th': "c\<noteq>0" by auto
       from prems have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
-	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
+        have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
     with prems have ?case
       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
@@ -4764,7 +4764,7 @@
     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
     moreover {assume y: "y < (-?N x e)/ real c"
       hence "y * real c < - ?N x e"
-	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
+        by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
       hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y > (- ?N x e) / real c" 
@@ -4783,7 +4783,7 @@
     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
     moreover {assume y: "y < (-?N x e)/ real c"
       hence "y * real c < - ?N x e"
-	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
+        by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
       hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y > (- ?N x e) / real c" 
@@ -4802,7 +4802,7 @@
     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
     moreover {assume y: "y > (-?N x e)/ real c"
       hence "y * real c > - ?N x e"
-	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
+        by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
       hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y < (- ?N x e) / real c" 
@@ -4821,7 +4821,7 @@
     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
     moreover {assume y: "y > (-?N x e)/ real c"
       hence "y * real c > - ?N x e"
-	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
+        by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
       hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y < (- ?N x e) / real c" 
@@ -4948,8 +4948,8 @@
     moreover{
       assume "\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p"
       then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M" 
-	and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
-	by blast
+        and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
+        by blast
       from t1M have "\<exists> (t1u,t1n) \<in> ?U. t1 = ?N a t1u / real t1n" by auto
       then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \<in> ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast
       from t2M have "\<exists> (t2u,t2n) \<in> ?U. t2 = ?N a t2u / real t2n" by auto
@@ -5005,13 +5005,13 @@
     let ?N = "\<lambda> t. Inum (x#bs) t"
     {fix t n s m assume "(t,n)\<in> set (\<Upsilon> p)" and "(s,m) \<in> set (\<Upsilon> p)"
       with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0"
-	by auto
+        by auto
       let ?st = "Add (Mul m t) (Mul n s)"
       from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
-	by (simp add: mult_commute)
+        by (simp add: mult_commute)
       from tnb snb have st_nb: "numbound0 ?st" by simp
       have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-	using mnp mp np by (simp add: algebra_simps add_divide_distrib)
+        using mnp mp np by (simp add: algebra_simps add_divide_distrib)
       from \<upsilon>_I[OF lp mnp st_nb, where x="x" and bs="bs"] 
       have "?I x (\<upsilon> p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])}
     with rinf_\<Upsilon>[OF lp nmi npi px] have "?F" by blast hence "?D" by blast}
@@ -5298,7 +5298,7 @@
     { fix t n assume tnY: "(t,n) \<in> set ?Y" 
       hence "(t,n) \<in> set ?SS" by simp
       hence "\<exists> (t',n') \<in> set ?S. simp_num_pair (t',n') = (t,n)"
-	by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all)
+        by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all)
       then obtain t' n' where tn'S: "(t',n') \<in> set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast
       from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto
       from simp_num_pair_l[OF tnb np tns]
--- a/src/HOL/Decision_Procs/cooper_tac.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Decision_Procs/cooper_tac.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -74,16 +74,16 @@
     val (t,np,nh) = prepare_for_linz q g
     (* Some simpsets for dealing with mod div abs and nat*)
     val mod_div_simpset = HOL_basic_ss
-			addsimps [refl,mod_add_eq, mod_add_left_eq,
-				  mod_add_right_eq,
-				  nat_div_add_eq, int_div_add_eq,
-				  @{thm mod_self}, @{thm "zmod_self"},
-				  @{thm mod_by_0}, @{thm div_by_0},
-				  @{thm "zdiv_zero"}, @{thm "zmod_zero"}, @{thm "div_0"}, @{thm "mod_0"},
-				  @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, @{thm "mod_1"},
-				  Suc_eq_plus1]
-			addsimps @{thms add_ac}
-			addsimprocs [cancel_div_mod_nat_proc, cancel_div_mod_int_proc]
+      addsimps [refl,mod_add_eq, mod_add_left_eq,
+          mod_add_right_eq,
+          nat_div_add_eq, int_div_add_eq,
+          @{thm mod_self}, @{thm "zmod_self"},
+          @{thm mod_by_0}, @{thm div_by_0},
+          @{thm "zdiv_zero"}, @{thm "zmod_zero"}, @{thm "div_0"}, @{thm "mod_0"},
+          @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, @{thm "mod_1"},
+          Suc_eq_plus1]
+      addsimps @{thms add_ac}
+      addsimprocs [cancel_div_mod_nat_proc, cancel_div_mod_int_proc]
     val simpset0 = HOL_basic_ss
       addsimps [mod_div_equality', Suc_eq_plus1]
       addsimps comp_arith
--- a/src/HOL/Decision_Procs/ferrack_tac.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Decision_Procs/ferrack_tac.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -16,9 +16,9 @@
 fun trace_msg s = if !trace then tracing s else ();
 
 val ferrack_ss = let val ths = [@{thm real_of_int_inject}, @{thm real_of_int_less_iff}, 
-				@{thm real_of_int_le_iff}]
-	     in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths)
-	     end;
+                                @{thm real_of_int_le_iff}]
+             in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths)
+             end;
 
 val binarith =
   @{thms normalize_bin_simps} @ @{thms pred_bin_simps} @ @{thms succ_bin_simps} @
@@ -74,8 +74,8 @@
 
 fun linr_tac ctxt q i = 
     (ObjectLogic.atomize_prems_tac i) 
-	THEN (REPEAT_DETERM (split_tac [@{thm split_min}, @{thm split_max}, @{thm abs_split}] i))
-	THEN (fn st =>
+        THEN (REPEAT_DETERM (split_tac [@{thm split_min}, @{thm split_max}, @{thm abs_split}] i))
+        THEN (fn st =>
   let
     val g = List.nth (prems_of st, i - 1)
     val thy = ProofContext.theory_of ctxt
--- a/src/HOL/Extraction/Euclid.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Extraction/Euclid.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,7 +1,7 @@
 (*  Title:      HOL/Extraction/Euclid.thy
     Author:     Markus Wenzel, TU Muenchen
-                Freek Wiedijk, Radboud University Nijmegen
-                Stefan Berghofer, TU Muenchen
+    Author:     Freek Wiedijk, Radboud University Nijmegen
+    Author:     Stefan Berghofer, TU Muenchen
 *)
 
 header {* Euclid's theorem *}
@@ -71,23 +71,23 @@
       assume nmk: "n = m * k"
       assume m: "Suc 0 < m"
       from n m nmk have k: "0 < k"
-	by (cases k) auto
+        by (cases k) auto
       moreover from n have n: "0 < n" by simp
       moreover note m
       moreover from nmk have "m * k = n" by simp
       ultimately have kn: "k < n" by (rule prod_mn_less_k)
       show "m = n"
       proof (cases "k = Suc 0")
-	case True
-	with nmk show ?thesis by (simp only: mult_Suc_right)
+        case True
+        with nmk show ?thesis by (simp only: mult_Suc_right)
       next
-	case False
-	from m have "0 < m" by simp
-	moreover note n
-	moreover from False n nmk k have "Suc 0 < k" by auto
-	moreover from nmk have "k * m = n" by (simp only: mult_ac)
-	ultimately have mn: "m < n" by (rule prod_mn_less_k)
-	with kn A nmk show ?thesis by iprover
+        case False
+        from m have "0 < m" by simp
+        moreover note n
+        moreover from False n nmk k have "Suc 0 < k" by auto
+        moreover from nmk have "k * m = n" by (simp only: mult_ac)
+        ultimately have mn: "m < n" by (rule prod_mn_less_k)
+        with kn A nmk show ?thesis by iprover
       qed
     qed
     with n have "prime n"
--- a/src/HOL/Extraction/Greatest_Common_Divisor.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Extraction/Greatest_Common_Divisor.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -24,7 +24,7 @@
     moreover have "Suc m * 1 = Suc m" by simp
     moreover {
       fix l2 have "\<And>l l1. l * l1 = n \<Longrightarrow> l * l2 = Suc m \<Longrightarrow> l \<le> Suc m"
-	by (cases l2) simp_all }
+        by (cases l2) simp_all }
     ultimately show ?thesis by iprover
   next
     case (Suc nat)
@@ -47,7 +47,7 @@
       assume ll1n: "l * l1 = n"
       assume ll2m: "l * l2 = Suc m"
       moreover have "l * (l1 - l2 * q) = Suc nat"
-	by (simp add: diff_mult_distrib2 h1 Suc [symmetric] mn ll1n ll2m [symmetric])
+        by (simp add: diff_mult_distrib2 h1 Suc [symmetric] mn ll1n ll2m [symmetric])
       ultimately show "l \<le> k" by (rule h3')
     qed
     ultimately show ?thesis using h1' by iprover
--- a/src/HOL/Extraction/Higman.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Extraction/Higman.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Extraction/Higman.thy
     Author:     Stefan Berghofer, TU Muenchen
-                Monika Seisenberger, LMU Muenchen
+    Author:     Monika Seisenberger, LMU Muenchen
 *)
 
 header {* Higman's lemma *}
@@ -167,21 +167,21 @@
       fix w
       show "bar (w # zs)"
       proof (cases w)
-	case Nil
-	thus ?thesis by simp (rule prop1)
+        case Nil
+        thus ?thesis by simp (rule prop1)
       next
-	case (Cons c cs)
-	from letter_eq_dec show ?thesis
-	proof
-	  assume ca: "c = a"
-	  from ab have "bar ((a # cs) # zs)" by (iprover intro: I ys Ta Tb)
-	  thus ?thesis by (simp add: Cons ca)
-	next
-	  assume "c \<noteq> a"
-	  with ab have cb: "c = b" by (rule letter_neq)
-	  from ab have "bar ((b # cs) # zs)" by (iprover intro: I' Ta Tb)
-	  thus ?thesis by (simp add: Cons cb)
-	qed
+        case (Cons c cs)
+        from letter_eq_dec show ?thesis
+        proof
+          assume ca: "c = a"
+          from ab have "bar ((a # cs) # zs)" by (iprover intro: I ys Ta Tb)
+          thus ?thesis by (simp add: Cons ca)
+        next
+          assume "c \<noteq> a"
+          with ab have cb: "c = b" by (rule letter_neq)
+          from ab have "bar ((b # cs) # zs)" by (iprover intro: I' Ta Tb)
+          thus ?thesis by (simp add: Cons cb)
+        qed
       qed
     qed
   qed
@@ -210,12 +210,12 @@
       case (Cons c cs)
       from letter_eq_dec show ?case
       proof
-	assume "c = a"
-	thus ?thesis by (iprover intro: I [simplified] R)
+        assume "c = a"
+        thus ?thesis by (iprover intro: I [simplified] R)
       next
-	from R xsn have T: "T a xs zs" by (rule lemma4)
-	assume "c \<noteq> a"
-	thus ?thesis by (iprover intro: prop2 Cons xsb xsn R T)
+        from R xsn have T: "T a xs zs" by (rule lemma4)
+        assume "c \<noteq> a"
+        thus ?thesis by (iprover intro: prop2 Cons xsb xsn R T)
       qed
     qed
   qed
--- a/src/HOL/Extraction/Pigeonhole.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Extraction/Pigeonhole.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -36,94 +36,94 @@
       let ?f = "\<lambda>i. if f i = Suc n then f (Suc (Suc n)) else f i"
       have "\<not> (\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j)"
       proof
-	assume "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j"
-      	then obtain i j where i: "i \<le> Suc n" and j: "j < i"
-	  and f: "?f i = ?f j" by iprover
-      	from j have i_nz: "Suc 0 \<le> i" by simp
-      	from i have iSSn: "i \<le> Suc (Suc n)" by simp
-      	have S0SSn: "Suc 0 \<le> Suc (Suc n)" by simp
-      	show False
-      	proof cases
-	  assume fi: "f i = Suc n"
-	  show False
-	  proof cases
-	    assume fj: "f j = Suc n"
-	    from i_nz and iSSn and j have "f i \<noteq> f j" by (rule 0)
-	    moreover from fi have "f i = f j"
-	      by (simp add: fj [symmetric])
-	    ultimately show ?thesis ..
-	  next
-	    from i and j have "j < Suc (Suc n)" by simp
-	    with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f j"
-	      by (rule 0)
-	    moreover assume "f j \<noteq> Suc n"
-	    with fi and f have "f (Suc (Suc n)) = f j" by simp
-	    ultimately show False ..
-	  qed
-      	next
-	  assume fi: "f i \<noteq> Suc n"
-	  show False
-	  proof cases
-	    from i have "i < Suc (Suc n)" by simp
-	    with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f i"
-	      by (rule 0)
-	    moreover assume "f j = Suc n"
-	    with fi and f have "f (Suc (Suc n)) = f i" by simp
-	    ultimately show False ..
-	  next
-	    from i_nz and iSSn and j
-	    have "f i \<noteq> f j" by (rule 0)
-	    moreover assume "f j \<noteq> Suc n"
-	    with fi and f have "f i = f j" by simp
-	    ultimately show False ..
-	  qed
-      	qed
+        assume "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j"
+        then obtain i j where i: "i \<le> Suc n" and j: "j < i"
+          and f: "?f i = ?f j" by iprover
+        from j have i_nz: "Suc 0 \<le> i" by simp
+        from i have iSSn: "i \<le> Suc (Suc n)" by simp
+        have S0SSn: "Suc 0 \<le> Suc (Suc n)" by simp
+        show False
+        proof cases
+          assume fi: "f i = Suc n"
+          show False
+          proof cases
+            assume fj: "f j = Suc n"
+            from i_nz and iSSn and j have "f i \<noteq> f j" by (rule 0)
+            moreover from fi have "f i = f j"
+              by (simp add: fj [symmetric])
+            ultimately show ?thesis ..
+          next
+            from i and j have "j < Suc (Suc n)" by simp
+            with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f j"
+              by (rule 0)
+            moreover assume "f j \<noteq> Suc n"
+            with fi and f have "f (Suc (Suc n)) = f j" by simp
+            ultimately show False ..
+          qed
+        next
+          assume fi: "f i \<noteq> Suc n"
+          show False
+          proof cases
+            from i have "i < Suc (Suc n)" by simp
+            with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f i"
+              by (rule 0)
+            moreover assume "f j = Suc n"
+            with fi and f have "f (Suc (Suc n)) = f i" by simp
+            ultimately show False ..
+          next
+            from i_nz and iSSn and j
+            have "f i \<noteq> f j" by (rule 0)
+            moreover assume "f j \<noteq> Suc n"
+            with fi and f have "f i = f j" by simp
+            ultimately show False ..
+          qed
+        qed
       qed
       moreover have "\<And>i. i \<le> Suc n \<Longrightarrow> ?f i \<le> n"
       proof -
-	fix i assume "i \<le> Suc n"
-	hence i: "i < Suc (Suc n)" by simp
-	have "f (Suc (Suc n)) \<noteq> f i"
-	  by (rule 0) (simp_all add: i)
-	moreover have "f (Suc (Suc n)) \<le> Suc n"
-	  by (rule Suc) simp
-	moreover from i have "i \<le> Suc (Suc n)" by simp
-	hence "f i \<le> Suc n" by (rule Suc)
-	ultimately show "?thesis i"
-	  by simp
+        fix i assume "i \<le> Suc n"
+        hence i: "i < Suc (Suc n)" by simp
+        have "f (Suc (Suc n)) \<noteq> f i"
+          by (rule 0) (simp_all add: i)
+        moreover have "f (Suc (Suc n)) \<le> Suc n"
+          by (rule Suc) simp
+        moreover from i have "i \<le> Suc (Suc n)" by simp
+        hence "f i \<le> Suc n" by (rule Suc)
+        ultimately show "?thesis i"
+          by simp
       qed
       hence "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j"
-      	by (rule Suc)
+        by (rule Suc)
       ultimately show ?case ..
     next
       case (Suc k)
       from search [OF nat_eq_dec] show ?case
       proof
-	assume "\<exists>j<Suc k. f (Suc k) = f j"
-	thus ?case by (iprover intro: le_refl)
+        assume "\<exists>j<Suc k. f (Suc k) = f j"
+        thus ?case by (iprover intro: le_refl)
       next
-	assume nex: "\<not> (\<exists>j<Suc k. f (Suc k) = f j)"
-	have "\<exists>i j. i \<le> k \<and> j < i \<and> f i = f j"
-	proof (rule Suc)
-	  from Suc show "k \<le> Suc (Suc n)" by simp
-	  fix i j assume k: "Suc k \<le> i" and i: "i \<le> Suc (Suc n)"
-	    and j: "j < i"
-	  show "f i \<noteq> f j"
-	  proof cases
-	    assume eq: "i = Suc k"
-	    show ?thesis
-	    proof
-	      assume "f i = f j"
-	      hence "f (Suc k) = f j" by (simp add: eq)
-	      with nex and j and eq show False by iprover
-	    qed
-	  next
-	    assume "i \<noteq> Suc k"
-	    with k have "Suc (Suc k) \<le> i" by simp
-	    thus ?thesis using i and j by (rule Suc)
-	  qed
-	qed
-	thus ?thesis by (iprover intro: le_SucI)
+        assume nex: "\<not> (\<exists>j<Suc k. f (Suc k) = f j)"
+        have "\<exists>i j. i \<le> k \<and> j < i \<and> f i = f j"
+        proof (rule Suc)
+          from Suc show "k \<le> Suc (Suc n)" by simp
+          fix i j assume k: "Suc k \<le> i" and i: "i \<le> Suc (Suc n)"
+            and j: "j < i"
+          show "f i \<noteq> f j"
+          proof cases
+            assume eq: "i = Suc k"
+            show ?thesis
+            proof
+              assume "f i = f j"
+              hence "f (Suc k) = f j" by (simp add: eq)
+              with nex and j and eq show False by iprover
+            qed
+          next
+            assume "i \<noteq> Suc k"
+            with k have "Suc (Suc k) \<le> i" by simp
+            thus ?thesis using i and j by (rule Suc)
+          qed
+        qed
+        thus ?thesis by (iprover intro: le_SucI)
       qed
     qed
   }
@@ -159,16 +159,16 @@
       fix i assume i: "i \<le> Suc n"
       show "?thesis i"
       proof (cases "f i = Suc n")
-	case True
-	from i and nex have "f (Suc (Suc n)) \<noteq> f i" by simp
-	with True have "f (Suc (Suc n)) \<noteq> Suc n" by simp
-	moreover from Suc have "f (Suc (Suc n)) \<le> Suc n" by simp
-	ultimately have "f (Suc (Suc n)) \<le> n" by simp
-	with True show ?thesis by simp
+        case True
+        from i and nex have "f (Suc (Suc n)) \<noteq> f i" by simp
+        with True have "f (Suc (Suc n)) \<noteq> Suc n" by simp
+        moreover from Suc have "f (Suc (Suc n)) \<le> Suc n" by simp
+        ultimately have "f (Suc (Suc n)) \<le> n" by simp
+        with True show ?thesis by simp
       next
-	case False
-	from Suc and i have "f i \<le> Suc n" by simp
-	with False show ?thesis by simp
+        case False
+        from Suc and i have "f i \<le> Suc n" by simp
+        with False show ?thesis by simp
       qed
     qed
     hence "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j" by (rule Suc)
@@ -179,23 +179,23 @@
       case True
       show ?thesis
       proof (cases "f j = Suc n")
-	assume "f j = Suc n"
-	with True show ?thesis by simp
+        assume "f j = Suc n"
+        with True show ?thesis by simp
       next
-	assume "f j \<noteq> Suc n"
-	moreover from i ji nex have "f (Suc (Suc n)) \<noteq> f j" by simp
-	ultimately show ?thesis using True f by simp
+        assume "f j \<noteq> Suc n"
+        moreover from i ji nex have "f (Suc (Suc n)) \<noteq> f j" by simp
+        ultimately show ?thesis using True f by simp
       qed
     next
       case False
       show ?thesis
       proof (cases "f j = Suc n")
-	assume "f j = Suc n"
-	moreover from i nex have "f (Suc (Suc n)) \<noteq> f i" by simp
-	ultimately show ?thesis using False f by simp
+        assume "f j = Suc n"
+        moreover from i nex have "f (Suc (Suc n)) \<noteq> f i" by simp
+        ultimately show ?thesis using False f by simp
       next
-	assume "f j \<noteq> Suc n"
-	with False f show ?thesis by simp
+        assume "f j \<noteq> Suc n"
+        with False f show ?thesis by simp
       qed
     qed
     moreover from i have "i \<le> Suc (Suc n)" by simp
--- a/src/HOL/Extraction/Util.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Extraction/Util.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Extraction/Util.thy
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 *)
 
@@ -39,7 +38,7 @@
     proof
       assume ny: "n = y"
       have "P n"
-	by (rule R) (rule Suc)
+        by (rule R) (rule Suc)
       with ny show ?case by simp
     next
       assume "n \<noteq> y"
@@ -77,18 +76,18 @@
       assume nP: "\<not> P z"
       have "\<not> (\<exists>x<Suc z. P x)"
       proof
-	assume "\<exists>x<Suc z. P x"
-	then obtain x where le: "x < Suc z" and P: "P x" by iprover
-	have "x < z"
-	proof (cases "x = z")
-	  case True
-	  with nP and P show ?thesis by simp
-	next
-	  case False
-	  with le show ?thesis by simp
-	qed
-	with P have "\<exists>x<z. P x" by iprover
-	with nex show False ..
+        assume "\<exists>x<Suc z. P x"
+        then obtain x where le: "x < Suc z" and P: "P x" by iprover
+        have "x < z"
+        proof (cases "x = z")
+          case True
+          with nP and P show ?thesis by simp
+        next
+          case False
+          with le show ?thesis by simp
+        qed
+        with P have "\<exists>x<z. P x" by iprover
+        with nex show False ..
       qed
       thus ?case by iprover
     qed
--- a/src/HOL/Extraction/Warshall.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Extraction/Warshall.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -129,26 +129,26 @@
       case (Cons a as j)
       show ?case
       proof (cases "a=i")
-      	case True
-      	show ?thesis
-      	proof
-	  from True and Cons have "r j i = T" by simp
-	  thus "list_all (\<lambda>x. x < i) [] \<and> is_path' r j [] i" by simp
-      	qed
+        case True
+        show ?thesis
+        proof
+          from True and Cons have "r j i = T" by simp
+          thus "list_all (\<lambda>x. x < i) [] \<and> is_path' r j [] i" by simp
+        qed
       next
-      	case False
-      	have "PROP ?ih as" by (rule Cons)
-      	then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r a ys i"
-      	proof
-	  from Cons show "list_all (\<lambda>x. x < Suc i) as" by simp
-	  from Cons show "is_path' r a as k" by simp
-	  from Cons and False show "\<not> list_all (\<lambda>x. x < i) as" by (simp)
-      	qed
-      	show ?thesis
-      	proof
-	  from Cons False ys
-	  show "list_all (\<lambda>x. x<i) (a#ys) \<and> is_path' r j (a#ys) i" by simp
-      	qed
+        case False
+        have "PROP ?ih as" by (rule Cons)
+        then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r a ys i"
+        proof
+          from Cons show "list_all (\<lambda>x. x < Suc i) as" by simp
+          from Cons show "is_path' r a as k" by simp
+          from Cons and False show "\<not> list_all (\<lambda>x. x < i) as" by (simp)
+        qed
+        show ?thesis
+        proof
+          from Cons False ys
+          show "list_all (\<lambda>x. x<i) (a#ys) \<and> is_path' r j (a#ys) i" by simp
+        qed
       qed
     qed
     show "\<And>k. list_all (\<lambda>x. x < Suc i) xs \<Longrightarrow> is_path' r j xs k \<Longrightarrow>
@@ -161,27 +161,27 @@
       case (snoc a as k)
       show ?case
       proof (cases "a=i")
-      	case True
-      	show ?thesis
-      	proof
-	  from True and snoc have "r i k = T" by simp
-	  thus "list_all (\<lambda>x. x < i) [] \<and> is_path' r i [] k" by simp
-      	qed
+        case True
+        show ?thesis
+        proof
+          from True and snoc have "r i k = T" by simp
+          thus "list_all (\<lambda>x. x < i) [] \<and> is_path' r i [] k" by simp
+        qed
       next
-      	case False
-      	have "PROP ?ih as" by (rule snoc)
-      	then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys a"
-      	proof
-	  from snoc show "list_all (\<lambda>x. x < Suc i) as" by simp
-	  from snoc show "is_path' r j as a" by simp
-	  from snoc and False show "\<not> list_all (\<lambda>x. x < i) as" by simp
-      	qed
-      	show ?thesis
-      	proof
-	  from snoc False ys
-	  show "list_all (\<lambda>x. x < i) (ys @ [a]) \<and> is_path' r i (ys @ [a]) k"
-	    by simp
-      	qed
+        case False
+        have "PROP ?ih as" by (rule snoc)
+        then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys a"
+        proof
+          from snoc show "list_all (\<lambda>x. x < Suc i) as" by simp
+          from snoc show "is_path' r j as a" by simp
+          from snoc and False show "\<not> list_all (\<lambda>x. x < i) as" by simp
+        qed
+        show ?thesis
+        proof
+          from snoc False ys
+          show "list_all (\<lambda>x. x < i) (ys @ [a]) \<and> is_path' r i (ys @ [a]) k"
+            by simp
+        qed
       qed
     qed
   qed (rule asms)+
@@ -219,24 +219,24 @@
     proof
       assume "\<not> (\<exists>p. is_path r p i j i)"
       with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)"
-	by (iprover dest: lemma5')
+        by (iprover dest: lemma5')
       thus ?case ..
     next
       assume "\<exists>p. is_path r p i j i"
       then obtain p where h2: "is_path r p i j i" ..
       from Suc show ?case
       proof
-	assume "\<not> (\<exists>p. is_path r p i i k)"
-	with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)"
-	  by (iprover dest: lemma5')
-	thus ?case ..
+        assume "\<not> (\<exists>p. is_path r p i i k)"
+        with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)"
+          by (iprover dest: lemma5')
+        thus ?case ..
       next
-	assume "\<exists>q. is_path r q i i k"
-	then obtain q where "is_path r q i i k" ..
-	with h2 have "is_path r (conc p q) (Suc i) j k" 
-	  by (rule lemma3)
-	hence "\<exists>pq. is_path r pq (Suc i) j k" ..
-	thus ?case ..
+        assume "\<exists>q. is_path r q i i k"
+        then obtain q where "is_path r q i i k" ..
+        with h2 have "is_path r (conc p q) (Suc i) j k" 
+          by (rule lemma3)
+        hence "\<exists>pq. is_path r pq (Suc i) j k" ..
+        thus ?case ..
       qed
     qed
   next
--- a/src/HOL/Finite_Set.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Finite_Set.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -236,22 +236,22 @@
     {
       fix c b :: "'a set"
       assume c: "finite c" and b: "finite b"
-	and P1: "P b" and P2: "!!x y. finite y ==> x \<in> y ==> P y ==> P (y - {x})"
+        and P1: "P b" and P2: "!!x y. finite y ==> x \<in> y ==> P y ==> P (y - {x})"
       have "c \<subseteq> b ==> P (b - c)"
-	using c
+        using c
       proof induct
-	case empty
-	from P1 show ?case by simp
+        case empty
+        from P1 show ?case by simp
       next
-	case (insert x F)
-	have "P (b - F - {x})"
-	proof (rule P2)
+        case (insert x F)
+        have "P (b - F - {x})"
+        proof (rule P2)
           from _ b show "finite (b - F)" by (rule finite_subset) blast
           from insert show "x \<in> b - F" by simp
           from insert show "P (b - F)" by simp
-	qed
-	also have "b - F - {x} = b - insert x F" by (rule Diff_insert [symmetric])
-	finally show ?case .
+        qed
+        also have "b - F - {x} = b - insert x F" by (rule Diff_insert [symmetric])
+        finally show ?case .
       qed
     }
     then show ?thesis by this (simp_all add: assms)
@@ -617,9 +617,9 @@
       show "inj_on (Fun.swap k m h) {i. i < Suc m}" using inj_hm nSuc by simp
       show "?hm m \<notin> A" by (simp add: swap_def hkeq anot) 
       show "insert (?hm m) A = ?hm ` {i. i < Suc m}"
-	using aA hkeq nSuc klessn
-	by (auto simp add: swap_def image_less_Suc fun_upd_image 
-			   less_Suc_eq inj_on_image_set_diff [OF inj_on])
+        using aA hkeq nSuc klessn
+        by (auto simp add: swap_def image_less_Suc fun_upd_image 
+                           less_Suc_eq inj_on_image_set_diff [OF inj_on])
     qed
   qed
 qed
@@ -665,26 +665,26 @@
       show "x'=x"
       proof cases
         assume "b=c"
-	then moreover have "B = C" using AbB AcC notinB notinC by auto
-	ultimately show ?thesis  using Bu Cv x x' IH [OF lessC Ceq inj_onC]
+        then moreover have "B = C" using AbB AcC notinB notinC by auto
+        ultimately show ?thesis  using Bu Cv x x' IH [OF lessC Ceq inj_onC]
           by auto
       next
-	assume diff: "b \<noteq> c"
-	let ?D = "B - {c}"
-	have B: "B = insert c ?D" and C: "C = insert b ?D"
-	  using AbB AcC notinB notinC diff by(blast elim!:equalityE)+
-	have "finite A" by(rule fold_graph_imp_finite [OF Afoldx])
-	with AbB have "finite ?D" by simp
-	then obtain d where Dfoldd: "fold_graph f z ?D d"
-	  using finite_imp_fold_graph by iprover
-	moreover have cinB: "c \<in> B" using B by auto
-	ultimately have "fold_graph f z B (f c d)" by(rule Diff1_fold_graph)
-	hence "f c d = u" by (rule IH [OF lessB Beq inj_onB Bu]) 
+        assume diff: "b \<noteq> c"
+        let ?D = "B - {c}"
+        have B: "B = insert c ?D" and C: "C = insert b ?D"
+          using AbB AcC notinB notinC diff by(blast elim!:equalityE)+
+        have "finite A" by(rule fold_graph_imp_finite [OF Afoldx])
+        with AbB have "finite ?D" by simp
+        then obtain d where Dfoldd: "fold_graph f z ?D d"
+          using finite_imp_fold_graph by iprover
+        moreover have cinB: "c \<in> B" using B by auto
+        ultimately have "fold_graph f z B (f c d)" by(rule Diff1_fold_graph)
+        hence "f c d = u" by (rule IH [OF lessB Beq inj_onB Bu]) 
         moreover have "f b d = v"
-	proof (rule IH[OF lessC Ceq inj_onC Cv])
-	  show "fold_graph f z C (f b d)" using C notinB Dfoldd by fastsimp
-	qed
-	ultimately show ?thesis
+        proof (rule IH[OF lessC Ceq inj_onC Cv])
+          show "fold_graph f z C (f b d)" using C notinB Dfoldd by fastsimp
+        qed
+        ultimately show ?thesis
           using fun_left_comm [of c b] x x' by (auto simp add: o_def)
       qed
     qed
@@ -2588,7 +2588,7 @@
     next
       assume "A' \<noteq> {}"
       with prems show ?thesis
-	by (simp add: fold1_insert mult_assoc [symmetric] mult_idem)
+        by (simp add: fold1_insert mult_assoc [symmetric] mult_idem)
     qed
   next
     assume "a \<noteq> x"
--- a/src/HOL/GCD.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/GCD.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1190,35 +1190,35 @@
      moreover
      {assume db: "d=b"
        from prems have ?thesis apply simp
-	 apply (rule exI[where x = b], simp)
-	 apply (rule exI[where x = b])
-	by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
+         apply (rule exI[where x = b], simp)
+         apply (rule exI[where x = b])
+        by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
     moreover
     {assume db: "d < b"
-	{assume "x=0" hence ?thesis  using prems by simp }
-	moreover
-	{assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
-	  from db have "d \<le> b - 1" by simp
-	  hence "d*b \<le> b*(b - 1)" by simp
-	  with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
-	  have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
-	  from H (3) have "d + (b - 1) * (b*x) = d + (b - 1) * (a*y + d)"
+        {assume "x=0" hence ?thesis  using prems by simp }
+        moreover
+        {assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
+          from db have "d \<le> b - 1" by simp
+          hence "d*b \<le> b*(b - 1)" by simp
+          with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
+          have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
+          from H (3) have "d + (b - 1) * (b*x) = d + (b - 1) * (a*y + d)"
             by simp
-	  hence "d + (b - 1) * a * y + (b - 1) * d = d + (b - 1) * b * x"
-	    by (simp only: mult_assoc right_distrib)
-	  hence "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)"
+          hence "d + (b - 1) * a * y + (b - 1) * d = d + (b - 1) * b * x"
+            by (simp only: mult_assoc right_distrib)
+          hence "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)"
             by algebra
-	  hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
-	  hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)"
-	    by (simp only: diff_add_assoc[OF dble, of d, symmetric])
-	  hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
-	    by (simp only: diff_mult_distrib2 add_commute mult_ac)
-	  hence ?thesis using H(1,2)
-	    apply -
-	    apply (rule exI[where x=d], simp)
-	    apply (rule exI[where x="(b - 1) * y"])
-	    by (rule exI[where x="x*(b - 1) - d"], simp)}
-	ultimately have ?thesis by blast}
+          hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
+          hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)"
+            by (simp only: diff_add_assoc[OF dble, of d, symmetric])
+          hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
+            by (simp only: diff_mult_distrib2 add_commute mult_ac)
+          hence ?thesis using H(1,2)
+            apply -
+            apply (rule exI[where x=d], simp)
+            apply (rule exI[where x="(b - 1) * y"])
+            by (rule exI[where x="x*(b - 1) - d"], simp)}
+        ultimately have ?thesis by blast}
     ultimately have ?thesis by blast}
   ultimately have ?thesis by blast}
  ultimately show ?thesis by blast
@@ -1653,23 +1653,23 @@
     next
       assume "M\<noteq>{0}"
       with `M\<noteq>{}` assms show ?thesis
-	apply(subst Gcd_remove0_nat[OF assms])
-	apply(simp add:GCD_def)
-	apply(subst divisors_remove0_nat)
-	apply(simp add:LCM_def)
-	apply rule
-	 apply rule
-	 apply(subst Gcd_eq_Max)
-	    apply simp
-	   apply blast
-	  apply blast
-	 apply(rule Lcm_eq_Max_nat)
-	    apply simp
-	   apply blast
-	  apply fastsimp
-	 apply clarsimp
-	apply(fastsimp intro: finite_divisors_nat intro!: finite_INT)
-	done
+        apply(subst Gcd_remove0_nat[OF assms])
+        apply(simp add:GCD_def)
+        apply(subst divisors_remove0_nat)
+        apply(simp add:LCM_def)
+        apply rule
+         apply rule
+         apply(subst Gcd_eq_Max)
+            apply simp
+           apply blast
+          apply blast
+         apply(rule Lcm_eq_Max_nat)
+            apply simp
+           apply blast
+          apply fastsimp
+         apply clarsimp
+        apply(fastsimp intro: finite_divisors_nat intro!: finite_INT)
+        done
     qed
   qed
 qed
--- a/src/HOL/Hahn_Banach/Bounds.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hahn_Banach/Bounds.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -34,16 +34,16 @@
       fix x' assume lub': "lub A x'"
       show "x' = x"
       proof (rule order_antisym)
-	from lub' show "x' \<le> x"
-	proof
+        from lub' show "x' \<le> x"
+        proof
           fix a assume "a \<in> A"
           then show "a \<le> x" ..
-	qed
-	show "x \<le> x'"
-	proof
+        qed
+        show "x \<le> x'"
+        proof
           fix a assume "a \<in> A"
           with lub' show "a \<le> x'" ..
-	qed
+        qed
       qed
     qed
   qed
--- a/src/HOL/Hahn_Banach/Function_Norm.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hahn_Banach/Function_Norm.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -224,9 +224,9 @@
     proof (rule mult_right_mono)
       from x show "0 \<le> \<parallel>x\<parallel>" ..
       from x and neq have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<in> B V f"
-	by (auto simp add: B_def real_divide_def)
+        by (auto simp add: B_def real_divide_def)
       with `continuous V norm f` show "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>V"
-	by (rule fn_norm_ub)
+        by (rule fn_norm_ub)
     qed
     finally show ?thesis .
   qed
@@ -257,18 +257,18 @@
       assume "b \<noteq> 0"
       with b obtain x where b_rep: "b = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
         and x_neq: "x \<noteq> 0" and x: "x \<in> V"
-	by (auto simp add: B_def real_divide_def)
+        by (auto simp add: B_def real_divide_def)
       note b_rep
       also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
       proof (rule mult_right_mono)
-	have "0 < \<parallel>x\<parallel>" using x x_neq ..
-	then show "0 \<le> inverse \<parallel>x\<parallel>" by simp
-	from ineq and x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
+        have "0 < \<parallel>x\<parallel>" using x x_neq ..
+        then show "0 \<le> inverse \<parallel>x\<parallel>" by simp
+        from ineq and x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
       qed
       also have "\<dots> = c"
       proof -
-	from x_neq and x have "\<parallel>x\<parallel> \<noteq> 0" by simp
-	then show ?thesis by simp
+        from x_neq and x have "\<parallel>x\<parallel> \<noteq> 0" by simp
+        then show ?thesis by simp
       qed
       finally show ?thesis .
     qed
--- a/src/HOL/Hahn_Banach/Hahn_Banach.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -202,7 +202,7 @@
           proof (rule graph_extI)
             fix t assume t: "t \<in> H"
             from E HE t have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
-	      using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
+              using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
             with h'_def show "h t = h' t" by (simp add: Let_def)
           next
             from HH' show "H \<subseteq> H'" ..
@@ -216,7 +216,7 @@
           proof
             assume eq: "graph H h = graph H' h'"
             have "x' \<in> H'"
-	      unfolding H'_def
+              unfolding H'_def
             proof
               from H show "0 \<in> H" by (rule vectorspace.zero)
               from x'E show "x' \<in> lin x'" by (rule x_lin_x)
@@ -236,10 +236,10 @@
         show "graph H' h' \<in> norm_pres_extensions E p F f"
         proof (rule norm_pres_extensionI2)
           show "linearform H' h'"
-	    using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
-	    by (rule h'_lf)
+            using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
+            by (rule h'_lf)
           show "H' \<unlhd> E"
-	  unfolding H'_def
+          unfolding H'_def
           proof
             show "H \<unlhd> E" by fact
             show "vectorspace E" by fact
@@ -256,12 +256,12 @@
               by (simp add: Let_def)
             also have "(x, 0) =
                 (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
-	      using E HE
+              using E HE
             proof (rule decomp_H'_H [symmetric])
               from FH x show "x \<in> H" ..
               from x' show "x' \<noteq> 0" .
-	      show "x' \<notin> H" by fact
-	      show "x' \<in> E" by fact
+              show "x' \<notin> H" by fact
+              show "x' \<in> E" by fact
             qed
             also have
               "(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
@@ -271,9 +271,9 @@
             from FH' show "F \<subseteq> H'" ..
           qed
           show "\<forall>x \<in> H'. h' x \<le> p x"
-	    using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
-	      `seminorm E p` linearform and hp xi
-	    by (rule h'_norm_pres)
+            using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
+              `seminorm E p` linearform and hp xi
+            by (rule h'_norm_pres)
         qed
       qed
       ultimately show ?thesis ..
@@ -483,23 +483,23 @@
 
     show "\<parallel>f\<parallel>\<hyphen>F \<le> \<parallel>g\<parallel>\<hyphen>E"
     proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
-	[OF normed_vectorspace_with_fn_norm.intro,
-	 OF F_norm, folded B_def fn_norm_def])
+        [OF normed_vectorspace_with_fn_norm.intro,
+         OF F_norm, folded B_def fn_norm_def])
       show "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
       proof
-	fix x assume x: "x \<in> F"
-	from a x have "g x = f x" ..
-	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
-	also from g_cont
-	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
-	proof (rule fn_norm_le_cong [OF _ linearformE, folded B_def fn_norm_def])
-	  from FE x show "x \<in> E" ..
-	qed
-	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
+        fix x assume x: "x \<in> F"
+        from a x have "g x = f x" ..
+        then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
+        also from g_cont
+        have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
+        proof (rule fn_norm_le_cong [OF _ linearformE, folded B_def fn_norm_def])
+          from FE x show "x \<in> E" ..
+        qed
+        finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
       qed
       show "0 \<le> \<parallel>g\<parallel>\<hyphen>E"
-	using g_cont
-	by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
+        using g_cont
+        by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
       show "continuous F norm f" by fact
     qed
   qed
--- a/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -112,16 +112,16 @@
       fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
       show "h' (x1 + x2) = h' x1 + h' x2"
       proof -
-	from H' x1 x2 have "x1 + x2 \<in> H'"
+        from H' x1 x2 have "x1 + x2 \<in> H'"
           by (rule vectorspace.add_closed)
-	with x1 x2 obtain y y1 y2 a a1 a2 where
+        with x1 x2 obtain y y1 y2 a a1 a2 where
           x1x2: "x1 + x2 = y + a \<cdot> x0" and y: "y \<in> H"
           and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
           and x2_rep: "x2 = y2 + a2 \<cdot> x0" and y2: "y2 \<in> H"
           unfolding H'_def sum_def lin_def by blast
-	
-	have ya: "y1 + y2 = y \<and> a1 + a2 = a" using E HE _ y x0
-	proof (rule decomp_H') txt_raw {* \label{decomp-H-use} *}
+        
+        have ya: "y1 + y2 = y \<and> a1 + a2 = a" using E HE _ y x0
+        proof (rule decomp_H') txt_raw {* \label{decomp-H-use} *}
           from HE y1 y2 show "y1 + y2 \<in> H"
             by (rule subspace.add_closed)
           from x0 and HE y y1 y2
@@ -130,38 +130,38 @@
             by (simp add: add_ac add_mult_distrib2)
           also note x1x2
           finally show "(y1 + y2) + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0" .
-	qed
-	
-	from h'_def x1x2 E HE y x0
-	have "h' (x1 + x2) = h y + a * xi"
+        qed
+        
+        from h'_def x1x2 E HE y x0
+        have "h' (x1 + x2) = h y + a * xi"
           by (rule h'_definite)
-	also have "\<dots> = h (y1 + y2) + (a1 + a2) * xi"
+        also have "\<dots> = h (y1 + y2) + (a1 + a2) * xi"
           by (simp only: ya)
-	also from y1 y2 have "h (y1 + y2) = h y1 + h y2"
+        also from y1 y2 have "h (y1 + y2) = h y1 + h y2"
           by simp
-	also have "\<dots> + (a1 + a2) * xi = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
+        also have "\<dots> + (a1 + a2) * xi = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
           by (simp add: left_distrib)
-	also from h'_def x1_rep E HE y1 x0
-	have "h y1 + a1 * xi = h' x1"
+        also from h'_def x1_rep E HE y1 x0
+        have "h y1 + a1 * xi = h' x1"
           by (rule h'_definite [symmetric])
-	also from h'_def x2_rep E HE y2 x0
-	have "h y2 + a2 * xi = h' x2"
+        also from h'_def x2_rep E HE y2 x0
+        have "h y2 + a2 * xi = h' x2"
           by (rule h'_definite [symmetric])
-	finally show ?thesis .
+        finally show ?thesis .
       qed
     next
       fix x1 c assume x1: "x1 \<in> H'"
       show "h' (c \<cdot> x1) = c * (h' x1)"
       proof -
-	from H' x1 have ax1: "c \<cdot> x1 \<in> H'"
+        from H' x1 have ax1: "c \<cdot> x1 \<in> H'"
           by (rule vectorspace.mult_closed)
-	with x1 obtain y a y1 a1 where
+        with x1 obtain y a y1 a1 where
             cx1_rep: "c \<cdot> x1 = y + a \<cdot> x0" and y: "y \<in> H"
           and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
           unfolding H'_def sum_def lin_def by blast
-	
-	have ya: "c \<cdot> y1 = y \<and> c * a1 = a" using E HE _ y x0
-	proof (rule decomp_H')
+        
+        have ya: "c \<cdot> y1 = y \<and> c * a1 = a" using E HE _ y x0
+        proof (rule decomp_H')
           from HE y1 show "c \<cdot> y1 \<in> H"
             by (rule subspace.mult_closed)
           from x0 and HE y y1
@@ -170,19 +170,19 @@
             by (simp add: mult_assoc add_mult_distrib1)
           also note cx1_rep
           finally show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0" .
-	qed
-	
-	from h'_def cx1_rep E HE y x0 have "h' (c \<cdot> x1) = h y + a * xi"
+        qed
+        
+        from h'_def cx1_rep E HE y x0 have "h' (c \<cdot> x1) = h y + a * xi"
           by (rule h'_definite)
-	also have "\<dots> = h (c \<cdot> y1) + (c * a1) * xi"
+        also have "\<dots> = h (c \<cdot> y1) + (c * a1) * xi"
           by (simp only: ya)
-	also from y1 have "h (c \<cdot> y1) = c * h y1"
+        also from y1 have "h (c \<cdot> y1) = c * h y1"
           by simp
-	also have "\<dots> + (c * a1) * xi = c * (h y1 + a1 * xi)"
+        also have "\<dots> + (c * a1) * xi = c * (h y1 + a1 * xi)"
           by (simp only: right_distrib)
-	also from h'_def x1_rep E HE y1 x0 have "h y1 + a1 * xi = h' x1"
+        also from h'_def x1_rep E HE y1 x0 have "h y1 + a1 * xi = h' x1"
           by (rule h'_definite [symmetric])
-	finally show ?thesis .
+        finally show ?thesis .
       qed
     }
   qed
@@ -212,64 +212,64 @@
     show "h' x \<le> p x"
     proof -
       from a' have a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya \<le> xi"
-	and a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya" by auto
+        and a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya" by auto
       from x' obtain y a where
           x_rep: "x = y + a \<cdot> x0" and y: "y \<in> H"
-	unfolding H'_def sum_def lin_def by blast
+        unfolding H'_def sum_def lin_def by blast
       from y have y': "y \<in> E" ..
       from y have ay: "inverse a \<cdot> y \<in> H" by simp
       
       from h'_def x_rep E HE y x0 have "h' x = h y + a * xi"
-	by (rule h'_definite)
+        by (rule h'_definite)
       also have "\<dots> \<le> p (y + a \<cdot> x0)"
       proof (rule linorder_cases)
-	assume z: "a = 0"
-	then have "h y + a * xi = h y" by simp
-	also from a y have "\<dots> \<le> p y" ..
-	also from x0 y' z have "p y = p (y + a \<cdot> x0)" by simp
-	finally show ?thesis .
+        assume z: "a = 0"
+        then have "h y + a * xi = h y" by simp
+        also from a y have "\<dots> \<le> p y" ..
+        also from x0 y' z have "p y = p (y + a \<cdot> x0)" by simp
+        finally show ?thesis .
       next
-	txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
+        txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
           with @{text ya} taken as @{text "y / a"}: *}
-	assume lz: "a < 0" then have nz: "a \<noteq> 0" by simp
-	from a1 ay
-	have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi" ..
-	with lz have "a * xi \<le>
+        assume lz: "a < 0" then have nz: "a \<noteq> 0" by simp
+        from a1 ay
+        have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi" ..
+        with lz have "a * xi \<le>
           a * (- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
           by (simp add: mult_left_mono_neg order_less_imp_le)
-	
-	also have "\<dots> =
+        
+        also have "\<dots> =
           - a * (p (inverse a \<cdot> y + x0)) - a * (h (inverse a \<cdot> y))"
-	  by (simp add: right_diff_distrib)
-	also from lz x0 y' have "- a * (p (inverse a \<cdot> y + x0)) =
+          by (simp add: right_diff_distrib)
+        also from lz x0 y' have "- a * (p (inverse a \<cdot> y + x0)) =
           p (a \<cdot> (inverse a \<cdot> y + x0))"
           by (simp add: abs_homogenous)
-	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
+        also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
           by (simp add: add_mult_distrib1 mult_assoc [symmetric])
-	also from nz y have "a * (h (inverse a \<cdot> y)) =  h y"
+        also from nz y have "a * (h (inverse a \<cdot> y)) =  h y"
           by simp
-	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
-	then show ?thesis by simp
+        finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
+        then show ?thesis by simp
       next
-	txt {* In the case @{text "a > 0"}, we use @{text "a\<^sub>2"}
+        txt {* In the case @{text "a > 0"}, we use @{text "a\<^sub>2"}
           with @{text ya} taken as @{text "y / a"}: *}
-	assume gz: "0 < a" then have nz: "a \<noteq> 0" by simp
-	from a2 ay
-	have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)" ..
-	with gz have "a * xi \<le>
+        assume gz: "0 < a" then have nz: "a \<noteq> 0" by simp
+        from a2 ay
+        have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)" ..
+        with gz have "a * xi \<le>
           a * (p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
           by simp
-	also have "\<dots> = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
-	  by (simp add: right_diff_distrib)
-	also from gz x0 y'
-	have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
+        also have "\<dots> = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
+          by (simp add: right_diff_distrib)
+        also from gz x0 y'
+        have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
           by (simp add: abs_homogenous)
-	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
+        also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
           by (simp add: add_mult_distrib1 mult_assoc [symmetric])
-	also from nz y have "a * h (inverse a \<cdot> y) = h y"
+        also from nz y have "a * h (inverse a \<cdot> y) = h y"
           by simp
-	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
-	then show ?thesis by simp
+        finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
+        then show ?thesis by simp
       qed
       also from x_rep have "\<dots> = p x" by (simp only:)
       finally show ?thesis .
--- a/src/HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -431,7 +431,7 @@
       from H x have "- x \<in> H" by (rule vectorspace.neg_closed)
       with r have "h (- x) \<le> p (- x)" ..
       also have "\<dots> = p x"
-	using `seminorm E p` `vectorspace E`
+        using `seminorm E p` `vectorspace E`
       proof (rule seminorm.minus)
         from x show "x \<in> E" ..
       qed
--- a/src/HOL/Hahn_Banach/Subspace.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hahn_Banach/Subspace.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -294,7 +294,7 @@
     proof
       fix x assume "x \<in> U + V"
       then obtain u v where "x = u + v" and
-	"u \<in> U" and "v \<in> V" ..
+        "u \<in> U" and "v \<in> V" ..
       then show "x \<in> E" by simp
     qed
     fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
@@ -305,16 +305,16 @@
       from y obtain uy vy where "y = uy + vy" and "uy \<in> U" and "vy \<in> V" ..
       ultimately
       have "ux + uy \<in> U"
-	and "vx + vy \<in> V"
-	and "x + y = (ux + uy) + (vx + vy)"
-	using x y by (simp_all add: add_ac)
+        and "vx + vy \<in> V"
+        and "x + y = (ux + uy) + (vx + vy)"
+        using x y by (simp_all add: add_ac)
       then show ?thesis ..
     qed
     fix a show "a \<cdot> x \<in> U + V"
     proof -
       from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
       then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
-	and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
+        and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
       then show ?thesis ..
     qed
   qed
@@ -404,8 +404,8 @@
       show "a2 \<cdot> x' \<in> lin x'" ..
       show "H \<inter> lin x' = {0}"
       proof
-	show "H \<inter> lin x' \<subseteq> {0}"
-	proof
+        show "H \<inter> lin x' \<subseteq> {0}"
+        proof
           fix x assume x: "x \<in> H \<inter> lin x'"
           then obtain a where xx': "x = a \<cdot> x'"
             by blast
@@ -421,13 +421,13 @@
             with `x' \<notin> H` show ?thesis by contradiction
           qed
           then show "x \<in> {0}" ..
-	qed
-	show "{0} \<subseteq> H \<inter> lin x'"
-	proof -
+        qed
+        show "{0} \<subseteq> H \<inter> lin x'"
+        proof -
           have "0 \<in> H" using `vectorspace E` ..
           moreover have "0 \<in> lin x'" using `x' \<in> E` ..
           ultimately show ?thesis by blast
-	qed
+        qed
       qed
       show "lin x' \<unlhd> E" using `x' \<in> E` ..
     qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq)
--- a/src/HOL/Hahn_Banach/Zorn_Lemma.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hahn_Banach/Zorn_Lemma.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -31,13 +31,13 @@
     proof cases
 
       txt {* If @{text c} is an empty chain, then every element in
-	@{text S} is an upper bound of @{text c}. *}
+        @{text S} is an upper bound of @{text c}. *}
 
       assume "c = {}"
       with aS show ?thesis by fast
 
       txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
-	bound of @{text c}, lying in @{text S}. *}
+        bound of @{text c}, lying in @{text S}. *}
 
     next
       assume "c \<noteq> {}"
@@ -47,7 +47,7 @@
         show "\<Union>c \<in> S"
         proof (rule r)
           from `c \<noteq> {}` show "\<exists>x. x \<in> c" by fast
-	  show "c \<in> chain S" by fact
+          show "c \<in> chain S" by fact
         qed
       qed
     qed
--- a/src/HOL/Hoare/SchorrWaite.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hoare/SchorrWaite.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -100,14 +100,14 @@
 apply (case_tac "a=q")
  apply auto
 done
-	      
+
 lemma restr_un: "((r \<union> s)|m) = (r|m) \<union> (s|m)"
   by (auto simp add:restr_def)
 
 lemma rel_upd3: "(a, b) \<notin> (r|(m(q := t))) \<Longrightarrow> (a,b) \<in> (r|m) \<Longrightarrow> a = q "
 apply (rule classical)
 apply (simp add:restr_def fun_upd_apply)
-done	
+done
 
 constdefs
   -- "A short form for the stack mapping function for List"
@@ -255,313 +255,313 @@
 
       show "(?ifB1 \<longrightarrow> (?ifB2 \<longrightarrow> (\<exists>stack.?popInv stack)) \<and> 
                             (\<not>?ifB2 \<longrightarrow> (\<exists>stack.?swInv stack)) ) \<and>
-	      (\<not>?ifB1 \<longrightarrow> (\<exists>stack.?puInv stack))"
+              (\<not>?ifB1 \<longrightarrow> (\<exists>stack.?puInv stack))"
       proof - 
-	{
-	  assume ifB1: "t = Null \<or> t^.m" and ifB2: "p^.c"
-	  from ifB1 whileB have pNotNull: "p \<noteq> Null" by auto
-	  then obtain addr_p where addr_p_eq: "p = Ref addr_p" by auto
-	  with i1 obtain stack_tl where stack_eq: "stack = (addr p) # stack_tl"
-	    by auto
-	  with i2 have m_addr_p: "p^.m" by auto
-	  have stackDist: "distinct (stack)" using i1 by (rule List_distinct)
-	  from stack_eq stackDist have p_notin_stack_tl: "addr p \<notin> set stack_tl" by simp
-	  let "?poI1\<and> ?poI2\<and> ?poI3\<and> ?poI4\<and> ?poI5\<and> ?poI6\<and> ?poI7" = "?popInv stack_tl"
-	  have "?popInv stack_tl"
-	  proof -
+        {
+          assume ifB1: "t = Null \<or> t^.m" and ifB2: "p^.c"
+          from ifB1 whileB have pNotNull: "p \<noteq> Null" by auto
+          then obtain addr_p where addr_p_eq: "p = Ref addr_p" by auto
+          with i1 obtain stack_tl where stack_eq: "stack = (addr p) # stack_tl"
+            by auto
+          with i2 have m_addr_p: "p^.m" by auto
+          have stackDist: "distinct (stack)" using i1 by (rule List_distinct)
+          from stack_eq stackDist have p_notin_stack_tl: "addr p \<notin> set stack_tl" by simp
+          let "?poI1\<and> ?poI2\<and> ?poI3\<and> ?poI4\<and> ?poI5\<and> ?poI6\<and> ?poI7" = "?popInv stack_tl"
+          have "?popInv stack_tl"
+          proof -
 
-	    -- {*List property is maintained:*}
-	    from i1 p_notin_stack_tl ifB2
-	    have poI1: "List (S c l (r(p \<rightarrow> t))) (p^.r) stack_tl" 
-	      by(simp add: addr_p_eq stack_eq, simp add: S_def)
+            -- {*List property is maintained:*}
+            from i1 p_notin_stack_tl ifB2
+            have poI1: "List (S c l (r(p \<rightarrow> t))) (p^.r) stack_tl" 
+              by(simp add: addr_p_eq stack_eq, simp add: S_def)
 
-	    moreover
-	    -- {*Everything on the stack is marked:*}
-	    from i2 have poI2: "\<forall> x \<in> set stack_tl. m x" by (simp add:stack_eq)
-	    moreover
+            moreover
+            -- {*Everything on the stack is marked:*}
+            from i2 have poI2: "\<forall> x \<in> set stack_tl. m x" by (simp add:stack_eq)
+            moreover
 
-	    -- {*Everything is still reachable:*}
-	    let "(R = reachable ?Ra ?A)" = "?I3"
-	    let "?Rb" = "(relS {l, r(p \<rightarrow> t)})"
-	    let "?B" = "{p, p^.r}"
-	    -- {*Our goal is @{text"R = reachable ?Rb ?B"}.*}
-	    have "?Ra\<^sup>* `` addrs ?A = ?Rb\<^sup>* `` addrs ?B" (is "?L = ?R")
-	    proof
-	      show "?L \<subseteq> ?R"
-	      proof (rule still_reachable)
-		show "addrs ?A \<subseteq> ?Rb\<^sup>* `` addrs ?B" by(fastsimp simp:addrs_def relS_def rel_def addr_p_eq 
-		     intro:oneStep_reachable Image_iff[THEN iffD2])
-		show "\<forall>(x,y) \<in> ?Ra-?Rb. y \<in> (?Rb\<^sup>* `` addrs ?B)" by (clarsimp simp:relS_def) 
-	             (fastsimp simp add:rel_def Image_iff addrs_def dest:rel_upd1)
-	      qed
-	      show "?R \<subseteq> ?L"
-	      proof (rule still_reachable)
-		show "addrs ?B \<subseteq> ?Ra\<^sup>* `` addrs ?A"
-		  by(fastsimp simp:addrs_def rel_defs addr_p_eq 
-		      intro:oneStep_reachable Image_iff[THEN iffD2])
-	      next
-		show "\<forall>(x, y)\<in>?Rb-?Ra. y\<in>(?Ra\<^sup>*``addrs ?A)"
-		  by (clarsimp simp:relS_def) 
-	             (fastsimp simp add:rel_def Image_iff addrs_def dest:rel_upd2)
-	      qed
-	    qed
-	    with i3 have poI3: "R = reachable ?Rb ?B"  by (simp add:reachable_def) 
-	    moreover
+            -- {*Everything is still reachable:*}
+            let "(R = reachable ?Ra ?A)" = "?I3"
+            let "?Rb" = "(relS {l, r(p \<rightarrow> t)})"
+            let "?B" = "{p, p^.r}"
+            -- {*Our goal is @{text"R = reachable ?Rb ?B"}.*}
+            have "?Ra\<^sup>* `` addrs ?A = ?Rb\<^sup>* `` addrs ?B" (is "?L = ?R")
+            proof
+              show "?L \<subseteq> ?R"
+              proof (rule still_reachable)
+                show "addrs ?A \<subseteq> ?Rb\<^sup>* `` addrs ?B" by(fastsimp simp:addrs_def relS_def rel_def addr_p_eq 
+                     intro:oneStep_reachable Image_iff[THEN iffD2])
+                show "\<forall>(x,y) \<in> ?Ra-?Rb. y \<in> (?Rb\<^sup>* `` addrs ?B)" by (clarsimp simp:relS_def) 
+                     (fastsimp simp add:rel_def Image_iff addrs_def dest:rel_upd1)
+              qed
+              show "?R \<subseteq> ?L"
+              proof (rule still_reachable)
+                show "addrs ?B \<subseteq> ?Ra\<^sup>* `` addrs ?A"
+                  by(fastsimp simp:addrs_def rel_defs addr_p_eq 
+                      intro:oneStep_reachable Image_iff[THEN iffD2])
+              next
+                show "\<forall>(x, y)\<in>?Rb-?Ra. y\<in>(?Ra\<^sup>*``addrs ?A)"
+                  by (clarsimp simp:relS_def) 
+                     (fastsimp simp add:rel_def Image_iff addrs_def dest:rel_upd2)
+              qed
+            qed
+            with i3 have poI3: "R = reachable ?Rb ?B"  by (simp add:reachable_def) 
+            moreover
 
-	    -- "If it is reachable and not marked, it is still reachable using..."
-	    let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Ra ?A"  =  ?I4	    
-	    let "?Rb" = "relS {l, r(p \<rightarrow> t)} | m"
-	    let "?B" = "{p} \<union> set (map (r(p \<rightarrow> t)) stack_tl)"
-	    -- {*Our goal is @{text"\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B"}.*}
-	    let ?T = "{t, p^.r}"
+            -- "If it is reachable and not marked, it is still reachable using..."
+            let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Ra ?A"  =  ?I4        
+            let "?Rb" = "relS {l, r(p \<rightarrow> t)} | m"
+            let "?B" = "{p} \<union> set (map (r(p \<rightarrow> t)) stack_tl)"
+            -- {*Our goal is @{text"\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B"}.*}
+            let ?T = "{t, p^.r}"
 
-	    have "?Ra\<^sup>* `` addrs ?A \<subseteq> ?Rb\<^sup>* `` (addrs ?B \<union> addrs ?T)"
-	    proof (rule still_reachable)
-	      have rewrite: "\<forall>s\<in>set stack_tl. (r(p \<rightarrow> t)) s = r s"
-		by (auto simp add:p_notin_stack_tl intro:fun_upd_other)	
-	      show "addrs ?A \<subseteq> ?Rb\<^sup>* `` (addrs ?B \<union> addrs ?T)"
-		by (fastsimp cong:map_cong simp:stack_eq addrs_def rewrite intro:self_reachable)
-	      show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``(addrs ?B \<union> addrs ?T))"
-		by (clarsimp simp:restr_def relS_def) 
-	          (fastsimp simp add:rel_def Image_iff addrs_def dest:rel_upd1)
- 	    qed
-	    -- "We now bring a term from the right to the left of the subset relation."
-	    hence subset: "?Ra\<^sup>* `` addrs ?A - ?Rb\<^sup>* `` addrs ?T \<subseteq> ?Rb\<^sup>* `` addrs ?B"
-	      by blast
-	    have poI4: "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B"
-	    proof (rule allI, rule impI)
-	      fix x
-	      assume a: "x \<in> R \<and> \<not> m x"
-	      -- {*First, a disjunction on @{term"p^.r"} used later in the proof*}
-	      have pDisj:"p^.r = Null \<or> (p^.r \<noteq> Null \<and> p^.r^.m)" using poI1 poI2 
-		by auto
-	      -- {*@{term x} belongs to the left hand side of @{thm[source] subset}:*}
-	      have incl: "x \<in> ?Ra\<^sup>*``addrs ?A" using  a i4 by (simp only:reachable_def, clarsimp)
-	      have excl: "x \<notin> ?Rb\<^sup>*`` addrs ?T" using pDisj ifB1 a by (auto simp add:addrs_def)
-	      -- {*And therefore also belongs to the right hand side of @{thm[source]subset},*}
-	      -- {*which corresponds to our goal.*}
-	      from incl excl subset  show "x \<in> reachable ?Rb ?B" by (auto simp add:reachable_def)
-	    qed
-	    moreover
+            have "?Ra\<^sup>* `` addrs ?A \<subseteq> ?Rb\<^sup>* `` (addrs ?B \<union> addrs ?T)"
+            proof (rule still_reachable)
+              have rewrite: "\<forall>s\<in>set stack_tl. (r(p \<rightarrow> t)) s = r s"
+                by (auto simp add:p_notin_stack_tl intro:fun_upd_other) 
+              show "addrs ?A \<subseteq> ?Rb\<^sup>* `` (addrs ?B \<union> addrs ?T)"
+                by (fastsimp cong:map_cong simp:stack_eq addrs_def rewrite intro:self_reachable)
+              show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``(addrs ?B \<union> addrs ?T))"
+                by (clarsimp simp:restr_def relS_def) 
+                  (fastsimp simp add:rel_def Image_iff addrs_def dest:rel_upd1)
+            qed
+            -- "We now bring a term from the right to the left of the subset relation."
+            hence subset: "?Ra\<^sup>* `` addrs ?A - ?Rb\<^sup>* `` addrs ?T \<subseteq> ?Rb\<^sup>* `` addrs ?B"
+              by blast
+            have poI4: "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B"
+            proof (rule allI, rule impI)
+              fix x
+              assume a: "x \<in> R \<and> \<not> m x"
+              -- {*First, a disjunction on @{term"p^.r"} used later in the proof*}
+              have pDisj:"p^.r = Null \<or> (p^.r \<noteq> Null \<and> p^.r^.m)" using poI1 poI2 
+                by auto
+              -- {*@{term x} belongs to the left hand side of @{thm[source] subset}:*}
+              have incl: "x \<in> ?Ra\<^sup>*``addrs ?A" using  a i4 by (simp only:reachable_def, clarsimp)
+              have excl: "x \<notin> ?Rb\<^sup>*`` addrs ?T" using pDisj ifB1 a by (auto simp add:addrs_def)
+              -- {*And therefore also belongs to the right hand side of @{thm[source]subset},*}
+              -- {*which corresponds to our goal.*}
+              from incl excl subset  show "x \<in> reachable ?Rb ?B" by (auto simp add:reachable_def)
+            qed
+            moreover
 
-	    -- "If it is marked, then it is reachable"
-	    from i5 have poI5: "\<forall>x. m x \<longrightarrow> x \<in> R" .
-	    moreover
+            -- "If it is marked, then it is reachable"
+            from i5 have poI5: "\<forall>x. m x \<longrightarrow> x \<in> R" .
+            moreover
 
-	    -- {*If it is not on the stack, then its @{term l} and @{term r} fields are unchanged*}
-	    from i7 i6 ifB2 
-	    have poI6: "\<forall>x. x \<notin> set stack_tl \<longrightarrow> (r(p \<rightarrow> t)) x = iR x \<and> l x = iL x" 
-	      by(auto simp: addr_p_eq stack_eq fun_upd_apply)
+            -- {*If it is not on the stack, then its @{term l} and @{term r} fields are unchanged*}
+            from i7 i6 ifB2 
+            have poI6: "\<forall>x. x \<notin> set stack_tl \<longrightarrow> (r(p \<rightarrow> t)) x = iR x \<and> l x = iL x" 
+              by(auto simp: addr_p_eq stack_eq fun_upd_apply)
 
-	    moreover
+            moreover
 
-	    -- {*If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed*}
-	    from p_notin_stack_tl i7 have poI7: "stkOk c l (r(p \<rightarrow> t)) iL iR p stack_tl"
-	      by (clarsimp simp:stack_eq addr_p_eq)
+            -- {*If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed*}
+            from p_notin_stack_tl i7 have poI7: "stkOk c l (r(p \<rightarrow> t)) iL iR p stack_tl"
+              by (clarsimp simp:stack_eq addr_p_eq)
 
-	    ultimately show "?popInv stack_tl" by simp
-	  qed
-	  hence "\<exists>stack. ?popInv stack" ..
-	}
-	moreover
+            ultimately show "?popInv stack_tl" by simp
+          qed
+          hence "\<exists>stack. ?popInv stack" ..
+        }
+        moreover
 
-	-- "Proofs of the Swing and Push arm follow."
-	-- "Since they are in principle simmilar to the Pop arm proof,"
-	-- "we show fewer comments and use frequent pattern matching."
-	{
-	  -- "Swing arm"
-	  assume ifB1: "?ifB1" and nifB2: "\<not>?ifB2"
-	  from ifB1 whileB have pNotNull: "p \<noteq> Null" by clarsimp
-	  then obtain addr_p where addr_p_eq: "p = Ref addr_p" by clarsimp
-	  with i1 obtain stack_tl where stack_eq: "stack = (addr p) # stack_tl" by clarsimp
-	  with i2 have m_addr_p: "p^.m" by clarsimp
-	  from stack_eq stackDist have p_notin_stack_tl: "(addr p) \<notin> set stack_tl"
-	    by simp
-	  let "?swI1\<and>?swI2\<and>?swI3\<and>?swI4\<and>?swI5\<and>?swI6\<and>?swI7" = "?swInv stack"
-	  have "?swInv stack"
-	  proof -
-	    
-	    -- {*List property is maintained:*}
-	    from i1 p_notin_stack_tl nifB2
-	    have swI1: "?swI1"
-	      by (simp add:addr_p_eq stack_eq, simp add:S_def)
-	    moreover
-	    
-	    -- {*Everything on the stack is marked:*}
-	    from i2
-	    have swI2: "?swI2" .
-	    moreover
-	    
-	    -- {*Everything is still reachable:*}
-	    let "R = reachable ?Ra ?A" = "?I3"
-	    let "R = reachable ?Rb ?B" = "?swI3"
-	    have "?Ra\<^sup>* `` addrs ?A = ?Rb\<^sup>* `` addrs ?B"
-	    proof (rule still_reachable_eq)
-	      show "addrs ?A \<subseteq> ?Rb\<^sup>* `` addrs ?B"
-		by(fastsimp simp:addrs_def rel_defs addr_p_eq intro:oneStep_reachable Image_iff[THEN iffD2])
-	    next
-	      show "addrs ?B \<subseteq> ?Ra\<^sup>* `` addrs ?A"
-		by(fastsimp simp:addrs_def rel_defs addr_p_eq intro:oneStep_reachable Image_iff[THEN iffD2])
-	    next
-	      show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``addrs ?B)"
-		by (clarsimp simp:relS_def) (fastsimp simp add:rel_def Image_iff addrs_def fun_upd_apply dest:rel_upd1)
-	    next
-	      show "\<forall>(x, y)\<in>?Rb-?Ra. y\<in>(?Ra\<^sup>*``addrs ?A)"
-		by (clarsimp simp:relS_def) (fastsimp simp add:rel_def Image_iff addrs_def fun_upd_apply dest:rel_upd2)
-	    qed
-	    with i3
-	    have swI3: "?swI3" by (simp add:reachable_def) 
-	    moreover
+        -- "Proofs of the Swing and Push arm follow."
+        -- "Since they are in principle simmilar to the Pop arm proof,"
+        -- "we show fewer comments and use frequent pattern matching."
+        {
+          -- "Swing arm"
+          assume ifB1: "?ifB1" and nifB2: "\<not>?ifB2"
+          from ifB1 whileB have pNotNull: "p \<noteq> Null" by clarsimp
+          then obtain addr_p where addr_p_eq: "p = Ref addr_p" by clarsimp
+          with i1 obtain stack_tl where stack_eq: "stack = (addr p) # stack_tl" by clarsimp
+          with i2 have m_addr_p: "p^.m" by clarsimp
+          from stack_eq stackDist have p_notin_stack_tl: "(addr p) \<notin> set stack_tl"
+            by simp
+          let "?swI1\<and>?swI2\<and>?swI3\<and>?swI4\<and>?swI5\<and>?swI6\<and>?swI7" = "?swInv stack"
+          have "?swInv stack"
+          proof -
+            
+            -- {*List property is maintained:*}
+            from i1 p_notin_stack_tl nifB2
+            have swI1: "?swI1"
+              by (simp add:addr_p_eq stack_eq, simp add:S_def)
+            moreover
+            
+            -- {*Everything on the stack is marked:*}
+            from i2
+            have swI2: "?swI2" .
+            moreover
+            
+            -- {*Everything is still reachable:*}
+            let "R = reachable ?Ra ?A" = "?I3"
+            let "R = reachable ?Rb ?B" = "?swI3"
+            have "?Ra\<^sup>* `` addrs ?A = ?Rb\<^sup>* `` addrs ?B"
+            proof (rule still_reachable_eq)
+              show "addrs ?A \<subseteq> ?Rb\<^sup>* `` addrs ?B"
+                by(fastsimp simp:addrs_def rel_defs addr_p_eq intro:oneStep_reachable Image_iff[THEN iffD2])
+            next
+              show "addrs ?B \<subseteq> ?Ra\<^sup>* `` addrs ?A"
+                by(fastsimp simp:addrs_def rel_defs addr_p_eq intro:oneStep_reachable Image_iff[THEN iffD2])
+            next
+              show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``addrs ?B)"
+                by (clarsimp simp:relS_def) (fastsimp simp add:rel_def Image_iff addrs_def fun_upd_apply dest:rel_upd1)
+            next
+              show "\<forall>(x, y)\<in>?Rb-?Ra. y\<in>(?Ra\<^sup>*``addrs ?A)"
+                by (clarsimp simp:relS_def) (fastsimp simp add:rel_def Image_iff addrs_def fun_upd_apply dest:rel_upd2)
+            qed
+            with i3
+            have swI3: "?swI3" by (simp add:reachable_def) 
+            moreover
 
-	    -- "If it is reachable and not marked, it is still reachable using..."
-	    let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Ra ?A" = ?I4
-	    let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B" = ?swI4
-	    let ?T = "{t}"
-	    have "?Ra\<^sup>*``addrs ?A \<subseteq> ?Rb\<^sup>*``(addrs ?B \<union> addrs ?T)"
-	    proof (rule still_reachable)
-	      have rewrite: "(\<forall>s\<in>set stack_tl. (r(addr p := l(addr p))) s = r s)"
-		by (auto simp add:p_notin_stack_tl intro:fun_upd_other)
-	      show "addrs ?A \<subseteq> ?Rb\<^sup>* `` (addrs ?B \<union> addrs ?T)"
-		by (fastsimp cong:map_cong simp:stack_eq addrs_def rewrite intro:self_reachable)
-	    next
-	      show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``(addrs ?B \<union> addrs ?T))"
-		by (clarsimp simp:relS_def restr_def) (fastsimp simp add:rel_def Image_iff addrs_def fun_upd_apply dest:rel_upd1)
-	    qed
-	    then have subset: "?Ra\<^sup>*``addrs ?A - ?Rb\<^sup>*``addrs ?T \<subseteq> ?Rb\<^sup>*``addrs ?B"
-	      by blast
-	    have ?swI4
-	    proof (rule allI, rule impI)
-	      fix x
-	      assume a: "x \<in> R \<and>\<not> m x"
-	      with i4 addr_p_eq stack_eq  have inc: "x \<in> ?Ra\<^sup>*``addrs ?A" 
-		by (simp only:reachable_def, clarsimp)
-	      with ifB1 a 
-	      have exc: "x \<notin> ?Rb\<^sup>*`` addrs ?T" 
-		by (auto simp add:addrs_def)
-	      from inc exc subset  show "x \<in> reachable ?Rb ?B" 
-		by (auto simp add:reachable_def)
-	    qed
-	    moreover
-	    
-	    -- "If it is marked, then it is reachable"
-	    from i5
-	    have "?swI5" .
-	    moreover
+            -- "If it is reachable and not marked, it is still reachable using..."
+            let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Ra ?A" = ?I4
+            let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B" = ?swI4
+            let ?T = "{t}"
+            have "?Ra\<^sup>*``addrs ?A \<subseteq> ?Rb\<^sup>*``(addrs ?B \<union> addrs ?T)"
+            proof (rule still_reachable)
+              have rewrite: "(\<forall>s\<in>set stack_tl. (r(addr p := l(addr p))) s = r s)"
+                by (auto simp add:p_notin_stack_tl intro:fun_upd_other)
+              show "addrs ?A \<subseteq> ?Rb\<^sup>* `` (addrs ?B \<union> addrs ?T)"
+                by (fastsimp cong:map_cong simp:stack_eq addrs_def rewrite intro:self_reachable)
+            next
+              show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``(addrs ?B \<union> addrs ?T))"
+                by (clarsimp simp:relS_def restr_def) (fastsimp simp add:rel_def Image_iff addrs_def fun_upd_apply dest:rel_upd1)
+            qed
+            then have subset: "?Ra\<^sup>*``addrs ?A - ?Rb\<^sup>*``addrs ?T \<subseteq> ?Rb\<^sup>*``addrs ?B"
+              by blast
+            have ?swI4
+            proof (rule allI, rule impI)
+              fix x
+              assume a: "x \<in> R \<and>\<not> m x"
+              with i4 addr_p_eq stack_eq  have inc: "x \<in> ?Ra\<^sup>*``addrs ?A" 
+                by (simp only:reachable_def, clarsimp)
+              with ifB1 a 
+              have exc: "x \<notin> ?Rb\<^sup>*`` addrs ?T" 
+                by (auto simp add:addrs_def)
+              from inc exc subset  show "x \<in> reachable ?Rb ?B" 
+                by (auto simp add:reachable_def)
+            qed
+            moreover
+            
+            -- "If it is marked, then it is reachable"
+            from i5
+            have "?swI5" .
+            moreover
 
-	    -- {*If it is not on the stack, then its @{term l} and @{term r} fields are unchanged*}
-	    from i6 stack_eq
-	    have "?swI6"
-	      by clarsimp 	    
-	    moreover
+            -- {*If it is not on the stack, then its @{term l} and @{term r} fields are unchanged*}
+            from i6 stack_eq
+            have "?swI6"
+              by clarsimp           
+            moreover
 
-	    -- {*If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed*}
-	    from stackDist i7 nifB2 
-	    have "?swI7"
-	      by (clarsimp simp:addr_p_eq stack_eq)
+            -- {*If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed*}
+            from stackDist i7 nifB2 
+            have "?swI7"
+              by (clarsimp simp:addr_p_eq stack_eq)
 
-	    ultimately show ?thesis by auto
-	  qed
-	  then have "\<exists>stack. ?swInv stack" by blast
-	}
-	moreover
+            ultimately show ?thesis by auto
+          qed
+          then have "\<exists>stack. ?swInv stack" by blast
+        }
+        moreover
 
-	{
-	  -- "Push arm"
-	  assume nifB1: "\<not>?ifB1"
-	  from nifB1 whileB have tNotNull: "t \<noteq> Null" by clarsimp
-	  then obtain addr_t where addr_t_eq: "t = Ref addr_t" by clarsimp
-	  with i1 obtain new_stack where new_stack_eq: "new_stack = (addr t) # stack" by clarsimp
-	  from tNotNull nifB1 have n_m_addr_t: "\<not> (t^.m)" by clarsimp
-	  with i2 have t_notin_stack: "(addr t) \<notin> set stack" by blast
-	  let "?puI1\<and>?puI2\<and>?puI3\<and>?puI4\<and>?puI5\<and>?puI6\<and>?puI7" = "?puInv new_stack"
-	  have "?puInv new_stack"
-	  proof -
-	    
-	    -- {*List property is maintained:*}
-	    from i1 t_notin_stack
-	    have puI1: "?puI1"
-	      by (simp add:addr_t_eq new_stack_eq, simp add:S_def)
-	    moreover
-	    
-	    -- {*Everything on the stack is marked:*}
-	    from i2
-	    have puI2: "?puI2" 
-	      by (simp add:new_stack_eq fun_upd_apply)
-	    moreover
-	    
-	    -- {*Everything is still reachable:*}
-	    let "R = reachable ?Ra ?A" = "?I3"
-	    let "R = reachable ?Rb ?B" = "?puI3"
-	    have "?Ra\<^sup>* `` addrs ?A = ?Rb\<^sup>* `` addrs ?B"
-	    proof (rule still_reachable_eq)
-	      show "addrs ?A \<subseteq> ?Rb\<^sup>* `` addrs ?B"
-		by(fastsimp simp:addrs_def rel_defs addr_t_eq intro:oneStep_reachable Image_iff[THEN iffD2])
-	    next
-	      show "addrs ?B \<subseteq> ?Ra\<^sup>* `` addrs ?A"
-		by(fastsimp simp:addrs_def rel_defs addr_t_eq intro:oneStep_reachable Image_iff[THEN iffD2])
-	    next
-	      show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``addrs ?B)"
-		by (clarsimp simp:relS_def) (fastsimp simp add:rel_def Image_iff addrs_def dest:rel_upd1)
-	    next
-	      show "\<forall>(x, y)\<in>?Rb-?Ra. y\<in>(?Ra\<^sup>*``addrs ?A)"
-		by (clarsimp simp:relS_def) (fastsimp simp add:rel_def Image_iff addrs_def fun_upd_apply dest:rel_upd2)
-	    qed
-	    with i3
-	    have puI3: "?puI3" by (simp add:reachable_def) 
-	    moreover
-	    
-	    -- "If it is reachable and not marked, it is still reachable using..."
-	    let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Ra ?A" = ?I4
-	    let "\<forall>x. x \<in> R \<and> \<not> ?new_m x \<longrightarrow> x \<in> reachable ?Rb ?B" = ?puI4
-	    let ?T = "{t}"
-	    have "?Ra\<^sup>*``addrs ?A \<subseteq> ?Rb\<^sup>*``(addrs ?B \<union> addrs ?T)"
-	    proof (rule still_reachable)
-	      show "addrs ?A \<subseteq> ?Rb\<^sup>* `` (addrs ?B \<union> addrs ?T)"
-		by (fastsimp simp:new_stack_eq addrs_def intro:self_reachable)
-	    next
-	      show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``(addrs ?B \<union> addrs ?T))"
-		by (clarsimp simp:relS_def new_stack_eq restr_un restr_upd) 
-	           (fastsimp simp add:rel_def Image_iff restr_def addrs_def fun_upd_apply addr_t_eq dest:rel_upd3)
-	    qed
-	    then have subset: "?Ra\<^sup>*``addrs ?A - ?Rb\<^sup>*``addrs ?T \<subseteq> ?Rb\<^sup>*``addrs ?B"
-	      by blast
-	    have ?puI4
-	    proof (rule allI, rule impI)
-	      fix x
-	      assume a: "x \<in> R \<and> \<not> ?new_m x"
-	      have xDisj: "x=(addr t) \<or> x\<noteq>(addr t)" by simp
-	      with i4 a have inc: "x \<in> ?Ra\<^sup>*``addrs ?A"
-		by (fastsimp simp:addr_t_eq addrs_def reachable_def intro:self_reachable)
-	      have exc: "x \<notin> ?Rb\<^sup>*`` addrs ?T"
-		using xDisj a n_m_addr_t
-		by (clarsimp simp add:addrs_def addr_t_eq) 
-	      from inc exc subset  show "x \<in> reachable ?Rb ?B" 
-		by (auto simp add:reachable_def)
-	    qed  
-	    moreover
-	    
-	    -- "If it is marked, then it is reachable"
-	    from i5
-	    have "?puI5"
-	      by (auto simp:addrs_def i3 reachable_def addr_t_eq fun_upd_apply intro:self_reachable)
-	    moreover
-	    
-	    -- {*If it is not on the stack, then its @{term l} and @{term r} fields are unchanged*}
-	    from i6 
-	    have "?puI6"
-	      by (simp add:new_stack_eq)
-	    moreover
+        {
+          -- "Push arm"
+          assume nifB1: "\<not>?ifB1"
+          from nifB1 whileB have tNotNull: "t \<noteq> Null" by clarsimp
+          then obtain addr_t where addr_t_eq: "t = Ref addr_t" by clarsimp
+          with i1 obtain new_stack where new_stack_eq: "new_stack = (addr t) # stack" by clarsimp
+          from tNotNull nifB1 have n_m_addr_t: "\<not> (t^.m)" by clarsimp
+          with i2 have t_notin_stack: "(addr t) \<notin> set stack" by blast
+          let "?puI1\<and>?puI2\<and>?puI3\<and>?puI4\<and>?puI5\<and>?puI6\<and>?puI7" = "?puInv new_stack"
+          have "?puInv new_stack"
+          proof -
+            
+            -- {*List property is maintained:*}
+            from i1 t_notin_stack
+            have puI1: "?puI1"
+              by (simp add:addr_t_eq new_stack_eq, simp add:S_def)
+            moreover
+            
+            -- {*Everything on the stack is marked:*}
+            from i2
+            have puI2: "?puI2" 
+              by (simp add:new_stack_eq fun_upd_apply)
+            moreover
+            
+            -- {*Everything is still reachable:*}
+            let "R = reachable ?Ra ?A" = "?I3"
+            let "R = reachable ?Rb ?B" = "?puI3"
+            have "?Ra\<^sup>* `` addrs ?A = ?Rb\<^sup>* `` addrs ?B"
+            proof (rule still_reachable_eq)
+              show "addrs ?A \<subseteq> ?Rb\<^sup>* `` addrs ?B"
+                by(fastsimp simp:addrs_def rel_defs addr_t_eq intro:oneStep_reachable Image_iff[THEN iffD2])
+            next
+              show "addrs ?B \<subseteq> ?Ra\<^sup>* `` addrs ?A"
+                by(fastsimp simp:addrs_def rel_defs addr_t_eq intro:oneStep_reachable Image_iff[THEN iffD2])
+            next
+              show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``addrs ?B)"
+                by (clarsimp simp:relS_def) (fastsimp simp add:rel_def Image_iff addrs_def dest:rel_upd1)
+            next
+              show "\<forall>(x, y)\<in>?Rb-?Ra. y\<in>(?Ra\<^sup>*``addrs ?A)"
+                by (clarsimp simp:relS_def) (fastsimp simp add:rel_def Image_iff addrs_def fun_upd_apply dest:rel_upd2)
+            qed
+            with i3
+            have puI3: "?puI3" by (simp add:reachable_def) 
+            moreover
+            
+            -- "If it is reachable and not marked, it is still reachable using..."
+            let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Ra ?A" = ?I4
+            let "\<forall>x. x \<in> R \<and> \<not> ?new_m x \<longrightarrow> x \<in> reachable ?Rb ?B" = ?puI4
+            let ?T = "{t}"
+            have "?Ra\<^sup>*``addrs ?A \<subseteq> ?Rb\<^sup>*``(addrs ?B \<union> addrs ?T)"
+            proof (rule still_reachable)
+              show "addrs ?A \<subseteq> ?Rb\<^sup>* `` (addrs ?B \<union> addrs ?T)"
+                by (fastsimp simp:new_stack_eq addrs_def intro:self_reachable)
+            next
+              show "\<forall>(x, y)\<in>?Ra-?Rb. y\<in>(?Rb\<^sup>*``(addrs ?B \<union> addrs ?T))"
+                by (clarsimp simp:relS_def new_stack_eq restr_un restr_upd) 
+                   (fastsimp simp add:rel_def Image_iff restr_def addrs_def fun_upd_apply addr_t_eq dest:rel_upd3)
+            qed
+            then have subset: "?Ra\<^sup>*``addrs ?A - ?Rb\<^sup>*``addrs ?T \<subseteq> ?Rb\<^sup>*``addrs ?B"
+              by blast
+            have ?puI4
+            proof (rule allI, rule impI)
+              fix x
+              assume a: "x \<in> R \<and> \<not> ?new_m x"
+              have xDisj: "x=(addr t) \<or> x\<noteq>(addr t)" by simp
+              with i4 a have inc: "x \<in> ?Ra\<^sup>*``addrs ?A"
+                by (fastsimp simp:addr_t_eq addrs_def reachable_def intro:self_reachable)
+              have exc: "x \<notin> ?Rb\<^sup>*`` addrs ?T"
+                using xDisj a n_m_addr_t
+                by (clarsimp simp add:addrs_def addr_t_eq) 
+              from inc exc subset  show "x \<in> reachable ?Rb ?B" 
+                by (auto simp add:reachable_def)
+            qed  
+            moreover
+            
+            -- "If it is marked, then it is reachable"
+            from i5
+            have "?puI5"
+              by (auto simp:addrs_def i3 reachable_def addr_t_eq fun_upd_apply intro:self_reachable)
+            moreover
+            
+            -- {*If it is not on the stack, then its @{term l} and @{term r} fields are unchanged*}
+            from i6 
+            have "?puI6"
+              by (simp add:new_stack_eq)
+            moreover
 
-	    -- {*If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed*}
-	    from stackDist i6 t_notin_stack i7
-	    have "?puI7" by (clarsimp simp:addr_t_eq new_stack_eq)
+            -- {*If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed*}
+            from stackDist i6 t_notin_stack i7
+            have "?puI7" by (clarsimp simp:addr_t_eq new_stack_eq)
 
-	    ultimately show ?thesis by auto
-	  qed
-	  then have "\<exists>stack. ?puInv stack" by blast
+            ultimately show ?thesis by auto
+          qed
+          then have "\<exists>stack. ?puInv stack" by blast
 
-	}
-	ultimately show ?thesis by blast
+        }
+        ultimately show ?thesis by blast
       qed
     }
   qed
--- a/src/HOL/Hoare_Parallel/Graph.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hoare_Parallel/Graph.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -29,7 +29,7 @@
   Reach :: "edges \<Rightarrow> nat set"
   "Reach E \<equiv> {x. (\<exists>path. 1<length path \<and> path!(length path - 1)\<in>Roots \<and> x=path!0
               \<and> (\<forall>i<length path - 1. (\<exists>j<length E. E!j=(path!(i+1), path!i))))
-	      \<or> x\<in>Roots}"
+              \<or> x\<in>Roots}"
 
 text{* Reach: the set of reachable nodes is the set of Roots together with the
 nodes reachable from some Root by a path represented by a list of
--- a/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -246,7 +246,7 @@
    ELSE .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
          \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
          \<and> \<acute>Mul_PBInv \<and> \<acute>ind<length \<acute>E}.
-	 \<langle>IF \<acute>M!(fst (\<acute>E!\<acute>ind))\<noteq>Black THEN \<acute>ind:=\<acute>ind+1 FI\<rangle> FI
+         \<langle>IF \<acute>M!(fst (\<acute>E!\<acute>ind))\<noteq>Black THEN \<acute>ind:=\<acute>ind+1 FI\<rangle> FI
  OD"
 
 lemma Mul_Propagate_Black: 
@@ -318,7 +318,7 @@
           \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M  
           \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv \<acute>ind 
           \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))
-	  \<and> \<acute>q<n+1 \<and> \<acute>ind\<le>length \<acute>M}.
+          \<and> \<acute>q<n+1 \<and> \<acute>ind\<le>length \<acute>M}.
   DO .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
        \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
        \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv \<acute>ind 
--- a/src/HOL/Hoare_Parallel/OG_Hoare.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -57,7 +57,7 @@
 
 
 | Parallel: "\<lbrakk> \<forall>i<length Ts. \<exists>c q. Ts!i = (Some c, q) \<and> \<turnstile> c q; interfree Ts \<rbrakk>
-	   \<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts!i)))) 
+           \<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts!i)))) 
                      Parallel Ts 
                   (\<Inter>i\<in>{i. i<length Ts}. post(Ts!i))"
 
@@ -70,7 +70,7 @@
 | While:  "\<lbrakk> \<parallel>- (p \<inter> b) c p \<rbrakk> \<Longrightarrow> \<parallel>- p (While b i c) (p \<inter> -b)"
 
 | Conseq: "\<lbrakk> p' \<subseteq> p; \<parallel>- p c q ; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<parallel>- p' c q'"
-					    
+
 section {* Soundness *}
 (* In the version Isabelle-10-Sep-1999: HOL: The THEN and ELSE
 parts of conditional expressions (if P then x else y) are no longer
--- a/src/HOL/Hoare_Parallel/OG_Tran.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Hoare_Parallel/OG_Tran.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -51,7 +51,7 @@
                          (Some (AnnSeq c (AnnWhile i b i c)), s)"
 
 | AnnAwait: "\<lbrakk> s \<in> b; atom_com c; (c, s) -P*\<rightarrow> (Parallel [], t) \<rbrakk> \<Longrightarrow>
-	           (Some (AnnAwait r b c), s) -1\<rightarrow> (None, t)" 
+                   (Some (AnnAwait r b c), s) -1\<rightarrow> (None, t)" 
 
 | Parallel: "\<lbrakk> i<length Ts; Ts!i = (Some c, q); (Some c, s) -1\<rightarrow> (r, t) \<rbrakk>
               \<Longrightarrow> (Parallel Ts, s) -P1\<rightarrow> (Parallel (Ts [i:=(r, q)]), t)"
--- a/src/HOL/IMP/Compiler.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/IMP/Compiler.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/IMP/Compiler.thy
-    ID:         $Id$
     Author:     Tobias Nipkow, TUM
     Copyright   1996 TUM
 *)
@@ -262,32 +261,32 @@
       assume H: "\<langle>?W,[],s\<rangle> -n\<rightarrow> \<langle>[],rev ?W,t\<rangle>"
       show "\<langle>?w,s\<rangle> \<longrightarrow>\<^sub>c t"
       proof cases
-	assume b: "b s"
-	then obtain m where m: "n = Suc m"
+        assume b: "b s"
+        then obtain m where m: "n = Suc m"
           and "\<langle>?C @ [?j2],[?j1],s\<rangle> -m\<rightarrow> \<langle>[],rev ?W,t\<rangle>"
-	  using H by fastsimp
-	then obtain r n1 n2 where n1: "\<langle>?C,[],s\<rangle> -n1\<rightarrow> \<langle>[],rev ?C,r\<rangle>"
+          using H by fastsimp
+        then obtain r n1 n2 where n1: "\<langle>?C,[],s\<rangle> -n1\<rightarrow> \<langle>[],rev ?C,r\<rangle>"
           and n2: "\<langle>[?j2],rev ?C @ [?j1],r\<rangle> -n2\<rightarrow> \<langle>[],rev ?W,t\<rangle>"
           and n12: "m = n1+n2"
-	  using execn_decomp[of _ "[?j2]"]
-	  by(simp del: execn_simp) fast
-	have n2n: "n2 - 1 < n" using m n12 by arith
-	note b
-	moreover
-	{ from n1 have "\<langle>?C,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C,r\<rangle>"
-	    by (simp add:rtrancl_is_UN_rel_pow) fast
-	  hence "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c r" by(rule IHc)
-	}
-	moreover
-	{ have "n2 - 1 < n" using m n12 by arith
-	  moreover from n2 have "\<langle>?W,[],r\<rangle> -n2- 1\<rightarrow> \<langle>[],rev ?W,t\<rangle>" by fastsimp
-	  ultimately have "\<langle>?w,r\<rangle> \<longrightarrow>\<^sub>c t" by(rule IHm)
-	}
-	ultimately show ?thesis ..
+          using execn_decomp[of _ "[?j2]"]
+          by(simp del: execn_simp) fast
+        have n2n: "n2 - 1 < n" using m n12 by arith
+        note b
+        moreover
+        { from n1 have "\<langle>?C,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C,r\<rangle>"
+            by (simp add:rtrancl_is_UN_rel_pow) fast
+          hence "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c r" by(rule IHc)
+        }
+        moreover
+        { have "n2 - 1 < n" using m n12 by arith
+          moreover from n2 have "\<langle>?W,[],r\<rangle> -n2- 1\<rightarrow> \<langle>[],rev ?W,t\<rangle>" by fastsimp
+          ultimately have "\<langle>?w,r\<rangle> \<longrightarrow>\<^sub>c t" by(rule IHm)
+        }
+        ultimately show ?thesis ..
       next
-	assume b: "\<not> b s"
-	hence "t = s" using H by simp
-	with b show ?thesis by simp
+        assume b: "\<not> b s"
+        hence "t = s" using H by simp
+        with b show ?thesis by simp
       qed
     qed
   }
--- a/src/HOL/IMP/Live.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/IMP/Live.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -93,8 +93,8 @@
       then obtain t'' where "\<langle>c,t\<rangle> \<longrightarrow>\<^sub>c t''" and "\<langle>While b c,t''\<rangle> \<longrightarrow>\<^sub>c t'"
         using WhileTrue(6,7) by auto
       have "\<forall>x\<in>Dep b \<union> A \<union> L c A. s'' x = t'' x"
-	using IH(1)[OF _ `\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s''` `\<langle>c,t\<rangle> \<longrightarrow>\<^sub>c t''`] WhileTrue(6,8)
-	by (auto simp:L_gen_kill)
+        using IH(1)[OF _ `\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s''` `\<langle>c,t\<rangle> \<longrightarrow>\<^sub>c t''`] WhileTrue(6,8)
+        by (auto simp:L_gen_kill)
       moreover then have "\<forall>x\<in>L (While b c) A. s'' x = t'' x" by auto
       ultimately show ?case using WhileTrue(5,6) `\<langle>While b c,t''\<rangle> \<longrightarrow>\<^sub>c t'` by metis
     qed auto }
@@ -155,14 +155,14 @@
       have "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s''" using WhileTrue(2,6) by simp
       have "b t" using WhileTrue by (simp add: ball_Un)(blast dest:dep_on)
       then obtain t'' where tt'': "\<langle>bury c (Dep b \<union> A \<union> L c A),t\<rangle> \<longrightarrow>\<^sub>c t''"
-	and "\<langle>bury (While b c) A,t''\<rangle> \<longrightarrow>\<^sub>c t'"
+        and "\<langle>bury (While b c) A,t''\<rangle> \<longrightarrow>\<^sub>c t'"
         using WhileTrue(6,7) by auto
       have "\<forall>x\<in>Dep b \<union> A \<union> L c A. s'' x = t'' x"
-	using IH(1)[OF _ `\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s''` tt''] WhileTrue(6,8)
-	by (auto simp:L_gen_kill)
+        using IH(1)[OF _ `\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s''` tt''] WhileTrue(6,8)
+        by (auto simp:L_gen_kill)
       moreover then have "\<forall>x\<in>L (While b c) A. s'' x = t'' x" by auto
       ultimately show ?case
-	using WhileTrue(5,6) `\<langle>bury (While b c) A,t''\<rangle> \<longrightarrow>\<^sub>c t'` by metis
+        using WhileTrue(5,6) `\<langle>bury (While b c) A,t''\<rangle> \<longrightarrow>\<^sub>c t'` by metis
     qed auto }
   from this[OF IH(3) _ IH(4,2)] show ?case by metis
 qed
--- a/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -153,8 +153,8 @@
         by (elim crelE crel_nth crel_if crel_return) auto
       from swp rec_condition have
         "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array a h ! i = get_array a h1 ! i"
-	unfolding swap_def
-	by (elim crelE crel_nth crel_upd crel_return) auto
+        unfolding swap_def
+        by (elim crelE crel_nth crel_upd crel_return) auto
       with 1(2) [OF rec_condition False rec2] show ?thesis by fastsimp
     qed
   qed
@@ -190,10 +190,10 @@
         unfolding part1.simps[of a l r p]
         by (elim crelE crel_nth crel_if crel_return) auto
       from True part_outer_remains[OF rec1] have a_l: "get_array a h' ! l \<le> p"
-	by fastsimp
+        by fastsimp
       have "\<forall>i. (l \<le> i = (l = i \<or> Suc l \<le> i))" by arith
       with 1(1)[OF False True rec1] a_l show ?thesis
-	by auto
+        by auto
     next
       case False
       with lr cr
@@ -202,13 +202,13 @@
         unfolding part1.simps[of a l r p]
         by (elim crelE crel_nth crel_if crel_return) auto
       from swp False have "get_array a h1 ! r \<ge> p"
-	unfolding swap_def
-	by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return)
+        unfolding swap_def
+        by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return)
       with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \<ge> p"
-	by fastsimp
+        by fastsimp
       have "\<forall>i. (i \<le> r = (i = r \<or> i \<le> r - 1))" by arith
       with 1(2)[OF lr False rec2] a_r show ?thesis
-	by auto
+        by auto
     qed
   qed
 qed
@@ -315,7 +315,7 @@
       with part_partitions[OF part] right_remains True
       have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
       with i_props h'_def in_bounds have "get_array a h' ! i \<le> get_array a h' ! rs"
-	unfolding Heap.upd_def Heap.length_def by simp
+        unfolding Heap.upd_def Heap.length_def by simp
     }
     moreover
     {
@@ -325,24 +325,24 @@
       hence "(rs < i \<and> i \<le> r - 1) \<or> (rs < i \<and> i = r)" by arith
       hence "get_array a h' ! rs \<le> get_array a h' ! i"
       proof
-	assume i_is: "rs < i \<and> i \<le> r - 1"
-	with swap_length_remains in_bounds middle_in_bounds rs_equals
-	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-	from part_partitions[OF part] rs_equals right_remains i_is
-	have "get_array a h' ! rs \<le> get_array a h1 ! i"
-	  by fastsimp
-	with i_props h'_def show ?thesis by fastsimp
+        assume i_is: "rs < i \<and> i \<le> r - 1"
+        with swap_length_remains in_bounds middle_in_bounds rs_equals
+        have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+        from part_partitions[OF part] rs_equals right_remains i_is
+        have "get_array a h' ! rs \<le> get_array a h1 ! i"
+          by fastsimp
+        with i_props h'_def show ?thesis by fastsimp
       next
-	assume i_is: "rs < i \<and> i = r"
-	with rs_equals have "Suc middle \<noteq> r" by arith
-	with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith
-	with part_partitions[OF part] right_remains 
-	have "get_array a h' ! rs \<le> get_array a h1 ! (Suc middle)"
-	  by fastsimp
-	with i_is True rs_equals right_remains h'_def
-	show ?thesis using in_bounds
-	  unfolding Heap.upd_def Heap.length_def
-	  by auto
+        assume i_is: "rs < i \<and> i = r"
+        with rs_equals have "Suc middle \<noteq> r" by arith
+        with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith
+        with part_partitions[OF part] right_remains 
+        have "get_array a h' ! rs \<le> get_array a h1 ! (Suc middle)"
+          by fastsimp
+        with i_is True rs_equals right_remains h'_def
+        show ?thesis using in_bounds
+          unfolding Heap.upd_def Heap.length_def
+          by auto
       qed
     }
     ultimately show ?thesis by auto
@@ -357,7 +357,7 @@
       from part_partitions[OF part] rs_equals right_remains i_is_left
       have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
       with i_props h'_def have "get_array a h' ! i \<le> get_array a h' ! rs"
-	unfolding Heap.upd_def by simp
+        unfolding Heap.upd_def by simp
     }
     moreover
     {
@@ -366,19 +366,19 @@
       hence "(rs < i \<and> i \<le> r - 1) \<or> i = r" by arith
       hence "get_array a h' ! rs \<le> get_array a h' ! i"
       proof
-	assume i_is: "rs < i \<and> i \<le> r - 1"
-	with swap_length_remains in_bounds middle_in_bounds rs_equals
-	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-	from part_partitions[OF part] rs_equals right_remains i_is
-	have "get_array a h' ! rs \<le> get_array a h1 ! i"
-	  by fastsimp
-	with i_props h'_def show ?thesis by fastsimp
+        assume i_is: "rs < i \<and> i \<le> r - 1"
+        with swap_length_remains in_bounds middle_in_bounds rs_equals
+        have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+        from part_partitions[OF part] rs_equals right_remains i_is
+        have "get_array a h' ! rs \<le> get_array a h1 ! i"
+          by fastsimp
+        with i_props h'_def show ?thesis by fastsimp
       next
-	assume i_is: "i = r"
-	from i_is False rs_equals right_remains h'_def
-	show ?thesis using in_bounds
-	  unfolding Heap.upd_def Heap.length_def
-	  by auto
+        assume i_is: "i = r"
+        from i_is False rs_equals right_remains h'_def
+        show ?thesis using in_bounds
+          unfolding Heap.upd_def Heap.length_def
+          by auto
       qed
     }
     ultimately
@@ -507,7 +507,7 @@
       have IH1: "sorted (subarray l p a h2)"  by (cases p, auto simp add: subarray_Nil)
       from quicksort_outer_remains [OF qs2] length_remains
       have left_subarray_remains: "subarray l p a h2 = subarray l p a h'"
-	by (simp add: subarray_eq_samelength_iff)
+        by (simp add: subarray_eq_samelength_iff)
       with IH1 have IH1': "sorted (subarray l p a h')" by simp
       from 1(2)[OF True pivot qs2] pivot 1(5) length_remains
       have IH2: "sorted (subarray (p + 1) (r + 1) a h')"
@@ -520,14 +520,14 @@
       from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True
         length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"]
       have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)"
-	unfolding Heap.length_def subarray_def by (cases p, auto)
+        unfolding Heap.length_def subarray_def by (cases p, auto)
       with left_subarray_remains part_conds1 pivot_unchanged
       have part_conds2': "\<forall>j. j \<in> set (subarray l p a h') \<longrightarrow> j \<le> get_array a h' ! p"
         by (simp, subst set_of_multiset_of[symmetric], simp)
           (* -- These steps are the analogous for the right sublist \<dots> *)
       from quicksort_outer_remains [OF qs1] length_remains
       have right_subarray_remains: "subarray (p + 1) (r + 1) a h1 = subarray (p + 1) (r + 1) a h2"
-	by (auto simp add: subarray_eq_samelength_iff)
+        by (auto simp add: subarray_eq_samelength_iff)
       from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True
         length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"]
       have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)"
@@ -538,11 +538,11 @@
           (* -- Thirdly and finally, we show that the array is sorted
           following from the facts above. *)
       from True pivot 1(5) length_remains have "subarray l (r + 1) a h' = subarray l p a h' @ [get_array a h' ! p] @ subarray (p + 1) (r + 1) a h'"
-	by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil)
+        by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil)
       with IH1' IH2 part_conds1' part_conds2' pivot have ?thesis
-	unfolding subarray_def
-	apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv)
-	by (auto simp add: set_sublist' dest: le_trans [of _ "get_array a h' ! p"])
+        unfolding subarray_def
+        apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv)
+        by (auto simp add: set_sublist' dest: le_trans [of _ "get_array a h' ! p"])
     }
     with True cr show ?thesis
       unfolding quicksort.simps [of a l r]
--- a/src/HOL/Imperative_HOL/ex/Sublist.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Imperative_HOL/ex/Sublist.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -360,42 +360,42 @@
       case 0 note b = this
       show ?thesis
       proof (cases ys)
-	case Nil  with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3)
+        case Nil  with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3)
       next
-	case (Cons y ys)
-	show ?thesis
-	proof (cases j)
-	  case 0 with a b Cons.prems show ?thesis by simp
-	next
-	  case (Suc j') with a b Cons.prems Cons show ?thesis 
-	    apply -
-	    apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto)
-	    done
-	qed
+        case (Cons y ys)
+        show ?thesis
+        proof (cases j)
+          case 0 with a b Cons.prems show ?thesis by simp
+        next
+          case (Suc j') with a b Cons.prems Cons show ?thesis 
+            apply -
+            apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto)
+            done
+        qed
       qed
     next
       case (Suc n')
       show ?thesis
       proof (cases ys)
-	case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3)
+        case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3)
       next
-	case (Cons y ys) with Suc a Cons.prems show ?thesis
-	  apply -
-	  apply simp
-	  apply (cases j)
-	  apply simp
-	  apply (cases i)
-	  apply simp
-	  apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"])
-	  apply simp
-	  apply simp
-	  apply simp
-	  apply simp
-	  apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"])
-	  apply simp
-	  apply simp
-	  apply simp
-	  done
+        case (Cons y ys) with Suc a Cons.prems show ?thesis
+          apply -
+          apply simp
+          apply (cases j)
+          apply simp
+          apply (cases i)
+          apply simp
+          apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"])
+          apply simp
+          apply simp
+          apply simp
+          apply simp
+          apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"])
+          apply simp
+          apply simp
+          apply simp
+          done
       qed
     qed
   qed
--- a/src/HOL/Import/HOL4Compat.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/HOL4Compat.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -166,19 +166,19 @@
       assume ind: "!q. q + k = m \<longrightarrow> P q"
       show "!q. Suc (q + k) = m \<longrightarrow> P q"
       proof (rule+)
-	fix q
-	assume "Suc (q + k) = m"
-	hence "(Suc q) + k = m"
-	  by simp
-	with ind
-	have psq: "P (Suc q)"
-	  by simp
-	from alln
-	have "P (Suc q) --> P q"
-	  ..
-	with psq
-	show "P q"
-	  by simp
+        fix q
+        assume "Suc (q + k) = m"
+        hence "(Suc q) + k = m"
+          by simp
+        with ind
+        have psq: "P (Suc q)"
+          by simp
+        from alln
+        have "P (Suc q) --> P q"
+          ..
+        with psq
+        show "P q"
+          by simp
       qed
     qed
   qed
@@ -457,9 +457,9 @@
       assume "P x"
       from allx
       have "P x \<longrightarrow> 0 < x"
-	..
+        ..
       thus "0 < x"
-	by (simp add: prems)
+        by (simp add: prems)
     qed
   next
     from px
--- a/src/HOL/Import/hol4rews.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/hol4rews.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -7,7 +7,7 @@
 type holthm = (term * term) list * thm
 
 datatype ImportStatus =
-	 NoImport
+         NoImport
        | Generating of string
        | Replaying of string
 
@@ -81,10 +81,10 @@
 
 fun set_segment thyname segname thy =
     let
-	val imps = HOL4Imports.get thy
-	val imps' = Symtab.update_new (thyname,segname) imps
+        val imps = HOL4Imports.get thy
+        val imps' = Symtab.update_new (thyname,segname) imps
     in
-	HOL4Imports.put imps' thy
+        HOL4Imports.put imps' thy
     end
 
 structure HOL4CMoves = TheoryDataFun
@@ -176,42 +176,42 @@
 in
 fun add_hol4_rewrite (Context.Theory thy, th) =
     let
-	val _ = message "Adding HOL4 rewrite"
-	val th1 = th RS eq_reflection
-	val current_rews = HOL4Rewrites.get thy
-	val new_rews = insert Thm.eq_thm th1 current_rews
-	val updated_thy  = HOL4Rewrites.put new_rews thy
+        val _ = message "Adding HOL4 rewrite"
+        val th1 = th RS eq_reflection
+        val current_rews = HOL4Rewrites.get thy
+        val new_rews = insert Thm.eq_thm th1 current_rews
+        val updated_thy  = HOL4Rewrites.put new_rews thy
     in
-	(Context.Theory updated_thy,th1)
+        (Context.Theory updated_thy,th1)
     end
   | add_hol4_rewrite (context, th) = (context, th RS eq_reflection);
 end
 
 fun ignore_hol4 bthy bthm thy =
     let
-	val _ = message ("Ignoring " ^ bthy ^ "." ^ bthm)
-	val curmaps = HOL4Maps.get thy
-	val newmaps = StringPair.update_new ((bthy,bthm),NONE) curmaps
-	val upd_thy = HOL4Maps.put newmaps thy
+        val _ = message ("Ignoring " ^ bthy ^ "." ^ bthm)
+        val curmaps = HOL4Maps.get thy
+        val newmaps = StringPair.update_new ((bthy,bthm),NONE) curmaps
+        val upd_thy = HOL4Maps.put newmaps thy
     in
-	upd_thy
+        upd_thy
     end
 
 val opt_get_output_thy = #2 o HOL4Dump.get
 
 fun get_output_thy thy =
     case #2 (HOL4Dump.get thy) of
-	"" => error "No theory file being output"
+        "" => error "No theory file being output"
       | thyname => thyname
 
 val get_output_dir = #1 o HOL4Dump.get
 
 fun add_hol4_move bef aft thy =
     let
-	val curmoves = HOL4Moves.get thy
-	val newmoves = Symtab.update_new (bef, aft) curmoves
+        val curmoves = HOL4Moves.get thy
+        val newmoves = Symtab.update_new (bef, aft) curmoves
     in
-	HOL4Moves.put newmoves thy
+        HOL4Moves.put newmoves thy
     end
 
 fun get_hol4_move bef thy =
@@ -219,21 +219,21 @@
 
 fun follow_name thmname thy =
     let
-	val moves = HOL4Moves.get thy
-	fun F thmname =
-	    case Symtab.lookup moves thmname of
-		SOME name => F name
-	      | NONE => thmname
+        val moves = HOL4Moves.get thy
+        fun F thmname =
+            case Symtab.lookup moves thmname of
+                SOME name => F name
+              | NONE => thmname
     in
-	F thmname
+        F thmname
     end
 
 fun add_hol4_cmove bef aft thy =
     let
-	val curmoves = HOL4CMoves.get thy
-	val newmoves = Symtab.update_new (bef, aft) curmoves
+        val curmoves = HOL4CMoves.get thy
+        val newmoves = Symtab.update_new (bef, aft) curmoves
     in
-	HOL4CMoves.put newmoves thy
+        HOL4CMoves.put newmoves thy
     end
 
 fun get_hol4_cmove bef thy =
@@ -241,128 +241,128 @@
 
 fun follow_cname thmname thy =
     let
-	val moves = HOL4CMoves.get thy
-	fun F thmname =
-	    case Symtab.lookup moves thmname of
-		SOME name => F name
-	      | NONE => thmname
+        val moves = HOL4CMoves.get thy
+        fun F thmname =
+            case Symtab.lookup moves thmname of
+                SOME name => F name
+              | NONE => thmname
     in
-	F thmname
+        F thmname
     end
 
 fun add_hol4_mapping bthy bthm isathm thy =
-    let	
+    let 
        (* val _ = writeln ("Before follow_name: "^isathm)  *)
       val isathm = follow_name isathm thy
        (* val _ = writeln ("Adding theorem map: " ^ bthy ^ "." ^ bthm ^ " --> " ^ isathm)*)
-	val curmaps = HOL4Maps.get thy
-	val newmaps = StringPair.update_new ((bthy,bthm),SOME isathm) curmaps
-	val upd_thy = HOL4Maps.put newmaps thy
+        val curmaps = HOL4Maps.get thy
+        val newmaps = StringPair.update_new ((bthy,bthm),SOME isathm) curmaps
+        val upd_thy = HOL4Maps.put newmaps thy
     in
-	upd_thy
+        upd_thy
     end;
 
 fun get_hol4_type_mapping bthy tycon thy =
     let
-	val maps = HOL4TypeMaps.get thy
+        val maps = HOL4TypeMaps.get thy
     in
-	StringPair.lookup maps (bthy,tycon)
+        StringPair.lookup maps (bthy,tycon)
     end
 
 fun get_hol4_mapping bthy bthm thy =
     let
-	val curmaps = HOL4Maps.get thy
+        val curmaps = HOL4Maps.get thy
     in
-	StringPair.lookup curmaps (bthy,bthm)
+        StringPair.lookup curmaps (bthy,bthm)
     end;
 
 fun add_hol4_const_mapping bthy bconst internal isaconst thy =
     let
-	val thy = case opt_get_output_thy thy of
-		      "" => thy
-		    | output_thy => if internal
-				    then add_hol4_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
-				    else thy
-	val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
-	val curmaps = HOL4ConstMaps.get thy
-	val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,NONE)) curmaps
-	val upd_thy = HOL4ConstMaps.put newmaps thy
+        val thy = case opt_get_output_thy thy of
+                      "" => thy
+                    | output_thy => if internal
+                                    then add_hol4_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+                                    else thy
+        val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+        val curmaps = HOL4ConstMaps.get thy
+        val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,NONE)) curmaps
+        val upd_thy = HOL4ConstMaps.put newmaps thy
     in
-	upd_thy
+        upd_thy
     end;
 
 fun add_hol4_const_renaming bthy bconst newname thy =
     let
-	val currens = HOL4Rename.get thy
-	val _ = message ("Adding renaming " ^ bthy ^ "." ^ bconst ^ " -> " ^ newname)
-	val newrens = StringPair.update_new ((bthy,bconst),newname) currens
-	val upd_thy = HOL4Rename.put newrens thy
+        val currens = HOL4Rename.get thy
+        val _ = message ("Adding renaming " ^ bthy ^ "." ^ bconst ^ " -> " ^ newname)
+        val newrens = StringPair.update_new ((bthy,bconst),newname) currens
+        val upd_thy = HOL4Rename.put newrens thy
     in
-	upd_thy
+        upd_thy
     end;
 
 fun get_hol4_const_renaming bthy bconst thy =
     let
-	val currens = HOL4Rename.get thy
+        val currens = HOL4Rename.get thy
     in
-	StringPair.lookup currens (bthy,bconst)
+        StringPair.lookup currens (bthy,bconst)
     end;
 
 fun get_hol4_const_mapping bthy bconst thy =
     let
-	val bconst = case get_hol4_const_renaming bthy bconst thy of
-			 SOME name => name
-		       | NONE => bconst
-	val maps = HOL4ConstMaps.get thy
+        val bconst = case get_hol4_const_renaming bthy bconst thy of
+                         SOME name => name
+                       | NONE => bconst
+        val maps = HOL4ConstMaps.get thy
     in
-	StringPair.lookup maps (bthy,bconst)
+        StringPair.lookup maps (bthy,bconst)
     end
 
 fun add_hol4_const_wt_mapping bthy bconst internal isaconst typ thy =
     let
-	val thy = case opt_get_output_thy thy of
-		      "" => thy
-		    | output_thy => if internal
-				    then add_hol4_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
-				    else thy
-	val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
-	val curmaps = HOL4ConstMaps.get thy
-	val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,SOME typ)) curmaps
-	val upd_thy = HOL4ConstMaps.put newmaps thy
+        val thy = case opt_get_output_thy thy of
+                      "" => thy
+                    | output_thy => if internal
+                                    then add_hol4_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+                                    else thy
+        val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+        val curmaps = HOL4ConstMaps.get thy
+        val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,SOME typ)) curmaps
+        val upd_thy = HOL4ConstMaps.put newmaps thy
     in
-	upd_thy
+        upd_thy
     end;
 
 fun add_hol4_type_mapping bthy bconst internal isaconst thy =
     let
-	val curmaps = HOL4TypeMaps.get thy
-	val _ = writeln ("Adding tmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
-	val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst)) curmaps
+        val curmaps = HOL4TypeMaps.get thy
+        val _ = writeln ("Adding tmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+        val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst)) curmaps
                handle x => let val (internal, isaconst') = the (StringPair.lookup curmaps (bthy, bconst)) in
                       warning ("couldn't map type "^bthy^"."^bconst^" to "^isaconst^": already mapped to "^isaconst'); raise x end
         val upd_thy = HOL4TypeMaps.put newmaps thy
     in
-	upd_thy
+        upd_thy
     end;
 
 fun add_hol4_pending bthy bthm hth thy =
     let
-	val thmname = Sign.full_bname thy bthm
-	val _ = message ("Add pending " ^ bthy ^ "." ^ bthm)
-	val curpend = HOL4Pending.get thy
-	val newpend = StringPair.update_new ((bthy,bthm),hth) curpend
-	val upd_thy = HOL4Pending.put newpend thy
-	val thy' = case opt_get_output_thy upd_thy of
-		       "" => add_hol4_mapping bthy bthm thmname upd_thy
-		     | output_thy =>
-		       let
-			   val new_thmname = output_thy ^ "." ^ bthy ^ "." ^ bthm
-		       in
-			   upd_thy |> add_hol4_move thmname new_thmname
-				   |> add_hol4_mapping bthy bthm new_thmname
-		       end
+        val thmname = Sign.full_bname thy bthm
+        val _ = message ("Add pending " ^ bthy ^ "." ^ bthm)
+        val curpend = HOL4Pending.get thy
+        val newpend = StringPair.update_new ((bthy,bthm),hth) curpend
+        val upd_thy = HOL4Pending.put newpend thy
+        val thy' = case opt_get_output_thy upd_thy of
+                       "" => add_hol4_mapping bthy bthm thmname upd_thy
+                     | output_thy =>
+                       let
+                           val new_thmname = output_thy ^ "." ^ bthy ^ "." ^ bthm
+                       in
+                           upd_thy |> add_hol4_move thmname new_thmname
+                                   |> add_hol4_mapping bthy bthm new_thmname
+                       end
     in
-	thy'
+        thy'
     end;
 
 fun get_hol4_theorem thyname thmname thy =
@@ -386,7 +386,7 @@
     fun process ((bthy,bthm), hth as (_,thm)) thy =
       let
         val thm1 = rewrite_rule (map (Thm.transfer thy) rews) (Thm.transfer thy thm);
-        val thm2 = standard thm1;
+        val thm2 = Drule.standard thm1;
       in
         thy
         |> PureThy.store_thm (Binding.name bthm, thm2)
@@ -404,79 +404,79 @@
 
 fun add_dump str thy =
     let
-	val (dir,thyname,curdump) = HOL4Dump.get thy
+        val (dir,thyname,curdump) = HOL4Dump.get thy
     in
-	HOL4Dump.put (dir,thyname,str::curdump) thy
+        HOL4Dump.put (dir,thyname,str::curdump) thy
     end
 
 fun flush_dump thy =
     let
-	val (dir,thyname,dumpdata) = HOL4Dump.get thy
-	val os = TextIO.openOut (OS.Path.joinDirFile {dir=dir,
-						      file=thyname ^ ".thy"})
-	val _  = app (fn s => TextIO.output(os,s ^ "\n\n")) (rev dumpdata)
-	val  _ = TextIO.closeOut os
+        val (dir,thyname,dumpdata) = HOL4Dump.get thy
+        val os = TextIO.openOut (OS.Path.joinDirFile {dir=dir,
+                                                      file=thyname ^ ".thy"})
+        val _  = app (fn s => TextIO.output(os,s ^ "\n\n")) (rev dumpdata)
+        val  _ = TextIO.closeOut os
     in
-	HOL4Dump.put ("","",[]) thy
+        HOL4Dump.put ("","",[]) thy
     end
 
 fun set_generating_thy thyname thy =
     case HOL4DefThy.get thy of
-	NoImport => HOL4DefThy.put (Generating thyname) thy
+        NoImport => HOL4DefThy.put (Generating thyname) thy
       | _ => error "Import already in progess"
 
 fun set_replaying_thy thyname thy =
     case HOL4DefThy.get thy of
-	NoImport => HOL4DefThy.put (Replaying thyname) thy
+        NoImport => HOL4DefThy.put (Replaying thyname) thy
       | _ => error "Import already in progess"
 
 fun clear_import_thy thy =
     case HOL4DefThy.get thy of
-	NoImport => error "No import in progress"
+        NoImport => error "No import in progress"
       | _ => HOL4DefThy.put NoImport thy
 
 fun get_generating_thy thy =
     case HOL4DefThy.get thy of
-	Generating thyname => thyname
+        Generating thyname => thyname
       | _ => error "No theory being generated"
 
 fun get_replaying_thy thy =
     case HOL4DefThy.get thy of
-	Replaying thyname => thyname
+        Replaying thyname => thyname
       | _ => error "No theory being replayed"
 
 fun get_import_thy thy =
     case HOL4DefThy.get thy of
-	Replaying thyname => thyname
+        Replaying thyname => thyname
       | Generating thyname => thyname
       | _ => error "No theory being imported"
 
 fun should_ignore thyname thy thmname =
     case get_hol4_mapping thyname thmname thy of
-	SOME NONE => true 
+        SOME NONE => true 
       | _ => false
 
 val trans_string =
     let
-	fun quote s = "\"" ^ s ^ "\""
-	fun F [] = []
-	  | F (#"\\" :: cs) = patch #"\\" cs
-	  | F (#"\"" :: cs) = patch #"\"" cs
-	  | F (c     :: cs) = c :: F cs
-	and patch c rest = #"\\" :: c :: F rest
+        fun quote s = "\"" ^ s ^ "\""
+        fun F [] = []
+          | F (#"\\" :: cs) = patch #"\\" cs
+          | F (#"\"" :: cs) = patch #"\"" cs
+          | F (c     :: cs) = c :: F cs
+        and patch c rest = #"\\" :: c :: F rest
     in
-	quote o String.implode o F o String.explode
+        quote o String.implode o F o String.explode
     end
 
 fun dump_import_thy thyname thy =
     let
-	val output_dir = get_output_dir thy
-	val output_thy = get_output_thy thy
-	val input_thy = Context.theory_name thy
-	val import_segment = get_import_segment thy
-	val os = TextIO.openOut (OS.Path.joinDirFile {dir=output_dir,
-						      file=thyname ^ ".imp"})
-	fun out s = TextIO.output(os,s)
+        val output_dir = get_output_dir thy
+        val output_thy = get_output_thy thy
+        val input_thy = Context.theory_name thy
+        val import_segment = get_import_segment thy
+        val os = TextIO.openOut (OS.Path.joinDirFile {dir=output_dir,
+                                                      file=thyname ^ ".imp"})
+        fun out s = TextIO.output(os,s)
     val (ignored, mapped) = StringPair.fold
       (fn ((bthy, bthm), v) => fn (ign, map) =>
         if bthy = thyname
@@ -490,141 +490,141 @@
     val constrenames = mk (HOL4Rename.get thy);
     val typemaps = mk (HOL4TypeMaps.get thy);
     val defmaps = mk (HOL4DefMaps.get thy);
-	fun new_name internal isa =
-	    if internal
-	    then
-		let
-		    val paths = Long_Name.explode isa
-		    val i = Library.drop(length paths - 2,paths)
-		in
-		    case i of
-			[seg,con] => output_thy ^ "." ^ seg ^ "." ^ con
-		      | _ => error "hol4rews.dump internal error"
-		end
-	    else
-		isa
+        fun new_name internal isa =
+            if internal
+            then
+                let
+                    val paths = Long_Name.explode isa
+                    val i = Library.drop(length paths - 2,paths)
+                in
+                    case i of
+                        [seg,con] => output_thy ^ "." ^ seg ^ "." ^ con
+                      | _ => error "hol4rews.dump internal error"
+                end
+            else
+                isa
 
-	val _ = out "import\n\n"
+        val _ = out "import\n\n"
 
-	val _ = out ("import_segment " ^ trans_string import_segment ^ "\n\n")
-	val _ = if null defmaps
-		then ()
-		else out "def_maps"
-	val _ = app (fn (hol,isa) =>
-			out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string isa))) defmaps
-	val _ = if null defmaps
-		then ()
-		else out "\n\n"
+        val _ = out ("import_segment " ^ trans_string import_segment ^ "\n\n")
+        val _ = if null defmaps
+                then ()
+                else out "def_maps"
+        val _ = app (fn (hol,isa) =>
+                        out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string isa))) defmaps
+        val _ = if null defmaps
+                then ()
+                else out "\n\n"
 
-	val _ = if null typemaps
-		then ()
-		else out "type_maps"
-	val _ = app (fn (hol,(internal,isa)) =>
-			out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (new_name internal isa)))) typemaps
-	val _ = if null typemaps
-		then ()
-		else out "\n\n"
+        val _ = if null typemaps
+                then ()
+                else out "type_maps"
+        val _ = app (fn (hol,(internal,isa)) =>
+                        out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (new_name internal isa)))) typemaps
+        val _ = if null typemaps
+                then ()
+                else out "\n\n"
 
-	val _ = if null constmaps
-		then ()
-		else out "const_maps"
-	val _ = app (fn (hol,(internal,isa,opt_ty)) =>
-			(out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (follow_cname isa thy)));
-			 case opt_ty of
-			     SOME ty => out (" :: \"" ^ Syntax.string_of_typ_global thy ty ^ "\"")
-			   | NONE => ())) constmaps
-	val _ = if null constmaps
-		then ()
-		else out "\n\n"
+        val _ = if null constmaps
+                then ()
+                else out "const_maps"
+        val _ = app (fn (hol,(internal,isa,opt_ty)) =>
+                        (out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (follow_cname isa thy)));
+                         case opt_ty of
+                             SOME ty => out (" :: \"" ^ Syntax.string_of_typ_global thy ty ^ "\"")
+                           | NONE => ())) constmaps
+        val _ = if null constmaps
+                then ()
+                else out "\n\n"
 
-	val _ = if null constrenames
-		then ()
-		else out "const_renames"
-	val _ = app (fn (old,new) =>
-			out ("\n  " ^ (trans_string old) ^ " > " ^ (trans_string new))) constrenames
-	val _ = if null constrenames
-		then ()
-		else out "\n\n"
+        val _ = if null constrenames
+                then ()
+                else out "const_renames"
+        val _ = app (fn (old,new) =>
+                        out ("\n  " ^ (trans_string old) ^ " > " ^ (trans_string new))) constrenames
+        val _ = if null constrenames
+                then ()
+                else out "\n\n"
 
-	fun gen2replay in_thy out_thy s = 
-	    let
-		val ss = Long_Name.explode s
-	    in
-		if (hd ss = in_thy) then 
-		    Long_Name.implode (out_thy::(tl ss))
-		else
-		    s
-	    end 
+        fun gen2replay in_thy out_thy s = 
+            let
+                val ss = Long_Name.explode s
+            in
+                if (hd ss = in_thy) then 
+                    Long_Name.implode (out_thy::(tl ss))
+                else
+                    s
+            end 
 
-	val _ = if null mapped
-		then ()
-		else out "thm_maps"
-	val _ = app (fn (hol,isa) => out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (gen2replay input_thy output_thy isa)))) mapped
-	val _ = if null mapped 
-		then ()
-		else out "\n\n"
+        val _ = if null mapped
+                then ()
+                else out "thm_maps"
+        val _ = app (fn (hol,isa) => out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (gen2replay input_thy output_thy isa)))) mapped
+        val _ = if null mapped 
+                then ()
+                else out "\n\n"
 
-	val _ = if null ignored
-		then ()
-		else out "ignore_thms"
-	val _ = app (fn ign => out ("\n  " ^ (trans_string ign))) ignored
-	val _ = if null ignored
-		then ()
-		else out "\n\n"
+        val _ = if null ignored
+                then ()
+                else out "ignore_thms"
+        val _ = app (fn ign => out ("\n  " ^ (trans_string ign))) ignored
+        val _ = if null ignored
+                then ()
+                else out "\n\n"
 
-	val _ = out "end\n"
-	val _ = TextIO.closeOut os
+        val _ = out "end\n"
+        val _ = TextIO.closeOut os
     in
-	thy
+        thy
     end
 
 fun set_used_names names thy =
     let
-	val unames = HOL4UNames.get thy
+        val unames = HOL4UNames.get thy
     in
-	case unames of
-	    [] => HOL4UNames.put names thy
-	  | _ => error "hol4rews.set_used_names called on initialized data!"
+        case unames of
+            [] => HOL4UNames.put names thy
+          | _ => error "hol4rews.set_used_names called on initialized data!"
     end
 
 val clear_used_names = HOL4UNames.put [];
 
 fun get_defmap thyname const thy =
     let
-	val maps = HOL4DefMaps.get thy
+        val maps = HOL4DefMaps.get thy
     in
-	StringPair.lookup maps (thyname,const)
+        StringPair.lookup maps (thyname,const)
     end
 
 fun add_defmap thyname const defname thy =
     let
-	val _ = message ("Adding defmap " ^ thyname ^ "." ^ const ^ " --> " ^ defname)
-	val maps = HOL4DefMaps.get thy
-	val maps' = StringPair.update_new ((thyname,const),defname) maps
-	val thy' = HOL4DefMaps.put maps' thy
+        val _ = message ("Adding defmap " ^ thyname ^ "." ^ const ^ " --> " ^ defname)
+        val maps = HOL4DefMaps.get thy
+        val maps' = StringPair.update_new ((thyname,const),defname) maps
+        val thy' = HOL4DefMaps.put maps' thy
     in
-	thy'
+        thy'
     end
 
 fun get_defname thyname name thy =
     let
-	val maps = HOL4DefMaps.get thy
-	fun F dname = (dname,add_defmap thyname name dname thy)
+        val maps = HOL4DefMaps.get thy
+        fun F dname = (dname,add_defmap thyname name dname thy)
     in
-	case StringPair.lookup maps (thyname,name) of
-	    SOME thmname => (thmname,thy)
-	  | NONE =>
-	    let
-		val used = HOL4UNames.get thy
-		val defname = Thm.def_name name
-		val pdefname = name ^ "_primdef"
-	    in
-		if not (defname mem used)
-		then F defname                 (* name_def *)
-		else if not (pdefname mem used)
-		then F pdefname                (* name_primdef *)
-		else F (Name.variant used pdefname) (* last resort *)
-	    end
+        case StringPair.lookup maps (thyname,name) of
+            SOME thmname => (thmname,thy)
+          | NONE =>
+            let
+                val used = HOL4UNames.get thy
+                val defname = Thm.def_name name
+                val pdefname = name ^ "_primdef"
+            in
+                if not (defname mem used)
+                then F defname                 (* name_def *)
+                else if not (pdefname mem used)
+                then F pdefname                (* name_primdef *)
+                else F (Name.variant used pdefname) (* last resort *)
+            end
     end
 
 local
@@ -639,12 +639,12 @@
 
 local
     fun initial_maps thy =
-	thy |> add_hol4_type_mapping "min" "bool" false "bool"
-	    |> add_hol4_type_mapping "min" "fun" false "fun"
-	    |> add_hol4_type_mapping "min" "ind" false "Nat.ind"
-	    |> add_hol4_const_mapping "min" "=" false "op ="
-	    |> add_hol4_const_mapping "min" "==>" false "op -->"
-	    |> add_hol4_const_mapping "min" "@" false "Hilbert_Choice.Eps"
+        thy |> add_hol4_type_mapping "min" "bool" false "bool"
+            |> add_hol4_type_mapping "min" "fun" false "fun"
+            |> add_hol4_type_mapping "min" "ind" false "Nat.ind"
+            |> add_hol4_const_mapping "min" "=" false "op ="
+            |> add_hol4_const_mapping "min" "==>" false "op -->"
+            |> add_hol4_const_mapping "min" "@" false "Hilbert_Choice.Eps"
 in
 val hol4_setup =
   initial_maps #>
--- a/src/HOL/Import/import_syntax.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/import_syntax.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Import/import_syntax.ML
-    ID:         $Id$
     Author:     Sebastian Skalberg (TU Muenchen)
 *)
 
@@ -43,10 +42,10 @@
 
 
 val import_theory = P.name >> (fn thyname =>
-			       fn thy =>
-				  thy |> set_generating_thy thyname
-				      |> Sign.add_path thyname
-				      |> add_dump (";setup_theory " ^ thyname))
+                               fn thy =>
+                                  thy |> set_generating_thy thyname
+                                      |> Sign.add_path thyname
+                                      |> add_dump (";setup_theory " ^ thyname))
 
 fun do_skip_import_dir s = (ImportRecorder.set_skip_import_dir (SOME s); I)
 val skip_import_dir : command = P.string >> do_skip_import_dir
@@ -57,167 +56,167 @@
 
 fun end_import toks =
     Scan.succeed
-	(fn thy =>
-	    let
-		val thyname = get_generating_thy thy
-		val segname = get_import_segment thy
-		val (int_thms,facts) = Replay.setup_int_thms thyname thy
-		val thmnames = List.filter (not o should_ignore thyname thy) facts
-		fun replay thy = 
-		    let
-			val _ = ImportRecorder.load_history thyname
-			val thy = Replay.import_thms thyname int_thms thmnames thy
-			    handle x => (ImportRecorder.save_history thyname; raise x)
-			val _ = ImportRecorder.save_history thyname
-			val _ = ImportRecorder.clear_history ()
-		    in
-			thy
-		    end					
-	    in
-		thy |> clear_import_thy
-		    |> set_segment thyname segname
-		    |> set_used_names facts
-		    |> replay 
-		    |> clear_used_names
-		    |> export_hol4_pending
-		    |> Sign.parent_path
-		    |> dump_import_thy thyname
-		    |> add_dump ";end_setup"
-	    end) toks
+        (fn thy =>
+            let
+                val thyname = get_generating_thy thy
+                val segname = get_import_segment thy
+                val (int_thms,facts) = Replay.setup_int_thms thyname thy
+                val thmnames = List.filter (not o should_ignore thyname thy) facts
+                fun replay thy = 
+                    let
+                        val _ = ImportRecorder.load_history thyname
+                        val thy = Replay.import_thms thyname int_thms thmnames thy
+                            handle x => (ImportRecorder.save_history thyname; raise x)
+                        val _ = ImportRecorder.save_history thyname
+                        val _ = ImportRecorder.clear_history ()
+                    in
+                        thy
+                    end                                 
+            in
+                thy |> clear_import_thy
+                    |> set_segment thyname segname
+                    |> set_used_names facts
+                    |> replay 
+                    |> clear_used_names
+                    |> export_hol4_pending
+                    |> Sign.parent_path
+                    |> dump_import_thy thyname
+                    |> add_dump ";end_setup"
+            end) toks
 
 val ignore_thms = Scan.repeat1 P.name
-			       >> (fn ignored =>
-				   fn thy =>
-				      let
-					  val thyname = get_import_thy thy
-				      in
-					  Library.foldl (fn (thy,thmname) => ignore_hol4 thyname thmname thy) (thy,ignored)
-				      end)
+                               >> (fn ignored =>
+                                   fn thy =>
+                                      let
+                                          val thyname = get_import_thy thy
+                                      in
+                                          Library.foldl (fn (thy,thmname) => ignore_hol4 thyname thmname thy) (thy,ignored)
+                                      end)
 
 val thm_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname)
-			    >> (fn thmmaps =>
-				fn thy =>
-				   let
-				       val thyname = get_import_thy thy
-				   in
-				       Library.foldl (fn (thy,(hol,isa)) => add_hol4_mapping thyname hol isa thy) (thy,thmmaps)
-				   end)
+                            >> (fn thmmaps =>
+                                fn thy =>
+                                   let
+                                       val thyname = get_import_thy thy
+                                   in
+                                       Library.foldl (fn (thy,(hol,isa)) => add_hol4_mapping thyname hol isa thy) (thy,thmmaps)
+                                   end)
 
 val type_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname)
-			     >> (fn typmaps =>
-				 fn thy =>
-				    let
-					val thyname = get_import_thy thy
-				    in
-					Library.foldl (fn (thy,(hol,isa)) => add_hol4_type_mapping thyname hol false isa thy) (thy,typmaps)
-				    end)
+                             >> (fn typmaps =>
+                                 fn thy =>
+                                    let
+                                        val thyname = get_import_thy thy
+                                    in
+                                        Library.foldl (fn (thy,(hol,isa)) => add_hol4_type_mapping thyname hol false isa thy) (thy,typmaps)
+                                    end)
 
 val def_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname)
-			    >> (fn defmaps =>
-				fn thy =>
-				   let
-				       val thyname = get_import_thy thy
-				   in
-				       Library.foldl (fn (thy,(hol,isa)) => add_defmap thyname hol isa thy) (thy,defmaps)
-				   end)
+                            >> (fn defmaps =>
+                                fn thy =>
+                                   let
+                                       val thyname = get_import_thy thy
+                                   in
+                                       Library.foldl (fn (thy,(hol,isa)) => add_defmap thyname hol isa thy) (thy,defmaps)
+                                   end)
 
 val const_renames = Scan.repeat1 (P.name --| P.$$$ ">" -- P.name)
-				 >> (fn renames =>
-				     fn thy =>
-					let
-					    val thyname = get_import_thy thy
-					in
-					    Library.foldl (fn (thy,(hol,isa)) => add_hol4_const_renaming thyname hol isa thy) (thy,renames)
-					end)
-																      
+                                 >> (fn renames =>
+                                     fn thy =>
+                                        let
+                                            val thyname = get_import_thy thy
+                                        in
+                                            Library.foldl (fn (thy,(hol,isa)) => add_hol4_const_renaming thyname hol isa thy) (thy,renames)
+                                        end)
+                                                                                                                                      
 val const_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
-			      >> (fn constmaps =>
-				  fn thy =>
-				     let
-					 val thyname = get_import_thy thy
-				     in
-					 Library.foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol false isa thy
-						 | (thy,((hol,isa),SOME ty)) => add_hol4_const_wt_mapping thyname hol false isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
-				     end)
+                              >> (fn constmaps =>
+                                  fn thy =>
+                                     let
+                                         val thyname = get_import_thy thy
+                                     in
+                                         Library.foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol false isa thy
+                                                 | (thy,((hol,isa),SOME ty)) => add_hol4_const_wt_mapping thyname hol false isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
+                                     end)
 
 val const_moves = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
-			       >> (fn constmaps =>
-				   fn thy =>
-				      let
-					  val thyname = get_import_thy thy
-				      in
-					  Library.foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol true isa thy
-						  | (thy,((hol,isa),SOME ty)) => add_hol4_const_wt_mapping thyname hol true isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
-				      end)
+                               >> (fn constmaps =>
+                                   fn thy =>
+                                      let
+                                          val thyname = get_import_thy thy
+                                      in
+                                          Library.foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol true isa thy
+                                                  | (thy,((hol,isa),SOME ty)) => add_hol4_const_wt_mapping thyname hol true isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
+                                      end)
 
 fun import_thy s =
     let
-	val is = TextIO.openIn(s ^ ".imp")
-	val inp = TextIO.inputAll is
-	val _ = TextIO.closeIn is
-	val orig_source = Source.of_string inp
-	val symb_source = Symbol.source {do_recover = false} orig_source
-	val lexes = Unsynchronized.ref
-	  (Scan.make_lexicon (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
-		  Scan.empty_lexicon)
-	val get_lexes = fn () => !lexes
-	val token_source = OuterLex.source {do_recover = NONE} get_lexes Position.start symb_source
-	val token_list = filter_out (OuterLex.is_kind OuterLex.Space) (Source.exhaust token_source)
-	val import_segmentP = OuterParse.$$$ "import_segment" |-- import_segment
-	val type_mapsP = OuterParse.$$$ "type_maps" |-- type_maps
-	val def_mapsP = OuterParse.$$$ "def_maps" |-- def_maps
-	val const_mapsP = OuterParse.$$$ "const_maps" |-- const_maps
-	val const_movesP = OuterParse.$$$ "const_moves" |-- const_moves
-	val const_renamesP = OuterParse.$$$ "const_renames" |-- const_renames
-	val ignore_thmsP = OuterParse.$$$ "ignore_thms" |-- ignore_thms
-	val thm_mapsP = OuterParse.$$$ "thm_maps" |-- thm_maps
-	val parser = type_mapsP || def_mapsP || const_mapsP || const_movesP || const_renamesP || thm_mapsP || ignore_thmsP || import_segmentP
-	val importP = OuterParse.$$$ "import" |-- Scan.repeat parser --| OuterParse.$$$ "end"
-	fun apply [] thy = thy
-	  | apply (f::fs) thy = apply fs (f thy)
+        val is = TextIO.openIn(s ^ ".imp")
+        val inp = TextIO.inputAll is
+        val _ = TextIO.closeIn is
+        val orig_source = Source.of_string inp
+        val symb_source = Symbol.source {do_recover = false} orig_source
+        val lexes = Unsynchronized.ref
+          (Scan.make_lexicon (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
+                  Scan.empty_lexicon)
+        val get_lexes = fn () => !lexes
+        val token_source = OuterLex.source {do_recover = NONE} get_lexes Position.start symb_source
+        val token_list = filter_out (OuterLex.is_kind OuterLex.Space) (Source.exhaust token_source)
+        val import_segmentP = OuterParse.$$$ "import_segment" |-- import_segment
+        val type_mapsP = OuterParse.$$$ "type_maps" |-- type_maps
+        val def_mapsP = OuterParse.$$$ "def_maps" |-- def_maps
+        val const_mapsP = OuterParse.$$$ "const_maps" |-- const_maps
+        val const_movesP = OuterParse.$$$ "const_moves" |-- const_moves
+        val const_renamesP = OuterParse.$$$ "const_renames" |-- const_renames
+        val ignore_thmsP = OuterParse.$$$ "ignore_thms" |-- ignore_thms
+        val thm_mapsP = OuterParse.$$$ "thm_maps" |-- thm_maps
+        val parser = type_mapsP || def_mapsP || const_mapsP || const_movesP || const_renamesP || thm_mapsP || ignore_thmsP || import_segmentP
+        val importP = OuterParse.$$$ "import" |-- Scan.repeat parser --| OuterParse.$$$ "end"
+        fun apply [] thy = thy
+          | apply (f::fs) thy = apply fs (f thy)
     in
-	apply (set_replaying_thy s::Sign.add_path s::(fst (importP token_list)))
+        apply (set_replaying_thy s::Sign.add_path s::(fst (importP token_list)))
     end
 
 fun create_int_thms thyname thy =
     if ! quick_and_dirty
     then thy
     else
-	case ImportData.get thy of
-	    NONE => ImportData.put (SOME (fst (Replay.setup_int_thms thyname thy))) thy
-	  | SOME _ => error "Import data not closed - forgotten an end_setup, mayhap?"
+        case ImportData.get thy of
+            NONE => ImportData.put (SOME (fst (Replay.setup_int_thms thyname thy))) thy
+          | SOME _ => error "Import data not closed - forgotten an end_setup, mayhap?"
 
 fun clear_int_thms thy =
     if ! quick_and_dirty
     then thy
     else
-	case ImportData.get thy of
-	    NONE => error "Import data already cleared - forgotten a setup_theory?"
-	  | SOME _ => ImportData.put NONE thy
+        case ImportData.get thy of
+            NONE => error "Import data already cleared - forgotten a setup_theory?"
+          | SOME _ => ImportData.put NONE thy
 
 val setup_theory = P.name
-		       >>
-		       (fn thyname =>
-			fn thy =>
-			   (case HOL4DefThy.get thy of
-				NoImport => thy
-			      | Generating _ => error "Currently generating a theory!"
-			      | Replaying _ => thy |> clear_import_thy)
-			       |> import_thy thyname
-			       |> create_int_thms thyname)
+                       >>
+                       (fn thyname =>
+                        fn thy =>
+                           (case HOL4DefThy.get thy of
+                                NoImport => thy
+                              | Generating _ => error "Currently generating a theory!"
+                              | Replaying _ => thy |> clear_import_thy)
+                               |> import_thy thyname
+                               |> create_int_thms thyname)
 
 fun end_setup toks =
     Scan.succeed
-	(fn thy =>
-	    let
-		val thyname = get_import_thy thy
-		val segname = get_import_segment thy
-	    in
-		thy |> set_segment thyname segname
-		    |> clear_import_thy
-		    |> clear_int_thms
-		    |> Sign.parent_path
-	    end) toks
+        (fn thy =>
+            let
+                val thyname = get_import_thy thy
+                val segname = get_import_segment thy
+            in
+                thy |> set_segment thyname segname
+                    |> clear_import_thy
+                    |> clear_int_thms
+                    |> Sign.parent_path
+            end) toks
 
 val set_dump  = P.string -- P.string   >> setup_dump
 fun fl_dump toks  = Scan.succeed flush_dump toks
@@ -226,56 +225,56 @@
 fun setup () =
   (OuterKeyword.keyword ">";
    OuterSyntax.command "import_segment"
-		       "Set import segment name"
-		       K.thy_decl (import_segment >> Toplevel.theory);
+                       "Set import segment name"
+                       K.thy_decl (import_segment >> Toplevel.theory);
    OuterSyntax.command "import_theory"
-		       "Set default HOL4 theory name"
-		       K.thy_decl (import_theory >> Toplevel.theory);
+                       "Set default HOL4 theory name"
+                       K.thy_decl (import_theory >> Toplevel.theory);
    OuterSyntax.command "end_import"
-		       "End HOL4 import"
-		       K.thy_decl (end_import >> Toplevel.theory);
+                       "End HOL4 import"
+                       K.thy_decl (end_import >> Toplevel.theory);
    OuterSyntax.command "skip_import_dir" 
                        "Sets caching directory for skipping importing"
                        K.thy_decl (skip_import_dir >> Toplevel.theory);
    OuterSyntax.command "skip_import" 
                        "Switches skipping importing on or off"
-                       K.thy_decl (skip_import >> Toplevel.theory);		      
+                       K.thy_decl (skip_import >> Toplevel.theory);                   
    OuterSyntax.command "setup_theory"
-		       "Setup HOL4 theory replaying"
-		       K.thy_decl (setup_theory >> Toplevel.theory);
+                       "Setup HOL4 theory replaying"
+                       K.thy_decl (setup_theory >> Toplevel.theory);
    OuterSyntax.command "end_setup"
-		       "End HOL4 setup"
-		       K.thy_decl (end_setup >> Toplevel.theory);
+                       "End HOL4 setup"
+                       K.thy_decl (end_setup >> Toplevel.theory);
    OuterSyntax.command "setup_dump"
-		       "Setup the dump file name"
-		       K.thy_decl (set_dump >> Toplevel.theory);
+                       "Setup the dump file name"
+                       K.thy_decl (set_dump >> Toplevel.theory);
    OuterSyntax.command "append_dump"
-		       "Add text to dump file"
-		       K.thy_decl (append_dump >> Toplevel.theory);
+                       "Add text to dump file"
+                       K.thy_decl (append_dump >> Toplevel.theory);
    OuterSyntax.command "flush_dump"
-		       "Write the dump to file"
-		       K.thy_decl (fl_dump >> Toplevel.theory);
+                       "Write the dump to file"
+                       K.thy_decl (fl_dump >> Toplevel.theory);
    OuterSyntax.command "ignore_thms"
-		       "Theorems to ignore in next HOL4 theory import"
-		       K.thy_decl (ignore_thms >> Toplevel.theory);
+                       "Theorems to ignore in next HOL4 theory import"
+                       K.thy_decl (ignore_thms >> Toplevel.theory);
    OuterSyntax.command "type_maps"
-		       "Map HOL4 type names to existing Isabelle/HOL type names"
-		       K.thy_decl (type_maps >> Toplevel.theory);
+                       "Map HOL4 type names to existing Isabelle/HOL type names"
+                       K.thy_decl (type_maps >> Toplevel.theory);
    OuterSyntax.command "def_maps"
-		       "Map HOL4 constant names to their primitive definitions"
-		       K.thy_decl (def_maps >> Toplevel.theory);
+                       "Map HOL4 constant names to their primitive definitions"
+                       K.thy_decl (def_maps >> Toplevel.theory);
    OuterSyntax.command "thm_maps"
-		       "Map HOL4 theorem names to existing Isabelle/HOL theorem names"
-		       K.thy_decl (thm_maps >> Toplevel.theory);
+                       "Map HOL4 theorem names to existing Isabelle/HOL theorem names"
+                       K.thy_decl (thm_maps >> Toplevel.theory);
    OuterSyntax.command "const_renames"
-		       "Rename HOL4 const names"
-		       K.thy_decl (const_renames >> Toplevel.theory);
+                       "Rename HOL4 const names"
+                       K.thy_decl (const_renames >> Toplevel.theory);
    OuterSyntax.command "const_moves"
-		       "Rename HOL4 const names to other HOL4 constants"
-		       K.thy_decl (const_moves >> Toplevel.theory);
+                       "Rename HOL4 const names to other HOL4 constants"
+                       K.thy_decl (const_moves >> Toplevel.theory);
    OuterSyntax.command "const_maps"
-		       "Map HOL4 const names to existing Isabelle/HOL const names"
-		       K.thy_decl (const_maps >> Toplevel.theory));
+                       "Map HOL4 const names to existing Isabelle/HOL const names"
+                       K.thy_decl (const_maps >> Toplevel.theory));
 
 end
 
--- a/src/HOL/Import/importrecorder.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/importrecorder.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -2,26 +2,26 @@
 sig
 
     datatype deltastate = Consts of (string * typ * mixfix) list
-			| Specification of (string * string * bool) list * term
-			| Hol_mapping of string * string * string
-			| Hol_pending of string * string * term
-			| Hol_const_mapping of string * string * string
-			| Hol_move of string * string
-			| Defs of string * term
-			| Hol_theorem of string * string * term
-			| Typedef of string option * (string * string list * mixfix) * term * (string * string) option * term 
-			| Hol_type_mapping of string * string * string
-			| Indexed_theorem of int * term
-	                | Protect_varname of string * string
-			| Dump of string
+                        | Specification of (string * string * bool) list * term
+                        | Hol_mapping of string * string * string
+                        | Hol_pending of string * string * term
+                        | Hol_const_mapping of string * string * string
+                        | Hol_move of string * string
+                        | Defs of string * term
+                        | Hol_theorem of string * string * term
+                        | Typedef of string option * (string * string list * mixfix) * term * (string * string) option * term 
+                        | Hol_type_mapping of string * string * string
+                        | Indexed_theorem of int * term
+                        | Protect_varname of string * string
+                        | Dump of string
 
     datatype history = History of history_entry list
-	 and history_entry = ThmEntry of string*string*bool*history | DeltaEntry of deltastate list
+         and history_entry = ThmEntry of string*string*bool*history | DeltaEntry of deltastate list
 
     val get_history : unit -> history
     val set_history : history -> unit
     val clear_history : unit -> unit
-				      
+                                      
     val start_replay_proof : string -> string -> unit
     val stop_replay_proof : string -> string -> unit
     val abort_replay_proof : string -> string -> unit
@@ -54,23 +54,23 @@
 struct
 
 datatype deltastate = Consts of (string * typ * mixfix) list
-		    | Specification of (string * string * bool) list * term
-		    | Hol_mapping of string * string * string
-		    | Hol_pending of string * string * term
-		    | Hol_const_mapping of string * string * string
-		    | Hol_move of string * string
-		    | Defs of string * term
-		    | Hol_theorem of string * string * term
-		    | Typedef of string option * (string * string list * mixfix) * term * (string * string) option * term 
-		    | Hol_type_mapping of string * string * string
-		    | Indexed_theorem of int * term
-	            | Protect_varname of string * string
-		    | Dump of string
+                    | Specification of (string * string * bool) list * term
+                    | Hol_mapping of string * string * string
+                    | Hol_pending of string * string * term
+                    | Hol_const_mapping of string * string * string
+                    | Hol_move of string * string
+                    | Defs of string * term
+                    | Hol_theorem of string * string * term
+                    | Typedef of string option * (string * string list * mixfix) * term * (string * string) option * term 
+                    | Hol_type_mapping of string * string * string
+                    | Indexed_theorem of int * term
+                    | Protect_varname of string * string
+                    | Dump of string
 
 datatype history_entry = StartReplay of string*string  
-		       | StopReplay of string*string
-		       | AbortReplay of string*string
-		       | Delta of deltastate list
+                       | StopReplay of string*string
+                       | AbortReplay of string*string
+                       | Delta of deltastate list
 
 val history = Unsynchronized.ref ([]:history_entry list)
 val history_dir = Unsynchronized.ref (SOME "")
@@ -184,7 +184,7 @@
 fun add_indexed_theorem i th = add_delta (Indexed_theorem (i, prop_of th))
 fun protect_varname s t = add_delta (Protect_varname (s,t))
 fun add_dump s = add_delta (Dump s)
-			    
+                            
 fun set_skip_import_dir dir = (history_dir := dir)
 fun get_skip_import_dir () = !history_dir
 
@@ -192,26 +192,26 @@
 
 fun save_history thyname = 
     if get_skip_import () then
-	XMLConv.write_to_file xml_of_history (get_filename thyname) (!history)
+        XMLConv.write_to_file xml_of_history (get_filename thyname) (!history)
     else 
-	()
+        ()
 
 fun load_history thyname = 
     if get_skip_import () then
-	let 
-	    val filename = get_filename thyname
-	    val _ = writeln "load_history / before"
-	    val _ = 
-		if File.exists (Path.explode filename) then					
-		    (history := XMLConv.read_from_file history_of_xml (get_filename thyname)) 
-		else
-		    clear_history ()
-	    val _ = writeln "load_history / after"
-	in
-	    ()
-	end
+        let 
+            val filename = get_filename thyname
+            val _ = writeln "load_history / before"
+            val _ = 
+                if File.exists (Path.explode filename) then                                     
+                    (history := XMLConv.read_from_file history_of_xml (get_filename thyname)) 
+                else
+                    clear_history ()
+            val _ = writeln "load_history / after"
+        in
+            ()
+        end
     else
-	()
+        ()
     
  
 datatype history = History of history_entry list
@@ -221,35 +221,35 @@
 
 fun conv_histories ((StartReplay (thyname, thmname))::rest) = 
     let
-	val (hs, rest) = conv_histories rest
-	fun continue thyname' thmname' aborted rest = 
-	    if thyname = thyname' andalso thmname = thmname' then
-		let
-		    val (hs', rest) = conv_histories rest
-		in
-		    ((ThmEntry (thyname, thmname, aborted, History hs))::hs', rest)
-		end
-	    else
-		raise CONV 
+        val (hs, rest) = conv_histories rest
+        fun continue thyname' thmname' aborted rest = 
+            if thyname = thyname' andalso thmname = thmname' then
+                let
+                    val (hs', rest) = conv_histories rest
+                in
+                    ((ThmEntry (thyname, thmname, aborted, History hs))::hs', rest)
+                end
+            else
+                raise CONV 
     in
-	case rest of 
-	    (StopReplay (thyname',thmname'))::rest =>
-	    continue thyname' thmname' false rest
-	  | (AbortReplay (thyname',thmname'))::rest =>
-	    continue thyname' thmname' true rest
-	  | [] => (hs, [])
-	  | _ => raise CONV
+        case rest of 
+            (StopReplay (thyname',thmname'))::rest =>
+            continue thyname' thmname' false rest
+          | (AbortReplay (thyname',thmname'))::rest =>
+            continue thyname' thmname' true rest
+          | [] => (hs, [])
+          | _ => raise CONV
     end
   | conv_histories ((Delta ds)::rest) = (conv_histories rest) |>> (fn hs => (DeltaEntry (rev ds))::hs)  
   | conv_histories rest = ([], rest)
 
 fun conv es = 
     let
-	val (h, rest) = conv_histories (rev es)
+        val (h, rest) = conv_histories (rev es)
     in
-	case rest of
-	    [] => h
-	  | _ => raise CONV
+        case rest of
+            [] => h
+          | _ => raise CONV
     end 
 
 fun get_history () = History (conv (!history))
--- a/src/HOL/Import/lazy_seq.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/lazy_seq.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -98,11 +98,11 @@
 
 fun s1 @ s2 =
     let
-	fun F NONE = getItem s2
-	  | F (SOME(x,s1')) = SOME(x,F' s1')
-	and F' s = make (fn () => F (getItem s))
+        fun F NONE = getItem s2
+          | F (SOME(x,s1')) = SOME(x,F' s1')
+        and F' s = make (fn () => F (getItem s))
     in
-	F' s1
+        F' s1
     end
 
 local
@@ -148,7 +148,7 @@
 fun take_at_most s n = 
     if n <= 0 then [] else
     case getItem s of 
-	NONE => []
+        NONE => []
       | SOME (x,s') => x::(take_at_most s' (n-1))
 
 local
@@ -180,117 +180,117 @@
 local
     fun F NONE = NONE
       | F (SOME(s1,s2)) =
-	let
-	    fun G NONE = F (getItem s2)
-	      | G (SOME(x,s)) = SOME(x,make (fn () => G (getItem s)))
-	in
-	    G (getItem s1)
-	end
+        let
+            fun G NONE = F (getItem s2)
+              | G (SOME(x,s)) = SOME(x,make (fn () => G (getItem s)))
+        in
+            G (getItem s1)
+        end
 in
 fun concat s = make (fn () => F (getItem s))
 end
 
 fun app f =
     let
-	fun F NONE = ()
-	  | F (SOME(x,s)) =
-	    (f x;
-	     F (getItem s))
+        fun F NONE = ()
+          | F (SOME(x,s)) =
+            (f x;
+             F (getItem s))
     in
-	F o getItem
+        F o getItem
     end
 
 fun map f =
     let
-	fun F NONE = NONE
-	  | F (SOME(x,s)) = SOME(f x,F' s)
-	and F' s = make (fn() => F (getItem s))
+        fun F NONE = NONE
+          | F (SOME(x,s)) = SOME(f x,F' s)
+        and F' s = make (fn() => F (getItem s))
     in
-	F'
+        F'
     end
 
 fun mapPartial f =
     let
-	fun F NONE = NONE
-	  | F (SOME(x,s)) =
-	    (case f x of
-		 SOME y => SOME(y,F' s)
-	       | NONE => F (getItem s))
-	and F' s = make (fn()=> F (getItem s))
+        fun F NONE = NONE
+          | F (SOME(x,s)) =
+            (case f x of
+                 SOME y => SOME(y,F' s)
+               | NONE => F (getItem s))
+        and F' s = make (fn()=> F (getItem s))
     in
-	F'
+        F'
     end
 
 fun find P =
     let
-	fun F NONE = NONE
-	  | F (SOME(x,s)) =
-	    if P x
-	    then SOME x
-	    else F (getItem s)
+        fun F NONE = NONE
+          | F (SOME(x,s)) =
+            if P x
+            then SOME x
+            else F (getItem s)
     in
-	F o getItem
+        F o getItem
     end
 
 (*fun filter p = mapPartial (fn x => if p x then SOME x else NONE)*)
 
 fun filter P =
     let
-	fun F NONE = NONE
-	  | F (SOME(x,s)) =
-	    if P x
-	    then SOME(x,F' s)
-	    else F (getItem s)
-	and F' s = make (fn () => F (getItem s))
+        fun F NONE = NONE
+          | F (SOME(x,s)) =
+            if P x
+            then SOME(x,F' s)
+            else F (getItem s)
+        and F' s = make (fn () => F (getItem s))
     in
-	F'
+        F'
     end
 
 fun partition f s =
     let
-	val s' = map (fn x => (x,f x)) s
+        val s' = map (fn x => (x,f x)) s
     in
-	(mapPartial (fn (x,true) => SOME x | _ => NONE) s',
-	 mapPartial (fn (x,false) => SOME x | _ => NONE) s')
+        (mapPartial (fn (x,true) => SOME x | _ => NONE) s',
+         mapPartial (fn (x,false) => SOME x | _ => NONE) s')
     end
 
 fun exists P =
     let
-	fun F NONE = false
-	  | F (SOME(x,s)) = P x orelse F (getItem s)
+        fun F NONE = false
+          | F (SOME(x,s)) = P x orelse F (getItem s)
     in
-	F o getItem
+        F o getItem
     end
 
 fun all P =
     let
-	fun F NONE = true
-	  | F (SOME(x,s)) = P x andalso F (getItem s)
+        fun F NONE = true
+          | F (SOME(x,s)) = P x andalso F (getItem s)
     in
-	F o getItem
+        F o getItem
     end
 
 (*fun tabulate f = map f (iterates (fn x => x + 1) 0)*)
 
 fun tabulate (n,f) =
     let
-	fun F n = make (fn () => SOME(f n,F (n+1)))
+        fun F n = make (fn () => SOME(f n,F (n+1)))
     in
-	F n
+        F n
     end
 
 fun collate c (s1,s2) =
     let
-	fun F NONE _ = LESS
-	  | F _ NONE = GREATER
-	  | F (SOME(x,s1)) (SOME(y,s2)) =
-	    (case c (x,y) of
-		 LESS => LESS
-	       | GREATER => GREATER
-	       | EQUAL => F' s1 s2)
-	and F' s1 s2 = F (getItem s1) (getItem s2)
+        fun F NONE _ = LESS
+          | F _ NONE = GREATER
+          | F (SOME(x,s1)) (SOME(y,s2)) =
+            (case c (x,y) of
+                 LESS => LESS
+               | GREATER => GREATER
+               | EQUAL => F' s1 s2)
+        and F' s1 s2 = F (getItem s1) (getItem s2)
     in
-	F' s1 s2
+        F' s1 s2
     end
 
 fun empty  _ = Seq (Lazy.value NONE)
@@ -299,25 +299,25 @@
 
 fun cycle seqfn =
     let
-	val knot = Unsynchronized.ref (Seq (Lazy.value NONE))
+        val knot = Unsynchronized.ref (Seq (Lazy.value NONE))
     in
-	knot := seqfn (fn () => !knot);
-	!knot
+        knot := seqfn (fn () => !knot);
+        !knot
     end
 
 fun iterates f =
     let
-	fun F x = make (fn() => SOME(x,F (f x)))
+        fun F x = make (fn() => SOME(x,F (f x)))
     in
-	F
+        F
     end
 
 fun interleave (s1,s2) =
     let
-	fun F NONE = getItem s2
-	  | F (SOME(x,s1')) = SOME(x,interleave(s2,s1'))
+        fun F NONE = getItem s2
+          | F (SOME(x,s1')) = SOME(x,interleave(s2,s1'))
     in
-	make (fn()=> F (getItem s1))
+        make (fn()=> F (getItem s1))
     end
 
 (* val force_all = app ignore *)
@@ -331,11 +331,11 @@
 
 fun of_function f =
     let
-	fun F () = case f () of
-		       SOME x => SOME(x,make F)
-		     | NONE => NONE
+        fun F () = case f () of
+                       SOME x => SOME(x,make F)
+                     | NONE => NONE
     in
-	make F
+        make F
     end
 
 local
@@ -350,17 +350,17 @@
 
 fun of_instream is =
     let
-	val buffer : char list Unsynchronized.ref = Unsynchronized.ref []
-	fun get_input () =
-	    case !buffer of
-		(c::cs) => (buffer:=cs;
-			    SOME c)
-	      |	[] => (case String.explode (TextIO.input is) of
-			   [] => NONE
-			 | (c::cs) => (buffer := cs;
-				       SOME c))
+        val buffer : char list Unsynchronized.ref = Unsynchronized.ref []
+        fun get_input () =
+            case !buffer of
+                (c::cs) => (buffer:=cs;
+                            SOME c)
+              | [] => (case String.explode (TextIO.input is) of
+                           [] => NONE
+                         | (c::cs) => (buffer := cs;
+                                       SOME c))
     in
-	of_function get_input
+        of_function get_input
     end
 
 local
@@ -379,32 +379,32 @@
 
 fun op THEN (f, g) =
     let
-	fun F NONE = NONE
-	  | F (SOME(x,xs)) =
-	    let
-		fun G NONE = F (getItem xs)
-		  | G (SOME(y,ys)) = SOME(y,make (fn () => G (getItem ys)))
-	    in
-		G (getItem (g x))
-	    end
+        fun F NONE = NONE
+          | F (SOME(x,xs)) =
+            let
+                fun G NONE = F (getItem xs)
+                  | G (SOME(y,ys)) = SOME(y,make (fn () => G (getItem ys)))
+            in
+                G (getItem (g x))
+            end
     in
-	fn x => make (fn () => F (getItem (f x)))
+        fn x => make (fn () => F (getItem (f x)))
     end
 
 fun op ORELSE (f, g) x =
     make (fn () =>
-	     case getItem (f x) of
-		 NONE => getItem (g x)
-	       | some => some)
+             case getItem (f x) of
+                 NONE => getItem (g x)
+               | some => some)
 
 fun op APPEND (f, g) x =
     let
-	fun copy s =
-	    case getItem s of
-		NONE => getItem (g x)
-	      | SOME(x,xs) => SOME(x,make (fn () => copy xs))
+        fun copy s =
+            case getItem s of
+                NONE => getItem (g x)
+              | SOME(x,xs) => SOME(x,make (fn () => copy xs))
     in
-	make (fn () => copy (f x))
+        make (fn () => copy (f x))
     end
 
 fun EVERY fs = List.foldr (op THEN) succeed fs
@@ -412,20 +412,20 @@
 
 fun TRY f x =
     make (fn () =>
-	     case getItem (f x) of
-		 NONE => SOME(x,Seq (Lazy.value NONE))
-	       | some => some)
+             case getItem (f x) of
+                 NONE => SOME(x,Seq (Lazy.value NONE))
+               | some => some)
 
 fun REPEAT f =
     let
-	fun rep qs x =
-	    case getItem (f x) of
-		NONE => SOME(x, make (fn () => repq qs))
-	      | SOME (x', q) => rep (q :: qs) x'
-	and repq [] = NONE
-	  | repq (q :: qs) =
+        fun rep qs x =
+            case getItem (f x) of
+                NONE => SOME(x, make (fn () => repq qs))
+              | SOME (x', q) => rep (q :: qs) x'
+        and repq [] = NONE
+          | repq (q :: qs) =
             case getItem q of
-		NONE => repq qs
+                NONE => repq qs
               | SOME (x, q) => rep (q :: qs) x
     in
      fn x => make (fn () => rep [] x)
@@ -435,63 +435,63 @@
 
 fun INTERVAL f i =
     let
-	fun F j =
-	    if i > j
-	    then single
-	    else op THEN (f j, F (j - 1))
+        fun F j =
+            if i > j
+            then single
+            else op THEN (f j, F (j - 1))
     in F end
 
 fun DETERM f x =
     make (fn () =>
-	     case getItem (f x) of
-		 NONE => NONE
-	       | SOME (x', _) => SOME(x',Seq (Lazy.value NONE)))
+             case getItem (f x) of
+                 NONE => NONE
+               | SOME (x', _) => SOME(x',Seq (Lazy.value NONE)))
 
 (*partial function as procedure*)
 fun try f x =
     make (fn () =>
-	     case Basics.try f x of
-		 SOME y => SOME(y,Seq (Lazy.value NONE))
-	       | NONE => NONE)
+             case Basics.try f x of
+                 SOME y => SOME(y,Seq (Lazy.value NONE))
+               | NONE => NONE)
 
 (*functional to print a sequence, up to "count" elements;
   the function prelem should print the element number and also the element*)
 fun print prelem count seq =
     let
-	fun pr k xq =
-	    if k > count
-	    then ()
-	    else
-		case getItem xq of
-		    NONE => ()
-		  | SOME (x, xq') =>
-		    (prelem k x;
-		     writeln "";
-		     pr (k + 1) xq')
+        fun pr k xq =
+            if k > count
+            then ()
+            else
+                case getItem xq of
+                    NONE => ()
+                  | SOME (x, xq') =>
+                    (prelem k x;
+                     writeln "";
+                     pr (k + 1) xq')
     in
-	pr 1 seq
+        pr 1 seq
     end
 
 (*accumulating a function over a sequence; this is lazy*)
 fun it_right f (xq, yq) =
     let
-	fun its s =
-	    make (fn () =>
-		     case getItem s of
-			 NONE => getItem yq
-		       | SOME (a, s') => getItem(f (a, its s')))
+        fun its s =
+            make (fn () =>
+                     case getItem s of
+                         NONE => getItem yq
+                       | SOME (a, s') => getItem(f (a, its s')))
     in
-	its xq
+        its xq
     end
 
 (*map over a sequence s1, append the sequence s2*)
 fun mapp f s1 s2 =
     let
-	fun F NONE = getItem s2
-	  | F (SOME(x,s1')) = SOME(f x,F' s1')
-	and F' s = make (fn () => F (getItem s))
+        fun F NONE = getItem s2
+          | F (SOME(x,s1')) = SOME(f x,F' s1')
+        and F' s = make (fn () => F (getItem s))
     in
-	F' s1
+        F' s1
     end
 
 (*turn a list of sequences into a sequence of lists*)
@@ -502,17 +502,17 @@
             NONE => NONE
           | SOME (x, xq') =>
             (case F xqs of
-		 NONE => NONE
-	       | SOME (xs, xsq) =>
-		 let
-		     fun G s =
-			 make (fn () =>
-				  case getItem s of
-				      NONE => F (xq' :: xqs)
-				    | SOME(ys,ysq) => SOME(x::ys,G ysq))
-		 in
-		     SOME (x :: xs, G xsq)
-		 end)
+                 NONE => NONE
+               | SOME (xs, xsq) =>
+                 let
+                     fun G s =
+                         make (fn () =>
+                                  case getItem s of
+                                      NONE => F (xq' :: xqs)
+                                    | SOME(ys,ysq) => SOME(x::ys,G ysq))
+                 in
+                     SOME (x :: xs, G xsq)
+                 end)
 in
 fun commute xqs = make (fn () => F xqs)
 end
@@ -523,24 +523,24 @@
     if n <= 0
     then ([], xq)
     else
-	case getItem xq of
-	    NONE => ([], xq)
-	  | SOME (x, xq') => apfst (Basics.cons x) (chop (n - 1, xq'))
+        case getItem xq of
+            NONE => ([], xq)
+          | SOME (x, xq') => apfst (Basics.cons x) (chop (n - 1, xq'))
 
 fun foldl f e s =
     let
-	fun F k NONE = k e
-	  | F k (SOME(x,s)) = F (fn y => k (f(x,y))) (getItem s)
+        fun F k NONE = k e
+          | F k (SOME(x,s)) = F (fn y => k (f(x,y))) (getItem s)
     in
-	F (fn x => x) (getItem s)
+        F (fn x => x) (getItem s)
     end
 
 fun foldr f e s =
     let
-	fun F e NONE = e
-	  | F e (SOME(x,s)) = F (f(x,e)) (getItem s)
+        fun F e NONE = e
+          | F e (SOME(x,s)) = F (f(x,e)) (getItem s)
     in
-	F e (getItem s)
+        F e (getItem s)
     end
 
 fun fromString s = of_list (explode s)
--- a/src/HOL/Import/mono_scan.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/mono_scan.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -15,13 +15,13 @@
     type 'a scanner = seq -> 'a * seq
 
     val :--      : 'a scanner * ('a -> 'b scanner)
-		   -> ('a*'b) scanner
+                   -> ('a*'b) scanner
     val --       : 'a scanner * 'b scanner -> ('a * 'b) scanner
     val >>       : 'a scanner * ('a -> 'b) -> 'b scanner
     val --|      : 'a scanner * 'b scanner -> 'a scanner
     val |--      : 'a scanner * 'b scanner -> 'b scanner
     val ^^       : string scanner * string scanner
-		   -> string scanner 
+                   -> string scanner 
     val ||       : 'a scanner * 'a scanner -> 'a scanner
     val one      : (item -> bool) -> item scanner
     val anyone   : item scanner
@@ -50,7 +50,7 @@
     val scan_nat : int scanner
 
     val this : item list -> item list scanner
-    val this_string : string -> string scanner					    
+    val this_string : string -> string scanner                                      
 
 end    
 
@@ -73,48 +73,48 @@
 
 fun (sc1 :-- sc2) toks =
     let
-	val (x,toks2) = sc1 toks
-	val (y,toks3) = sc2 x toks2
+        val (x,toks2) = sc1 toks
+        val (y,toks3) = sc2 x toks2
     in
-	((x,y),toks3)
+        ((x,y),toks3)
     end
 
 fun (sc1 -- sc2) toks =
     let
-	val (x,toks2) = sc1 toks
-	val (y,toks3) = sc2 toks2
+        val (x,toks2) = sc1 toks
+        val (y,toks3) = sc2 toks2
     in
-	((x,y),toks3)
+        ((x,y),toks3)
     end
 
 fun (sc >> f) toks =
     let
-	val (x,toks2) = sc toks
+        val (x,toks2) = sc toks
     in
-	(f x,toks2)
+        (f x,toks2)
     end
 
 fun (sc1 --| sc2) toks =
     let
-	val (x,toks2) = sc1 toks
-	val (_,toks3) = sc2 toks2
+        val (x,toks2) = sc1 toks
+        val (_,toks3) = sc2 toks2
     in
-	(x,toks3)
+        (x,toks3)
     end
 
 fun (sc1 |-- sc2) toks =
     let
-	val (_,toks2) = sc1 toks
+        val (_,toks2) = sc1 toks
     in
-	sc2 toks2
+        sc2 toks2
     end
 
 fun (sc1 ^^ sc2) toks =
     let
-	val (x,toks2) = sc1 toks
-	val (y,toks3) = sc2 toks2
+        val (x,toks2) = sc1 toks
+        val (y,toks3) = sc2 toks2
     in
-	(x^y,toks3)
+        (x^y,toks3)
     end
 
 fun (sc1 || sc2) toks =
@@ -129,22 +129,22 @@
 
 fun any p toks =
     case pull toks of
-	NONE =>  ([],toks)
+        NONE =>  ([],toks)
       | SOME(x,toks2) => if p x
-			 then
-			     let
-				 val (xs,toks3) = any p toks2
-			     in
-				 (x::xs,toks3)
-			     end
-			 else ([],toks)
+                         then
+                             let
+                                 val (xs,toks3) = any p toks2
+                             in
+                                 (x::xs,toks3)
+                             end
+                         else ([],toks)
 
 fun any1 p toks =
     let
-	val (x,toks2) = one p toks
-	val (xs,toks3) = any p toks2
+        val (x,toks2) = one p toks
+        val (xs,toks3) = any p toks2
     in
-	(x::xs,toks3)
+        (x::xs,toks3)
     end
 
 fun optional sc def =  sc || succeed def
@@ -153,16 +153,16 @@
 (*
 fun repeat sc =
     let
-	fun R toks =
-	    let
-		val (x,toks2) = sc toks
-		val (xs,toks3) = R toks2
-	    in
-		(x::xs,toks3)
-	    end
-	    handle SyntaxError => ([],toks)
+        fun R toks =
+            let
+                val (x,toks2) = sc toks
+                val (xs,toks3) = R toks2
+            in
+                (x::xs,toks3)
+            end
+            handle SyntaxError => ([],toks)
     in
-	R
+        R
     end
 *)
 
@@ -174,44 +174,44 @@
  *)
 fun repeat sc =
     let
-	fun R xs toks =
-	    case SOME (sc toks) handle SyntaxError => NONE of
-		SOME (x,toks2) => R (x::xs) toks2
-	      | NONE => (List.rev xs,toks)
+        fun R xs toks =
+            case SOME (sc toks) handle SyntaxError => NONE of
+                SOME (x,toks2) => R (x::xs) toks2
+              | NONE => (List.rev xs,toks)
     in
-	R []
+        R []
     end
 
 fun repeat1 sc toks =
     let
-	val (x,toks2) = sc toks
-	val (xs,toks3) = repeat sc toks2
+        val (x,toks2) = sc toks
+        val (xs,toks3) = repeat sc toks2
     in
-	(x::xs,toks3)
+        (x::xs,toks3)
     end
 
 fun repeat_fixed n sc =
     let
-	fun R n xs toks =
-	    if (n <= 0) then (List.rev xs, toks)
-	    else case (sc toks) of (x, toks2) => R (n-1) (x::xs) toks2
+        fun R n xs toks =
+            if (n <= 0) then (List.rev xs, toks)
+            else case (sc toks) of (x, toks2) => R (n-1) (x::xs) toks2
     in
-	R n []
+        R n []
     end
 
 fun ahead (sc:'a->'b*'a) toks = (#1 (sc toks),toks)
 
 fun unless test sc toks =
     let
-	val test_failed = (test toks;false) handle SyntaxError => true
+        val test_failed = (test toks;false) handle SyntaxError => true
     in
-	if test_failed
-	then sc toks
-	else raise SyntaxError
+        if test_failed
+        then sc toks
+        else raise SyntaxError
     end
 
 fun !! f sc toks = (sc toks
-		    handle SyntaxError => raise Fail (f toks))
+                    handle SyntaxError => raise Fail (f toks))
 
 end
 
--- a/src/HOL/Import/mono_seq.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/mono_seq.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -38,24 +38,24 @@
 fun take_at_most s n = 
     if n <= 0 then [] else
     case pull s of 
-	NONE => []
+        NONE => []
       | SOME (x,s') => x::(take_at_most s' (n-1))
 
 fun length' s n = 
     case pull s of
-	NONE => n
+        NONE => n
       | SOME (_, s') => length' s' (n+1)
 
 fun length s = length' s 0
-		
+                
 end  
 
 
 structure StringScannerSeq : 
-	  sig 
-	      include MONO_EXTENDED_SCANNER_SEQ 
-	      val fromString : string -> seq
-	  end
+          sig 
+              include MONO_EXTENDED_SCANNER_SEQ 
+              val fromString : string -> seq
+          end
   =
 struct
 
--- a/src/HOL/Import/replay.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/replay.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Import/replay.ML
-    ID:         $Id$
     Author:     Sebastian Skalberg (TU Muenchen)
 *)
 
@@ -17,406 +16,406 @@
 
 fun replay_proof int_thms thyname thmname prf thy =
     let
-	val _ = ImportRecorder.start_replay_proof thyname thmname 
-	fun rp (PRefl tm) thy = P.REFL tm thy
-	  | rp (PInstT(p,lambda)) thy =
-	    let
-		val (thy',th) = rp' p thy
-	    in
-		P.INST_TYPE lambda th thy'
-	    end
-	  | rp (PSubst(prfs,ctxt,prf)) thy =
-	    let
-		val (thy',ths) = fold_rev (fn p => fn (thy, ths) =>
-					   let
-					       val (thy',th) = rp' p thy
-					   in
-					       (thy',th::ths)
-					   end) prfs (thy,[])
-		val (thy'',th) = rp' prf thy'
-	    in
-		P.SUBST ths ctxt th thy''
-	    end
-	  | rp (PAbs(prf,v)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.ABS v th thy'
-	    end
-	  | rp (PDisch(prf,tm)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.DISCH tm th thy'
-	    end
-	  | rp (PMp(prf1,prf2)) thy =
-	    let
-		val (thy1,th1) = rp' prf1 thy
-		val (thy2,th2) = rp' prf2 thy1
-	    in
-		P.MP th1 th2 thy2
-	    end
-	  | rp (PHyp tm) thy = P.ASSUME tm thy
-	  | rp (PDef(seg,name,rhs)) thy =
-	    (case P.get_def seg name rhs thy of
-		 (thy',SOME res) => (thy',res)
-	       | (thy',NONE) => 
-		 if seg = thyname
-		 then P.new_definition seg name rhs thy'
-		 else raise ERR "replay_proof" ("Too late for term definition: "^seg^" != "^thyname))
-	  | rp (POracle(tg,asl,c)) thy = (*P.mk_oracle_thm tg (map be_contract asl,c)*) NY "ORACLE"
-	  | rp (PSpec(prf,tm)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.SPEC tm th thy'
-	    end
-	  | rp (PInst(prf,theta)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.INST theta th thy'
-	    end
-	  | rp (PGen(prf,v)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-		val p = P.GEN v th thy'
-	    in
-		p
-	    end
-	  | rp (PGenAbs(prf,opt,vl)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.GEN_ABS opt vl th thy'
-	    end
-	  | rp (PImpAS(prf1,prf2)) thy =
-	    let
-		val (thy1,th1) = rp' prf1 thy
-		val (thy2,th2) = rp' prf2 thy1
-	    in
-		P.IMP_ANTISYM th1 th2 thy2
-	    end
-	  | rp (PSym prf) thy =
-	    let
-		val (thy1,th) = rp' prf thy
-	    in
-		P.SYM th thy1
-	    end
-	  | rp (PTrans(prf1,prf2)) thy =
-	    let
-		val (thy1,th1) = rp' prf1 thy
-		val (thy2,th2) = rp' prf2 thy1
-	    in
-		P.TRANS th1 th2 thy2
-	    end
-	  | rp (PComb(prf1,prf2)) thy =
-	    let
-		val (thy1,th1) = rp' prf1 thy
-		val (thy2,th2) = rp' prf2 thy1
-	    in
-		P.COMB th1 th2 thy2
-	    end
-	  | rp (PEqMp(prf1,prf2)) thy =
-	    let
-		val (thy1,th1) = rp' prf1 thy
-		val (thy2,th2) = rp' prf2 thy1
-	    in
-		P.EQ_MP th1 th2 thy2
-	    end
-	  | rp (PEqImp prf) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.EQ_IMP_RULE th thy'
-	    end
-	  | rp (PExists(prf,ex,wit)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.EXISTS ex wit th thy'
-	    end
-	  | rp (PChoose(v,prf1,prf2)) thy =
-	    let
-		val (thy1,th1) = rp' prf1 thy
-		val (thy2,th2) = rp' prf2 thy1
-	    in
-		P.CHOOSE v th1 th2 thy2
-	    end
-	  | rp (PConj(prf1,prf2)) thy =
-	    let
-		val (thy1,th1) = rp' prf1 thy
-		val (thy2,th2) = rp' prf2 thy1
-	    in
-		P.CONJ th1 th2 thy2
-	    end
-	  | rp (PConjunct1 prf) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.CONJUNCT1 th thy'
-	    end
-	  | rp (PConjunct2 prf) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.CONJUNCT2 th thy'
-	    end
-	  | rp (PDisj1(prf,tm)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.DISJ1 th tm thy'
-	    end
-	  | rp (PDisj2(prf,tm)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.DISJ2 tm th thy'
-	    end
-	  | rp (PDisjCases(prf,prf1,prf2)) thy =
-	    let
-		val (thy',th)  = rp' prf  thy
-		val (thy1,th1) = rp' prf1 thy'
-		val (thy2,th2) = rp' prf2 thy1
-	    in
-		P.DISJ_CASES th th1 th2 thy2
-	    end
-	  | rp (PNotI prf) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.NOT_INTRO th thy'
-	    end
-	  | rp (PNotE prf) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.NOT_ELIM th thy'
-	    end
-	  | rp (PContr(prf,tm)) thy =
-	    let
-		val (thy',th) = rp' prf thy
-	    in
-		P.CCONTR tm th thy'
-	    end
-	  | rp (PTmSpec _) _ = raise ERR "rp" "Shouldn't reach here (PTmSpec)"
-	  | rp (PTyDef _) _ = raise ERR "rp" "Shouldn't reach here (PTyDef)"
-	  | rp (PTyIntro _) _ = raise ERR "rp" "Shouldn't reach here (PTyIntro)"
-	  | rp PDisk _ = raise ERR "rp" "Shouldn't reach here (PDisk)"
-	  | rp _ _ = raise ERR "rp" "What the hell is this? Which case did I forget?"
-	and rp' p thy =
-	    let
-		val pc = content_of p
-	    in
-		case pc of
-		    PDisk => (case disk_info_of p of
-				  SOME(thyname',thmname) =>
-				  (case Int.fromString thmname of
-				       SOME i =>
-				       if thyname' = thyname
-				       then
-					   (case Array.sub(int_thms,i-1) of
-						NONE =>
-						let
-						    val (thy',th) = rp' (snd (import_proof thyname' thmname thy) thy) thy
-						    val _ = Array.update(int_thms,i-1,SOME th)
-						in
-						    (thy',th)
-						end
-					      | SOME th => (thy,th))
-				       else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")
-				     | NONE => 
-				       (case P.get_thm thyname' thmname thy of
-					    (thy',SOME res) => (thy',res)
-					  | (thy',NONE) => 
-					    if thyname' = thyname
-					    then
-						let
-						    val _ = writeln ("Found no " ^ thmname ^ " theorem, replaying...")
-						    val (f_opt,prf) = import_proof thyname' thmname thy'
-						    val prf = prf thy'
-						    val (thy',th) = replay_proof int_thms thyname' thmname prf thy'
+        val _ = ImportRecorder.start_replay_proof thyname thmname 
+        fun rp (PRefl tm) thy = P.REFL tm thy
+          | rp (PInstT(p,lambda)) thy =
+            let
+                val (thy',th) = rp' p thy
+            in
+                P.INST_TYPE lambda th thy'
+            end
+          | rp (PSubst(prfs,ctxt,prf)) thy =
+            let
+                val (thy',ths) = fold_rev (fn p => fn (thy, ths) =>
+                                           let
+                                               val (thy',th) = rp' p thy
+                                           in
+                                               (thy',th::ths)
+                                           end) prfs (thy,[])
+                val (thy'',th) = rp' prf thy'
+            in
+                P.SUBST ths ctxt th thy''
+            end
+          | rp (PAbs(prf,v)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.ABS v th thy'
+            end
+          | rp (PDisch(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.DISCH tm th thy'
+            end
+          | rp (PMp(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                P.MP th1 th2 thy2
+            end
+          | rp (PHyp tm) thy = P.ASSUME tm thy
+          | rp (PDef(seg,name,rhs)) thy =
+            (case P.get_def seg name rhs thy of
+                 (thy',SOME res) => (thy',res)
+               | (thy',NONE) => 
+                 if seg = thyname
+                 then P.new_definition seg name rhs thy'
+                 else raise ERR "replay_proof" ("Too late for term definition: "^seg^" != "^thyname))
+          | rp (POracle(tg,asl,c)) thy = (*P.mk_oracle_thm tg (map be_contract asl,c)*) NY "ORACLE"
+          | rp (PSpec(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.SPEC tm th thy'
+            end
+          | rp (PInst(prf,theta)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.INST theta th thy'
+            end
+          | rp (PGen(prf,v)) thy =
+            let
+                val (thy',th) = rp' prf thy
+                val p = P.GEN v th thy'
+            in
+                p
+            end
+          | rp (PGenAbs(prf,opt,vl)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.GEN_ABS opt vl th thy'
+            end
+          | rp (PImpAS(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                P.IMP_ANTISYM th1 th2 thy2
+            end
+          | rp (PSym prf) thy =
+            let
+                val (thy1,th) = rp' prf thy
+            in
+                P.SYM th thy1
+            end
+          | rp (PTrans(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                P.TRANS th1 th2 thy2
+            end
+          | rp (PComb(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                P.COMB th1 th2 thy2
+            end
+          | rp (PEqMp(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                P.EQ_MP th1 th2 thy2
+            end
+          | rp (PEqImp prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.EQ_IMP_RULE th thy'
+            end
+          | rp (PExists(prf,ex,wit)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.EXISTS ex wit th thy'
+            end
+          | rp (PChoose(v,prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                P.CHOOSE v th1 th2 thy2
+            end
+          | rp (PConj(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                P.CONJ th1 th2 thy2
+            end
+          | rp (PConjunct1 prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.CONJUNCT1 th thy'
+            end
+          | rp (PConjunct2 prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.CONJUNCT2 th thy'
+            end
+          | rp (PDisj1(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.DISJ1 th tm thy'
+            end
+          | rp (PDisj2(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.DISJ2 tm th thy'
+            end
+          | rp (PDisjCases(prf,prf1,prf2)) thy =
+            let
+                val (thy',th)  = rp' prf  thy
+                val (thy1,th1) = rp' prf1 thy'
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                P.DISJ_CASES th th1 th2 thy2
+            end
+          | rp (PNotI prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.NOT_INTRO th thy'
+            end
+          | rp (PNotE prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.NOT_ELIM th thy'
+            end
+          | rp (PContr(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                P.CCONTR tm th thy'
+            end
+          | rp (PTmSpec _) _ = raise ERR "rp" "Shouldn't reach here (PTmSpec)"
+          | rp (PTyDef _) _ = raise ERR "rp" "Shouldn't reach here (PTyDef)"
+          | rp (PTyIntro _) _ = raise ERR "rp" "Shouldn't reach here (PTyIntro)"
+          | rp PDisk _ = raise ERR "rp" "Shouldn't reach here (PDisk)"
+          | rp _ _ = raise ERR "rp" "What the hell is this? Which case did I forget?"
+        and rp' p thy =
+            let
+                val pc = content_of p
+            in
+                case pc of
+                    PDisk => (case disk_info_of p of
+                                  SOME(thyname',thmname) =>
+                                  (case Int.fromString thmname of
+                                       SOME i =>
+                                       if thyname' = thyname
+                                       then
+                                           (case Array.sub(int_thms,i-1) of
+                                                NONE =>
+                                                let
+                                                    val (thy',th) = rp' (snd (import_proof thyname' thmname thy) thy) thy
+                                                    val _ = Array.update(int_thms,i-1,SOME th)
+                                                in
+                                                    (thy',th)
+                                                end
+                                              | SOME th => (thy,th))
+                                       else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")
+                                     | NONE => 
+                                       (case P.get_thm thyname' thmname thy of
+                                            (thy',SOME res) => (thy',res)
+                                          | (thy',NONE) => 
+                                            if thyname' = thyname
+                                            then
+                                                let
+                                                    val _ = writeln ("Found no " ^ thmname ^ " theorem, replaying...")
+                                                    val (f_opt,prf) = import_proof thyname' thmname thy'
+                                                    val prf = prf thy'
+                                                    val (thy',th) = replay_proof int_thms thyname' thmname prf thy'
                                                     val _ = writeln ("Successfully finished replaying "^thmname^" !")
-						in
-						    case content_of prf of
-							PTmSpec _ => (thy',th)
-						      | PTyDef  _ => (thy',th)
-						      | PTyIntro _ => (thy',th)
-						      | _ => P.store_thm thyname' thmname th thy'
-						end
-					    else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")))
-				| NONE => raise ERR "rp'.PDisk" "Not enough information")
-		  | PAxm(name,c) =>
-		    (case P.get_axiom thyname name thy of
-			    (thy',SOME res) => (thy',res)
-			  | (thy',NONE) => P.new_axiom name c thy')
-		  | PTmSpec(seg,names,prf') =>
-		    let
-			val (thy',th) = rp' prf' thy
-		    in
-			P.new_specification seg thmname names th thy'
-		    end
-		  | PTyDef(seg,name,prf') =>
-		    let
-			val (thy',th) = rp' prf' thy
-		    in
-			P.new_type_definition seg thmname name th thy'
-		    end
-		  | PTyIntro(seg,name,abs_name,rep_name,P,t,prf') =>
-		    let
-			val (thy',th) = rp' prf' thy
-		    in
-			P.type_introduction seg thmname name abs_name rep_name (P,t) th thy'
-		    end
-		  | _ => rp pc thy
-	    end
+                                                in
+                                                    case content_of prf of
+                                                        PTmSpec _ => (thy',th)
+                                                      | PTyDef  _ => (thy',th)
+                                                      | PTyIntro _ => (thy',th)
+                                                      | _ => P.store_thm thyname' thmname th thy'
+                                                end
+                                            else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")))
+                                | NONE => raise ERR "rp'.PDisk" "Not enough information")
+                  | PAxm(name,c) =>
+                    (case P.get_axiom thyname name thy of
+                            (thy',SOME res) => (thy',res)
+                          | (thy',NONE) => P.new_axiom name c thy')
+                  | PTmSpec(seg,names,prf') =>
+                    let
+                        val (thy',th) = rp' prf' thy
+                    in
+                        P.new_specification seg thmname names th thy'
+                    end
+                  | PTyDef(seg,name,prf') =>
+                    let
+                        val (thy',th) = rp' prf' thy
+                    in
+                        P.new_type_definition seg thmname name th thy'
+                    end
+                  | PTyIntro(seg,name,abs_name,rep_name,P,t,prf') =>
+                    let
+                        val (thy',th) = rp' prf' thy
+                    in
+                        P.type_introduction seg thmname name abs_name rep_name (P,t) th thy'
+                    end
+                  | _ => rp pc thy
+            end
     in
-	let
-	    val x = rp' prf thy handle e => (writeln "Exception in replay_proof"; OldGoals.print_exn e)
-	    val _ = ImportRecorder.stop_replay_proof thyname thmname
-	in
-	    x
-	end
+        let
+            val x = rp' prf thy handle e => (writeln "Exception in replay_proof"; OldGoals.print_exn e)
+            val _ = ImportRecorder.stop_replay_proof thyname thmname
+        in
+            x
+        end
     end handle x => (ImportRecorder.abort_replay_proof thyname thmname; raise x)
 
 fun setup_int_thms thyname thy =
     let
-	val fname =
-	    case P.get_proof_dir thyname thy of
-		SOME p => OS.Path.joinDirFile {dir=p,file=OS.Path.joinBaseExt{base = "facts",ext=SOME "lst"}}
-	      | NONE => error "Cannot find proof files"
-	val is = TextIO.openIn fname
-	val (num_int_thms,facts) =
-	    let
-		fun get_facts facts =
-		    case TextIO.inputLine is of
-			NONE => (case facts of
-				   i::facts => (valOf (Int.fromString i),map P.protect_factname (rev facts))
-				 | _ => raise ERR "replay_thm" "Bad facts.lst file")
-		      | SOME fact => get_facts ((String.substring(fact,0,String.size fact -1 ))::facts)
-	    in
-		get_facts []
-	    end
-	val _ = TextIO.closeIn is
-	val int_thms = Array.array(num_int_thms,NONE:thm option)
+        val fname =
+            case P.get_proof_dir thyname thy of
+                SOME p => OS.Path.joinDirFile {dir=p,file=OS.Path.joinBaseExt{base = "facts",ext=SOME "lst"}}
+              | NONE => error "Cannot find proof files"
+        val is = TextIO.openIn fname
+        val (num_int_thms,facts) =
+            let
+                fun get_facts facts =
+                    case TextIO.inputLine is of
+                        NONE => (case facts of
+                                   i::facts => (valOf (Int.fromString i),map P.protect_factname (rev facts))
+                                 | _ => raise ERR "replay_thm" "Bad facts.lst file")
+                      | SOME fact => get_facts ((String.substring(fact,0,String.size fact -1 ))::facts)
+            in
+                get_facts []
+            end
+        val _ = TextIO.closeIn is
+        val int_thms = Array.array(num_int_thms,NONE:thm option)
     in
-	(int_thms,facts)
+        (int_thms,facts)
     end
 
 fun import_single_thm thyname int_thms thmname thy =
     let
-	fun replay_fact (thmname,thy) =
-	    let
-		val prf = mk_proof PDisk
-		val _ = set_disk_info_of prf thyname thmname
+        fun replay_fact (thmname,thy) =
+            let
+                val prf = mk_proof PDisk
+                val _ = set_disk_info_of prf thyname thmname
                 val _ = writeln ("Replaying "^thmname^" ...")
-		val p = fst (replay_proof int_thms thyname thmname prf thy)
-	    in
-		p
-	    end
+                val p = fst (replay_proof int_thms thyname thmname prf thy)
+            in
+                p
+            end
     in
-	replay_fact (thmname,thy)
+        replay_fact (thmname,thy)
     end
 
 fun replay_chached_thm int_thms thyname thmname =
     let
-	fun th_of thy = setmp quick_and_dirty true (SkipProof.make_thm thy)
-	fun err msg = raise ERR "replay_cached_thm" msg
-	val _ = writeln ("Replaying (from cache) "^thmname^" ...")
-	fun rps [] thy = thy
-	  | rps (t::ts) thy = rps ts (rp t thy)
-	and rp (ThmEntry (thyname', thmname', aborted, History history)) thy = rps history thy	    
-	  | rp (DeltaEntry ds) thy = fold delta ds thy
-	and delta (Specification (names, th)) thy = 
-	    fst (Choice_Specification.add_specification NONE names (thy,th_of thy th))
-	  | delta (Hol_mapping (thyname, thmname, isaname)) thy = 
-	    add_hol4_mapping thyname thmname isaname thy
-	  | delta (Hol_pending (thyname, thmname, th)) thy = 
-	    add_hol4_pending thyname thmname ([], th_of thy th) thy
-	  | delta (Consts cs) thy = Sign.add_consts_i (map (fn (c, T, mx) => (Binding.name c, T, mx)) cs) thy
-	  | delta (Hol_const_mapping (thyname, constname, fullcname)) thy = 
-	    add_hol4_const_mapping thyname constname true fullcname thy
-	  | delta (Hol_move (fullname, moved_thmname)) thy = 
-	    add_hol4_move fullname moved_thmname thy
-	  | delta (Defs (thmname, eq)) thy =
-	    snd (PureThy.add_defs false [((Binding.name thmname, eq), [])] thy)
-	  | delta (Hol_theorem (thyname, thmname, th)) thy =
-	    add_hol4_theorem thyname thmname ([], th_of thy th) thy
-	  | delta (Typedef (thmname, (t, vs, mx), c, repabs, th)) thy = 
-	    snd (Typedef.add_typedef false (Option.map Binding.name thmname) (Binding.name t, vs, mx) c
+        fun th_of thy = setmp quick_and_dirty true (SkipProof.make_thm thy)
+        fun err msg = raise ERR "replay_cached_thm" msg
+        val _ = writeln ("Replaying (from cache) "^thmname^" ...")
+        fun rps [] thy = thy
+          | rps (t::ts) thy = rps ts (rp t thy)
+        and rp (ThmEntry (thyname', thmname', aborted, History history)) thy = rps history thy      
+          | rp (DeltaEntry ds) thy = fold delta ds thy
+        and delta (Specification (names, th)) thy = 
+            fst (Choice_Specification.add_specification NONE names (thy,th_of thy th))
+          | delta (Hol_mapping (thyname, thmname, isaname)) thy = 
+            add_hol4_mapping thyname thmname isaname thy
+          | delta (Hol_pending (thyname, thmname, th)) thy = 
+            add_hol4_pending thyname thmname ([], th_of thy th) thy
+          | delta (Consts cs) thy = Sign.add_consts_i (map (fn (c, T, mx) => (Binding.name c, T, mx)) cs) thy
+          | delta (Hol_const_mapping (thyname, constname, fullcname)) thy = 
+            add_hol4_const_mapping thyname constname true fullcname thy
+          | delta (Hol_move (fullname, moved_thmname)) thy = 
+            add_hol4_move fullname moved_thmname thy
+          | delta (Defs (thmname, eq)) thy =
+            snd (PureThy.add_defs false [((Binding.name thmname, eq), [])] thy)
+          | delta (Hol_theorem (thyname, thmname, th)) thy =
+            add_hol4_theorem thyname thmname ([], th_of thy th) thy
+          | delta (Typedef (thmname, (t, vs, mx), c, repabs, th)) thy = 
+            snd (Typedef.add_typedef false (Option.map Binding.name thmname) (Binding.name t, vs, mx) c
         (Option.map (pairself Binding.name) repabs) (rtac (th_of thy th) 1) thy)
-	  | delta (Hol_type_mapping (thyname, tycname, fulltyname)) thy =  
-	    add_hol4_type_mapping thyname tycname true fulltyname thy
-	  | delta (Indexed_theorem (i, th)) thy = 
-	    (Array.update (int_thms,i-1,SOME (P.to_hol_thm (th_of thy th))); thy)	    	    
+          | delta (Hol_type_mapping (thyname, tycname, fulltyname)) thy =  
+            add_hol4_type_mapping thyname tycname true fulltyname thy
+          | delta (Indexed_theorem (i, th)) thy = 
+            (Array.update (int_thms,i-1,SOME (P.to_hol_thm (th_of thy th))); thy)                   
           | delta (Protect_varname (s,t)) thy = (P.replay_protect_varname s t; thy)
-	  | delta (Dump s) thy = P.replay_add_dump s thy
+          | delta (Dump s) thy = P.replay_add_dump s thy
     in
-	rps
+        rps
     end
 
 fun import_thms thyname int_thms thmnames thy =
     let
-	fun zip names [] = ([], names)
-	  | zip [] _ = ([], [])
-	  | zip (thmname::names) ((ThmEntry (entry as (thyname',thmname',aborted,History history)))::ys) = 
-	    if thyname = thyname' andalso thmname = thmname' then
-		(if aborted then ([], thmname::names) else 
-		 let
-		     val _ = writeln ("theorem is in-sync: "^thmname)
-		     val (cached,normal) = zip names ys
-		 in
-		     (entry::cached, normal)
-		 end)
-	    else
-		let
-		    val _ = writeln ("cached theorems are not in-sync,  expected: "^thmname^", found: "^thmname')
-		    val _ = writeln ("proceeding with next uncached theorem...")
-		in
-		    ([], thmname::names)
-		end
-	(*	raise ERR "import_thms" ("cached theorems are not in-sync, expected: "^thmname^", found: "^thmname')*)
-	  | zip (thmname::_) (DeltaEntry _ :: _) = 
- 	    raise ERR "import_thms" ("expected theorem '"^thmname^"', but found a delta")
-	fun zip' xs (History ys) = 
-	    let
-		val _ = writeln ("length of xs: "^(string_of_int (length xs)))
-		val _ = writeln ("length of history: "^(string_of_int (length ys)))
-	    in
-		zip xs ys
-	    end
-	fun replay_fact thmname thy = 
-	    let
-		val prf = mk_proof PDisk	
-		val _ = set_disk_info_of prf thyname thmname
+        fun zip names [] = ([], names)
+          | zip [] _ = ([], [])
+          | zip (thmname::names) ((ThmEntry (entry as (thyname',thmname',aborted,History history)))::ys) = 
+            if thyname = thyname' andalso thmname = thmname' then
+                (if aborted then ([], thmname::names) else 
+                 let
+                     val _ = writeln ("theorem is in-sync: "^thmname)
+                     val (cached,normal) = zip names ys
+                 in
+                     (entry::cached, normal)
+                 end)
+            else
+                let
+                    val _ = writeln ("cached theorems are not in-sync,  expected: "^thmname^", found: "^thmname')
+                    val _ = writeln ("proceeding with next uncached theorem...")
+                in
+                    ([], thmname::names)
+                end
+        (*      raise ERR "import_thms" ("cached theorems are not in-sync, expected: "^thmname^", found: "^thmname')*)
+          | zip (thmname::_) (DeltaEntry _ :: _) = 
+            raise ERR "import_thms" ("expected theorem '"^thmname^"', but found a delta")
+        fun zip' xs (History ys) = 
+            let
+                val _ = writeln ("length of xs: "^(string_of_int (length xs)))
+                val _ = writeln ("length of history: "^(string_of_int (length ys)))
+            in
+                zip xs ys
+            end
+        fun replay_fact thmname thy = 
+            let
+                val prf = mk_proof PDisk        
+                val _ = set_disk_info_of prf thyname thmname
                 val _ = writeln ("Replaying "^thmname^" ...")
-		val p = fst (replay_proof int_thms thyname thmname prf thy)
-	    in
-		p
-	    end
-	fun replay_cache (_,thmname, _, History history) thy = replay_chached_thm int_thms thyname thmname history thy
-	val (cached, normal) = zip' thmnames (ImportRecorder.get_history ())
-	val _ = ImportRecorder.set_history (History (map ThmEntry cached))
-	val res_thy = fold replay_fact normal (fold replay_cache cached thy)
+                val p = fst (replay_proof int_thms thyname thmname prf thy)
+            in
+                p
+            end
+        fun replay_cache (_,thmname, _, History history) thy = replay_chached_thm int_thms thyname thmname history thy
+        val (cached, normal) = zip' thmnames (ImportRecorder.get_history ())
+        val _ = ImportRecorder.set_history (History (map ThmEntry cached))
+        val res_thy = fold replay_fact normal (fold replay_cache cached thy)
     in
-	res_thy
+        res_thy
     end
 
 fun import_thm thyname thmname thy =
     let
-	val int_thms = fst (setup_int_thms thyname thy)
-	fun replay_fact (thmname,thy) =
-	    let
-		val prf = mk_proof PDisk	
-		val _ = set_disk_info_of prf thyname thmname 	    
+        val int_thms = fst (setup_int_thms thyname thy)
+        fun replay_fact (thmname,thy) =
+            let
+                val prf = mk_proof PDisk        
+                val _ = set_disk_info_of prf thyname thmname        
                 val _ = writeln ("Replaying "^thmname^" ...")
-		val p = fst (replay_proof int_thms thyname thmname prf thy)
-	    in 
-		p
-	    end
+                val p = fst (replay_proof int_thms thyname thmname prf thy)
+            in 
+                p
+            end
     in
-	replay_fact (thmname,thy)
+        replay_fact (thmname,thy)
     end
 
 end
--- a/src/HOL/Import/scan.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/scan.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Import/scan.ML
-    ID:         $Id$
     Author:     Sebastian Skalberg, TU Muenchen / Steven Obua, TU Muenchen
 
 Scanner combinators for sequences.
@@ -15,13 +14,13 @@
     type ('a,'b) scanner = 'a seq -> 'b * 'a seq
 
     val :--      : ('a,'b) scanner * ('b -> ('a,'c) scanner)
-		   -> ('a,'b*'c) scanner
+                   -> ('a,'b*'c) scanner
     val --       : ('a,'b) scanner * ('a,'c) scanner -> ('a,'b*'c) scanner
     val >>       : ('a,'b) scanner * ('b -> 'c) -> ('a,'c) scanner
     val --|      : ('a,'b) scanner * ('a,'c) scanner -> ('a,'b) scanner
     val |--      : ('a,'b) scanner * ('a,'c) scanner -> ('a,'c) scanner
     val ^^       : ('a,string) scanner * ('a,string) scanner
-		   -> ('a,string) scanner 
+                   -> ('a,string) scanner 
     val ||       : ('a,'b) scanner * ('a,'b) scanner -> ('a,'b) scanner
     val one      : ('a -> bool) -> ('a,'a) scanner
     val anyone   : ('a,'a) scanner
@@ -63,48 +62,48 @@
 
 fun (sc1 :-- sc2) toks =
     let
-	val (x,toks2) = sc1 toks
-	val (y,toks3) = sc2 x toks2
+        val (x,toks2) = sc1 toks
+        val (y,toks3) = sc2 x toks2
     in
-	((x,y),toks3)
+        ((x,y),toks3)
     end
 
 fun (sc1 -- sc2) toks =
     let
-	val (x,toks2) = sc1 toks
-	val (y,toks3) = sc2 toks2
+        val (x,toks2) = sc1 toks
+        val (y,toks3) = sc2 toks2
     in
-	((x,y),toks3)
+        ((x,y),toks3)
     end
 
 fun (sc >> f) toks =
     let
-	val (x,toks2) = sc toks
+        val (x,toks2) = sc toks
     in
-	(f x,toks2)
+        (f x,toks2)
     end
 
 fun (sc1 --| sc2) toks =
     let
-	val (x,toks2) = sc1 toks
-	val (_,toks3) = sc2 toks2
+        val (x,toks2) = sc1 toks
+        val (_,toks3) = sc2 toks2
     in
-	(x,toks3)
+        (x,toks3)
     end
 
 fun (sc1 |-- sc2) toks =
     let
-	val (_,toks2) = sc1 toks
+        val (_,toks2) = sc1 toks
     in
-	sc2 toks2
+        sc2 toks2
     end
 
 fun (sc1 ^^ sc2) toks =
     let
-	val (x,toks2) = sc1 toks
-	val (y,toks3) = sc2 toks2
+        val (x,toks2) = sc1 toks
+        val (y,toks3) = sc2 toks2
     in
-	(x^y,toks3)
+        (x^y,toks3)
     end
 
 fun (sc1 || sc2) toks =
@@ -119,22 +118,22 @@
 
 fun any p toks =
     case pull toks of
-	NONE =>  ([],toks)
+        NONE =>  ([],toks)
       | SOME(x,toks2) => if p x
-			 then
-			     let
-				 val (xs,toks3) = any p toks2
-			     in
-				 (x::xs,toks3)
-			     end
-			 else ([],toks)
+                         then
+                             let
+                                 val (xs,toks3) = any p toks2
+                             in
+                                 (x::xs,toks3)
+                             end
+                         else ([],toks)
 
 fun any1 p toks =
     let
-	val (x,toks2) = one p toks
-	val (xs,toks3) = any p toks2
+        val (x,toks2) = one p toks
+        val (xs,toks3) = any p toks2
     in
-	(x::xs,toks3)
+        (x::xs,toks3)
     end
 
 fun optional sc def =  sc || succeed def
@@ -143,16 +142,16 @@
 (*
 fun repeat sc =
     let
-	fun R toks =
-	    let
-		val (x,toks2) = sc toks
-		val (xs,toks3) = R toks2
-	    in
-		(x::xs,toks3)
-	    end
-	    handle SyntaxError => ([],toks)
+        fun R toks =
+            let
+                val (x,toks2) = sc toks
+                val (xs,toks3) = R toks2
+            in
+                (x::xs,toks3)
+            end
+            handle SyntaxError => ([],toks)
     in
-	R
+        R
     end
 *)
 
@@ -164,46 +163,46 @@
  *)
 fun repeat sc =
     let
-	fun R xs toks =
-	    case SOME (sc toks) handle SyntaxError => NONE of
-		SOME (x,toks2) => R (x::xs) toks2
-	      | NONE => (List.rev xs,toks)
+        fun R xs toks =
+            case SOME (sc toks) handle SyntaxError => NONE of
+                SOME (x,toks2) => R (x::xs) toks2
+              | NONE => (List.rev xs,toks)
     in
-	R []
+        R []
     end
 
 fun repeat1 sc toks =
     let
-	val (x,toks2) = sc toks
-	val (xs,toks3) = repeat sc toks2
+        val (x,toks2) = sc toks
+        val (xs,toks3) = repeat sc toks2
     in
-	(x::xs,toks3)
+        (x::xs,toks3)
     end
 
 fun repeat_fixed n sc =
     let
-	fun R n xs toks =
-	    if (n <= 0) then (List.rev xs, toks)
-	    else case (sc toks) of (x, toks2) => R (n-1) (x::xs) toks2
+        fun R n xs toks =
+            if (n <= 0) then (List.rev xs, toks)
+            else case (sc toks) of (x, toks2) => R (n-1) (x::xs) toks2
     in
-	R n []
+        R n []
     end
 
 fun ahead (sc:'a->'b*'a) toks = (#1 (sc toks),toks)
 
 fun unless test sc toks =
     let
-	val test_failed = (test toks;false) handle SyntaxError => true
+        val test_failed = (test toks;false) handle SyntaxError => true
     in
-	if test_failed
-	then sc toks
-	else raise SyntaxError
+        if test_failed
+        then sc toks
+        else raise SyntaxError
     end
 
 fun $$ arg = one (fn x => x = arg)
 
 fun !! f sc toks = (sc toks
-		    handle SyntaxError => raise Fail (f toks))
+                    handle SyntaxError => raise Fail (f toks))
 
 val scan_id = one Symbol.is_letter ^^ (any Symbol.is_letdig >> implode);
 
--- a/src/HOL/Import/seq.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/seq.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -28,16 +28,16 @@
 fun take_at_most s n = 
     if n <= 0 then [] else
     case pull s of 
-	NONE => []
+        NONE => []
       | SOME (x,s') => x::(take_at_most s' (n-1))
 
 fun length' s n = 
     case pull s of
-	NONE => n
+        NONE => n
       | SOME (_, s') => length' s' (n+1)
 
 fun length s = length' s 0
-		
+                
 end  
 
 signature VECTOR_SCANNER_SEQ = 
--- a/src/HOL/Import/shuffler.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/shuffler.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -101,7 +101,7 @@
         val th4 = implies_elim_list (assume cPPQ) [th3,th3]
                                     |> implies_intr_list [cPPQ,cP]
     in
-        equal_intr th4 th1 |> standard
+        equal_intr th4 th1 |> Drule.standard
     end
 
 val imp_comm =
@@ -121,7 +121,7 @@
         val th2 = implies_elim_list (assume cQPR) [assume cQ,assume cP]
                                     |> implies_intr_list [cQPR,cP,cQ]
     in
-        equal_intr th1 th2 |> standard
+        equal_intr th1 th2 |> Drule.standard
     end
 
 val def_norm =
@@ -148,7 +148,7 @@
                               |> forall_intr cv
                               |> implies_intr cPQ
     in
-        equal_intr th1 th2 |> standard
+        equal_intr th1 th2 |> Drule.standard
     end
 
 val all_comm =
@@ -174,7 +174,7 @@
                          |> forall_intr_list [cy,cx]
                          |> implies_intr cl
     in
-        equal_intr th1 th2 |> standard
+        equal_intr th1 th2 |> Drule.standard
     end
 
 val equiv_comm =
@@ -188,7 +188,7 @@
         val th1  = assume ctu |> symmetric |> implies_intr ctu
         val th2  = assume cut |> symmetric |> implies_intr cut
     in
-        equal_intr th1 th2 |> standard
+        equal_intr th1 th2 |> Drule.standard
     end
 
 (* This simplification procedure rewrites !!x y. P x y
--- a/src/HOL/Import/xml.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/xml.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Import/xml.ML
-    ID:         $Id$
 
 Yet another version of XML support.
 *)
@@ -157,11 +156,11 @@
 
 fun tree_of_string s =
     let
-	val seq = Seq.fromString s
-	val scanner = !! (err "Malformed element") (scan_whspc |-- parse_elem --| scan_whspc)
-	val (x, toks) = scanner seq
+        val seq = Seq.fromString s
+        val scanner = !! (err "Malformed element") (scan_whspc |-- parse_elem --| scan_whspc)
+        val (x, toks) = scanner seq
     in
-	if Seq.null toks then x else error ("Unprocessed input: '"^(beginning 100 toks)^"'")
+        if Seq.null toks then x else error ("Unprocessed input: '"^(beginning 100 toks)^"'")
     end
 
 (* More efficient saving and loading of xml trees employing a proprietary external format *)
@@ -175,9 +174,9 @@
 fun write_tree (Text s) buf = buf |> Buffer.add "T" |> write_lstring s
   | write_tree (Elem (name, attrs, children)) buf = 
     buf |> Buffer.add "E" 
-	|> write_lstring name 
-	|> write_list (fn (a,b) => fn buf => buf |> write_lstring a |> write_lstring b) attrs 
-	|> write_list write_tree children
+        |> write_lstring name 
+        |> write_list (fn (a,b) => fn buf => buf |> write_lstring a |> write_lstring b) attrs 
+        |> write_list write_tree children
 
 fun parse_tree toks = (one (K true) :-- (fn "T" => parse_lstring >> Text | "E" => parse_elem | _ => raise SyntaxError) >> snd) toks
 and parse_elem toks = (parse_lstring -- parse_list (parse_lstring -- parse_lstring) -- parse_list parse_tree >> (fn ((n, a), c) => Elem (n,a,c))) toks
@@ -186,15 +185,15 @@
 
 fun tree_of_encoded_string s = 
     let
-	fun print (a,b) = writeln (a^" "^(string_of_int b))
-	val _ = print ("length of encoded string: ", size s)
-	val _ = writeln "Seq.fromString..."
-	val seq = timeit (fn () => Seq.fromString s)
-	val  _ = print ("length of sequence", timeit (fn () => Seq.length seq))
-	val scanner = !! (err "Malformed encoded xml") parse_tree
-	val (x, toks) = scanner seq
+        fun print (a,b) = writeln (a^" "^(string_of_int b))
+        val _ = print ("length of encoded string: ", size s)
+        val _ = writeln "Seq.fromString..."
+        val seq = timeit (fn () => Seq.fromString s)
+        val  _ = print ("length of sequence", timeit (fn () => Seq.length seq))
+        val scanner = !! (err "Malformed encoded xml") parse_tree
+        val (x, toks) = scanner seq
     in
-	if Seq.null toks then x else error ("Unprocessed input: '"^(beginning 100 toks)^"'")
-    end	       
+        if Seq.null toks then x else error ("Unprocessed input: '"^(beginning 100 toks)^"'")
+    end        
 
 end;
--- a/src/HOL/Import/xmlconv.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Import/xmlconv.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -18,7 +18,7 @@
 
   val xml_of_bool: bool output
   val bool_of_xml: bool input
-		  
+                  
   val xml_of_int: int output
   val int_of_xml: int input
 
@@ -86,7 +86,7 @@
 
 fun intofstr s i = 
     case Int.fromString i of 
-	NONE => parse_err (s^", invalid index: "^i)
+        NONE => parse_err (s^", invalid index: "^i)
       | SOME i => i
 
 fun typ_of_xml (XML.Elem ("TVar", [("name", s)], cs)) = TVar ((s,0), map class_of_xml cs)
@@ -166,7 +166,7 @@
     XML.Elem ("quadruple", [], [output1 x1, output2 x2, output3 x3, output4 x4])
 fun xml_of_quintuple output1 output2 output3 output4 output5 (x1, x2, x3, x4, x5) = 
     XML.Elem ("quintuple", [], [output1 x1, output2 x2, output3 x3, output4 x4, output5 x5])
-										  
+                                                                                  
 fun bool_of_xml (XML.Elem ("bool", [("value", "true")], [])) = true
   | bool_of_xml (XML.Elem ("bool", [("value", "false")], [])) = false
   | bool_of_xml _ = parse_err "bool"
@@ -225,10 +225,10 @@
   | xml_of_mixfix (Infixl i) = wrap "infixl" xml_of_int i
   | xml_of_mixfix (Infixr i) = wrap "infixr" xml_of_int i
   | xml_of_mixfix (Binder args) = wrap "binder" (xml_of_triple xml_of_string xml_of_int xml_of_int) args
-				  
+                                  
 fun mixfix_of_xml e = 
     (case name_of_wrap e of 
-	 "nosyn" => unwrap_empty NoSyn 
+         "nosyn" => unwrap_empty NoSyn 
        | "structure" => unwrap_empty Structure 
        | "mixfix" => unwrap Mixfix (triple_of_xml string_of_xml (list_of_xml int_of_xml) int_of_xml)
        | "delimfix" => unwrap Delimfix string_of_xml
@@ -247,15 +247,15 @@
  
 fun from_file f input fname = 
     let
-	val _ = writeln "read_from_file enter"
-	val s = File.read (Path.explode fname)
-	val _ = writeln "done: read file"
-	val tree = timeit (fn () => f s)
-	val _ = writeln "done: tree"
-	val x = timeit (fn () => input tree)
-	val _ = writeln "done: input"
+        val _ = writeln "read_from_file enter"
+        val s = File.read (Path.explode fname)
+        val _ = writeln "done: read file"
+        val tree = timeit (fn () => f s)
+        val _ = writeln "done: tree"
+        val x = timeit (fn () => input tree)
+        val _ = writeln "done: input"
     in
-	x
+        x
     end
 
 fun write_to_file x = to_file XML.string_of_tree x
--- a/src/HOL/Induct/LFilter.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Induct/LFilter.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/LList.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1997  University of Cambridge
 *)
@@ -10,8 +9,8 @@
 theory LFilter imports LList begin
 
 inductive_set
-  findRel	:: "('a => bool) => ('a llist * 'a llist)set"
-  for p :: "'a => bool"
+  findRel :: "('a => bool) => ('a llist * 'a llist)set"
+    for p :: "'a => bool"
   where
     found:  "p x ==> (LCons x l, LCons x l) \<in> findRel p"
   | seek:   "[| ~p x;  (l,l') \<in> findRel p |] ==> (LCons x l, l') \<in> findRel p"
--- a/src/HOL/Induct/QuoDataType.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Induct/QuoDataType.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      HOL/Induct/QuoDataType
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   2004  University of Cambridge
-
 *)
 
 header{*Defining an Initial Algebra by Quotienting a Free Algebra*}
@@ -14,9 +12,9 @@
 text{*Messages with encryption and decryption as free constructors.*}
 datatype
      freemsg = NONCE  nat
-	     | MPAIR  freemsg freemsg
-	     | CRYPT  nat freemsg  
-	     | DECRYPT  nat freemsg
+             | MPAIR  freemsg freemsg
+             | CRYPT  nat freemsg  
+             | DECRYPT  nat freemsg
 
 text{*The equivalence relation, which makes encryption and decryption inverses
 provided the keys are the same.
--- a/src/HOL/Induct/QuoNestedDataType.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Induct/QuoNestedDataType.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      HOL/Induct/QuoNestedDataType
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   2004  University of Cambridge
-
 *)
 
 header{*Quotienting a Free Algebra Involving Nested Recursion*}
@@ -14,8 +12,8 @@
 text{*Messages with encryption and decryption as free constructors.*}
 datatype
      freeExp = VAR  nat
-	     | PLUS  freeExp freeExp
-	     | FNCALL  nat "freeExp list"
+             | PLUS  freeExp freeExp
+             | FNCALL  nat "freeExp list"
 
 text{*The equivalence relation, which makes PLUS associative.*}
 
--- a/src/HOL/Induct/SList.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Induct/SList.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -276,7 +276,7 @@
         !!x xs. P(xs) ==> P(x # xs) |]  ==> P(l)"
 apply (unfold Nil_def Cons_def) 
 apply (rule Rep_List_inverse [THEN subst])
-			 (*types force good instantiation*)
+(*types force good instantiation*)
 apply (rule Rep_List [unfolded List_def, THEN list.induct], simp)
 apply (erule Abs_List_inverse [unfolded List_def, THEN subst], blast) 
 done
--- a/src/HOL/Induct/Sexp.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Induct/Sexp.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Induct/Sexp.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 
@@ -80,10 +79,10 @@
 (* (a,b) \<in> pred_sexp^+ ==> a \<in> sexp *)
 lemmas trancl_pred_sexpD1 = 
     pred_sexp_subset_Sigma
-	 [THEN trancl_subset_Sigma, THEN subsetD, THEN SigmaD1]
+         [THEN trancl_subset_Sigma, THEN subsetD, THEN SigmaD1]
 and trancl_pred_sexpD2 = 
     pred_sexp_subset_Sigma
-	 [THEN trancl_subset_Sigma, THEN subsetD, THEN SigmaD2]
+         [THEN trancl_subset_Sigma, THEN subsetD, THEN SigmaD2]
 
 lemma pred_sexpI1: 
     "[| M \<in> sexp;  N \<in> sexp |] ==> (M, Scons M N) \<in> pred_sexp"
--- a/src/HOL/IntDiv.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/IntDiv.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,7 +1,6 @@
 (*  Title:      HOL/IntDiv.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1999  University of Cambridge
-
 *)
 
 header{* The Division Operators div and mod *}
@@ -73,24 +72,24 @@
     fun posDivAlg (a,b) =
       if a<b then (0,a)
       else let val (q,r) = posDivAlg(a, 2*b)
-	       in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
-	   end
+               in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
+           end
 
     fun negDivAlg (a,b) =
       if 0\<le>a+b then (~1,a+b)
       else let val (q,r) = negDivAlg(a, 2*b)
-	       in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
-	   end;
+               in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
+           end;
 
     fun negateSnd (q,r:int) = (q,~r);
 
     fun divmod (a,b) = if 0\<le>a then 
-			  if b>0 then posDivAlg (a,b) 
-			   else if a=0 then (0,0)
-				else negateSnd (negDivAlg (~a,~b))
-		       else 
-			  if 0<b then negDivAlg (a,b)
-			  else        negateSnd (posDivAlg (~a,~b));
+                          if b>0 then posDivAlg (a,b) 
+                           else if a=0 then (0,0)
+                                else negateSnd (negDivAlg (~a,~b))
+                       else 
+                          if 0<b then negDivAlg (a,b)
+                          else        negateSnd (posDivAlg (~a,~b));
 \end{verbatim}
 *}
 
@@ -143,7 +142,7 @@
 lemma adjust_eq [simp]:
      "adjust b (q,r) = 
       (let diff = r-b in  
-	if 0 \<le> diff then (2*q + 1, diff)   
+        if 0 \<le> diff then (2*q + 1, diff)   
                      else (2*q, r))"
 by (simp add: Let_def adjust_def)
 
--- a/src/HOL/Integration.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Integration.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Author      : Jacques D. Fleuriot
-    Copyright   : 2000  University of Edinburgh
+(*  Author:     Jacques D. Fleuriot, University of Edinburgh
     Conversion to Isar and new proofs by Lawrence C Paulson, 2004
 *)
 
@@ -151,7 +150,7 @@
       case (Cons d1 D1')
       with induct(2)[OF `D2 \<noteq> []`, of D1'] induct(8)
       have "fine \<delta> (b, fst (hd D2)) D1'" and "fine \<delta> (fst (hd D2), c) D2" and
-	"d1 = (a', x, b)" by auto
+        "d1 = (a', x, b)" by auto
       with fine_Cons[OF this(1) induct(3,4,5), OF induct(6)] Cons
       show ?thesis by auto
     qed
@@ -381,13 +380,13 @@
     proof (rule fine_append)
       show "fine \<delta>1 (a, d) ?D1"
       proof (rule fine1[THEN fine_\<delta>_expand])
-	fix x assume "a \<le> x" "x \<le> d"
-	hence "x \<le> b" using `d < b` `x \<le> d` by auto
-	thus "\<delta> x \<le> \<delta>1 x" unfolding \<delta>_def by auto
+        fix x assume "a \<le> x" "x \<le> d"
+        hence "x \<le> b" using `d < b` `x \<le> d` by auto
+        thus "\<delta> x \<le> \<delta>1 x" unfolding \<delta>_def by auto
       qed
 
       have "b - d < \<delta>1 t"
-	using mem_fine3[OF fine in_D] \<delta>_def `b \<le> e` `t = b` by auto
+        using mem_fine3[OF fine in_D] \<delta>_def `b \<le> e` `t = b` by auto
       from `d < b` `d \<le> t` `t = b` this
       show "fine \<delta>1 (d, b) [(d, t, b)]" using fine_single by auto
     qed
@@ -410,8 +409,8 @@
       with *(2) have "fine \<delta> (e,c) (drop (Suc N) D)" by auto
       thus ?thesis
       proof (rule fine_\<delta>_expand)
-	fix x assume "e \<le> x" and "x \<le> c"
-	thus "\<delta> x \<le> \<delta>2 x" using `b \<le> e` unfolding \<delta>_def by auto
+        fix x assume "e \<le> x" and "x \<le> c"
+        thus "\<delta> x \<le> \<delta>2 x" using `b \<le> e` unfolding \<delta>_def by auto
       qed
     qed
 
@@ -421,11 +420,11 @@
     next
       case False
       have "e - b < \<delta>2 b"
-	using mem_fine3[OF fine in_D] \<delta>_def `d < b` `t = b` by auto
+        using mem_fine3[OF fine in_D] \<delta>_def `d < b` `t = b` by auto
       with False `t = b` `b \<le> e`
       show ?thesis using D2_def
-	by (auto intro!: fine_append[OF _ fine2] fine_single
-	       simp del: append_Cons)
+        by (auto intro!: fine_append[OF _ fine2] fine_single
+               simp del: append_Cons)
     qed
     note rsum2 = I2[OF this]
 
--- a/src/HOL/Isar_examples/Mutilated_Checkerboard.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Isar_examples/Mutilated_Checkerboard.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -43,12 +43,12 @@
     show "(a Un t) Un u : ?T"
     proof -
       have "a Un (t Un u) : ?T"
-	using `a : A`
+        using `a : A`
       proof (rule tiling.Un)
         from `(a Un t) Int u = {}` have "t Int u = {}" by blast
         then show "t Un u: ?T" by (rule Un)
         from `a <= - t` and `(a Un t) Int u = {}`
-	show "a <= - (t Un u)" by blast
+        show "a <= - (t Un u)" by blast
       qed
       also have "a Un (t Un u) = (a Un t) Un u"
         by (simp only: Un_assoc)
@@ -229,7 +229,7 @@
     moreover have "card ... = Suc (card (?e t b))"
     proof (rule card_insert_disjoint)
       from `t \<in> tiling domino` have "finite t"
-	by (rule tiling_domino_finite)
+        by (rule tiling_domino_finite)
       then show "finite (?e t b)"
         by (rule evnodd_finite)
       from e have "(i, j) : ?e a b" by simp
--- a/src/HOL/Isar_examples/Puzzle.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Isar_examples/Puzzle.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -27,19 +27,19 @@
       then have hyp: "\<And>m. f m < f n \<Longrightarrow> m \<le> f m" by (simp only:)
       show "n \<le> f n"
       proof (cases n)
-	case (Suc m)
-	from f_ax have "f (f m) < f n" by (simp only: Suc)
-	with hyp have "f m \<le> f (f m)" .
-	also from f_ax have "\<dots> < f n" by (simp only: Suc)
-	finally have "f m < f n" .
-	with hyp have "m \<le> f m" .
-	also note `\<dots> < f n`
-	finally have "m < f n" .
-	then have "n \<le> f n" by (simp only: Suc)
-	then show ?thesis .
+        case (Suc m)
+        from f_ax have "f (f m) < f n" by (simp only: Suc)
+        with hyp have "f m \<le> f (f m)" .
+        also from f_ax have "\<dots> < f n" by (simp only: Suc)
+        finally have "f m < f n" .
+        with hyp have "m \<le> f m" .
+        also note `\<dots> < f n`
+        finally have "m < f n" .
+        then have "n \<le> f n" by (simp only: Suc)
+        then show ?thesis .
       next
-	case 0
-	then show ?thesis by simp
+        case 0
+        then show ?thesis by simp
       qed
     qed
   } note ge = this
--- a/src/HOL/Lambda/NormalForm.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Lambda/NormalForm.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lambda/NormalForm.thy
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen, 2003
 *)
 
@@ -200,9 +199,9 @@
     proof
       assume "Var t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> t'"
       then obtain rs where "ts => rs"
-	by (iprover dest: head_Var_reduction)
+        by (iprover dest: head_Var_reduction)
       with App show False
-	by (induct rs arbitrary: ts) auto
+        by (induct rs arbitrary: ts) auto
     qed
   next
     case (Abs t)
@@ -229,7 +228,7 @@
     proof (cases ts)
       case Nil
       from 2 have "\<forall>u'. \<not> u \<rightarrow>\<^sub>\<beta> u'"
-	by (auto intro: apps_preserves_beta)
+        by (auto intro: apps_preserves_beta)
       then have "NF u" by (rule 2)
       then have "NF (Abs u)" by (rule NF.Abs)
       with Nil show ?thesis by simp
@@ -237,9 +236,9 @@
       case (Cons r rs)
       have "Abs u \<degree> r \<rightarrow>\<^sub>\<beta> u[r/0]" ..
       then have "Abs u \<degree> r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> u[r/0] \<degree>\<degree> rs"
-	by (rule apps_preserves_beta)
+        by (rule apps_preserves_beta)
       with Cons have "Abs u \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> u[r/0] \<degree>\<degree> rs"
-	by simp
+        by simp
       with 2 show ?thesis by iprover
     qed
   qed
--- a/src/HOL/Lambda/StrongNorm.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Lambda/StrongNorm.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lambda/StrongNorm.thy
-    ID:         $Id$
     Author:     Stefan Berghofer
     Copyright   2000 TU Muenchen
 *)
@@ -230,7 +229,7 @@
         with T have "e\<langle>i:T\<rangle> \<turnstile> r[a/0] \<degree>\<degree> as : T'"
           by (rule subject_reduction)
         hence "IT ((r[a/0] \<degree>\<degree> as)[u/i])"
-	  using uIT uT by (rule SI1)
+          using uIT uT by (rule SI1)
         thus "IT (r[lift u 0/Suc i][a[u/i]/0] \<degree>\<degree> map (\<lambda>t. t[u/i]) as)"
           by (simp del: subst_map add: subst_subst subst_map [symmetric])
         from T obtain U where "e\<langle>i:T\<rangle> \<turnstile> Abs r \<degree> a : U"
--- a/src/HOL/Lambda/WeakNorm.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Lambda/WeakNorm.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -79,112 +79,112 @@
       case (App ts x e_ T'_ u_ i_)
       assume "e\<langle>i:T\<rangle> \<turnstile> Var x \<degree>\<degree> ts : T'"
       then obtain Us
-	where varT: "e\<langle>i:T\<rangle> \<turnstile> Var x : Us \<Rrightarrow> T'"
-	and argsT: "e\<langle>i:T\<rangle> \<tturnstile> ts : Us"
-	by (rule var_app_typesE)
+        where varT: "e\<langle>i:T\<rangle> \<turnstile> Var x : Us \<Rrightarrow> T'"
+        and argsT: "e\<langle>i:T\<rangle> \<tturnstile> ts : Us"
+        by (rule var_app_typesE)
       from nat_eq_dec show "\<exists>t'. (Var x \<degree>\<degree> ts)[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'"
       proof
-	assume eq: "x = i"
-	show ?thesis
-	proof (cases ts)
-	  case Nil
-	  with eq have "(Var x \<degree>\<degree> [])[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* u" by simp
-	  with Nil and uNF show ?thesis by simp iprover
-	next
-	  case (Cons a as)
+        assume eq: "x = i"
+        show ?thesis
+        proof (cases ts)
+          case Nil
+          with eq have "(Var x \<degree>\<degree> [])[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* u" by simp
+          with Nil and uNF show ?thesis by simp iprover
+        next
+          case (Cons a as)
           with argsT obtain T'' Ts where Us: "Us = T'' # Ts"
-	    by (cases Us) (rule FalseE, simp+, erule that)
-	  from varT and Us have varT: "e\<langle>i:T\<rangle> \<turnstile> Var x : T'' \<Rightarrow> Ts \<Rrightarrow> T'"
-	    by simp
+            by (cases Us) (rule FalseE, simp+, erule that)
+          from varT and Us have varT: "e\<langle>i:T\<rangle> \<turnstile> Var x : T'' \<Rightarrow> Ts \<Rrightarrow> T'"
+            by simp
           from varT eq have T: "T = T'' \<Rightarrow> Ts \<Rrightarrow> T'" by cases auto
           with uT have uT': "e \<turnstile> u : T'' \<Rightarrow> Ts \<Rrightarrow> T'" by simp
-	  from argsT Us Cons have argsT': "e\<langle>i:T\<rangle> \<tturnstile> as : Ts" by simp
-	  from argsT Us Cons have argT: "e\<langle>i:T\<rangle> \<turnstile> a : T''" by simp
-	  from argT uT refl have aT: "e \<turnstile> a[u/i] : T''" by (rule subst_lemma)
-	  from App and Cons have "listall ?R as" by simp (iprover dest: listall_conj2)
-	  with lift_preserves_beta' lift_NF uNF uT argsT'
-	  have "\<exists>as'. \<forall>j. Var j \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as \<rightarrow>\<^sub>\<beta>\<^sup>*
+          from argsT Us Cons have argsT': "e\<langle>i:T\<rangle> \<tturnstile> as : Ts" by simp
+          from argsT Us Cons have argT: "e\<langle>i:T\<rangle> \<turnstile> a : T''" by simp
+          from argT uT refl have aT: "e \<turnstile> a[u/i] : T''" by (rule subst_lemma)
+          from App and Cons have "listall ?R as" by simp (iprover dest: listall_conj2)
+          with lift_preserves_beta' lift_NF uNF uT argsT'
+          have "\<exists>as'. \<forall>j. Var j \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as \<rightarrow>\<^sub>\<beta>\<^sup>*
             Var j \<degree>\<degree> map (\<lambda>t. lift t 0) as' \<and>
-	    NF (Var j \<degree>\<degree> map (\<lambda>t. lift t 0) as')" by (rule norm_list)
-	  then obtain as' where
-	    asred: "Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as \<rightarrow>\<^sub>\<beta>\<^sup>*
-	      Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as'"
-	    and asNF: "NF (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')" by iprover
-	  from App and Cons have "?R a" by simp
-	  with argT and uNF and uT have "\<exists>a'. a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* a' \<and> NF a'"
-	    by iprover
-	  then obtain a' where ared: "a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* a'" and aNF: "NF a'" by iprover
-	  from uNF have "NF (lift u 0)" by (rule lift_NF)
-	  hence "\<exists>u'. lift u 0 \<degree> Var 0 \<rightarrow>\<^sub>\<beta>\<^sup>* u' \<and> NF u'" by (rule app_Var_NF)
-	  then obtain u' where ured: "lift u 0 \<degree> Var 0 \<rightarrow>\<^sub>\<beta>\<^sup>* u'" and u'NF: "NF u'"
-	    by iprover
-	  from T and u'NF have "\<exists>ua. u'[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua \<and> NF ua"
-	  proof (rule MI1)
-	    have "e\<langle>0:T''\<rangle> \<turnstile> lift u 0 \<degree> Var 0 : Ts \<Rrightarrow> T'"
-	    proof (rule typing.App)
-	      from uT' show "e\<langle>0:T''\<rangle> \<turnstile> lift u 0 : T'' \<Rightarrow> Ts \<Rrightarrow> T'" by (rule lift_type)
-	      show "e\<langle>0:T''\<rangle> \<turnstile> Var 0 : T''" by (rule typing.Var) simp
-	    qed
-	    with ured show "e\<langle>0:T''\<rangle> \<turnstile> u' : Ts \<Rrightarrow> T'" by (rule subject_reduction')
-	    from ared aT show "e \<turnstile> a' : T''" by (rule subject_reduction')
-	    show "NF a'" by fact
-	  qed
-	  then obtain ua where uared: "u'[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" and uaNF: "NF ua"
-	    by iprover
-	  from ared have "(lift u 0 \<degree> Var 0)[a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>* (lift u 0 \<degree> Var 0)[a'/0]"
-	    by (rule subst_preserves_beta2')
-	  also from ured have "(lift u 0 \<degree> Var 0)[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* u'[a'/0]"
-	    by (rule subst_preserves_beta')
-	  also note uared
-	  finally have "(lift u 0 \<degree> Var 0)[a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" .
-	  hence uared': "u \<degree> a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" by simp
-	  from T asNF _ uaNF have "\<exists>r. (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r \<and> NF r"
-	  proof (rule MI2)
-	    have "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<turnstile> Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as : T'"
-	    proof (rule list_app_typeI)
-	      show "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<turnstile> Var 0 : Ts \<Rrightarrow> T'" by (rule typing.Var) simp
-	      from uT argsT' have "e \<tturnstile> map (\<lambda>t. t[u/i]) as : Ts"
-		by (rule substs_lemma)
-	      hence "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<tturnstile> map (\<lambda>t. lift t 0) (map (\<lambda>t. t[u/i]) as) : Ts"
-		by (rule lift_types)
-	      thus "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<tturnstile> map (\<lambda>t. lift (t[u/i]) 0) as : Ts"
-		by (simp_all add: map_compose [symmetric] o_def)
-	    qed
-	    with asred show "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<turnstile> Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as' : T'"
-	      by (rule subject_reduction')
-	    from argT uT refl have "e \<turnstile> a[u/i] : T''" by (rule subst_lemma)
-	    with uT' have "e \<turnstile> u \<degree> a[u/i] : Ts \<Rrightarrow> T'" by (rule typing.App)
-	    with uared' show "e \<turnstile> ua : Ts \<Rrightarrow> T'" by (rule subject_reduction')
-	  qed
-	  then obtain r where rred: "(Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r"
-	    and rnf: "NF r" by iprover
-	  from asred have
-	    "(Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as)[u \<degree> a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>*
-	    (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[u \<degree> a[u/i]/0]"
-	    by (rule subst_preserves_beta')
-	  also from uared' have "(Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[u \<degree> a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>*
-	    (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0]" by (rule subst_preserves_beta2')
-	  also note rred
-	  finally have "(Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as)[u \<degree> a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r" .
-	  with rnf Cons eq show ?thesis
-	    by (simp add: map_compose [symmetric] o_def) iprover
-	qed
+            NF (Var j \<degree>\<degree> map (\<lambda>t. lift t 0) as')" by (rule norm_list)
+          then obtain as' where
+            asred: "Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as \<rightarrow>\<^sub>\<beta>\<^sup>*
+              Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as'"
+            and asNF: "NF (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')" by iprover
+          from App and Cons have "?R a" by simp
+          with argT and uNF and uT have "\<exists>a'. a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* a' \<and> NF a'"
+            by iprover
+          then obtain a' where ared: "a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* a'" and aNF: "NF a'" by iprover
+          from uNF have "NF (lift u 0)" by (rule lift_NF)
+          hence "\<exists>u'. lift u 0 \<degree> Var 0 \<rightarrow>\<^sub>\<beta>\<^sup>* u' \<and> NF u'" by (rule app_Var_NF)
+          then obtain u' where ured: "lift u 0 \<degree> Var 0 \<rightarrow>\<^sub>\<beta>\<^sup>* u'" and u'NF: "NF u'"
+            by iprover
+          from T and u'NF have "\<exists>ua. u'[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua \<and> NF ua"
+          proof (rule MI1)
+            have "e\<langle>0:T''\<rangle> \<turnstile> lift u 0 \<degree> Var 0 : Ts \<Rrightarrow> T'"
+            proof (rule typing.App)
+              from uT' show "e\<langle>0:T''\<rangle> \<turnstile> lift u 0 : T'' \<Rightarrow> Ts \<Rrightarrow> T'" by (rule lift_type)
+              show "e\<langle>0:T''\<rangle> \<turnstile> Var 0 : T''" by (rule typing.Var) simp
+            qed
+            with ured show "e\<langle>0:T''\<rangle> \<turnstile> u' : Ts \<Rrightarrow> T'" by (rule subject_reduction')
+            from ared aT show "e \<turnstile> a' : T''" by (rule subject_reduction')
+            show "NF a'" by fact
+          qed
+          then obtain ua where uared: "u'[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" and uaNF: "NF ua"
+            by iprover
+          from ared have "(lift u 0 \<degree> Var 0)[a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>* (lift u 0 \<degree> Var 0)[a'/0]"
+            by (rule subst_preserves_beta2')
+          also from ured have "(lift u 0 \<degree> Var 0)[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* u'[a'/0]"
+            by (rule subst_preserves_beta')
+          also note uared
+          finally have "(lift u 0 \<degree> Var 0)[a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" .
+          hence uared': "u \<degree> a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" by simp
+          from T asNF _ uaNF have "\<exists>r. (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r \<and> NF r"
+          proof (rule MI2)
+            have "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<turnstile> Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as : T'"
+            proof (rule list_app_typeI)
+              show "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<turnstile> Var 0 : Ts \<Rrightarrow> T'" by (rule typing.Var) simp
+              from uT argsT' have "e \<tturnstile> map (\<lambda>t. t[u/i]) as : Ts"
+                by (rule substs_lemma)
+              hence "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<tturnstile> map (\<lambda>t. lift t 0) (map (\<lambda>t. t[u/i]) as) : Ts"
+                by (rule lift_types)
+              thus "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<tturnstile> map (\<lambda>t. lift (t[u/i]) 0) as : Ts"
+                by (simp_all add: map_compose [symmetric] o_def)
+            qed
+            with asred show "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<turnstile> Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as' : T'"
+              by (rule subject_reduction')
+            from argT uT refl have "e \<turnstile> a[u/i] : T''" by (rule subst_lemma)
+            with uT' have "e \<turnstile> u \<degree> a[u/i] : Ts \<Rrightarrow> T'" by (rule typing.App)
+            with uared' show "e \<turnstile> ua : Ts \<Rrightarrow> T'" by (rule subject_reduction')
+          qed
+          then obtain r where rred: "(Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r"
+            and rnf: "NF r" by iprover
+          from asred have
+            "(Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as)[u \<degree> a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>*
+            (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[u \<degree> a[u/i]/0]"
+            by (rule subst_preserves_beta')
+          also from uared' have "(Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[u \<degree> a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>*
+            (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0]" by (rule subst_preserves_beta2')
+          also note rred
+          finally have "(Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as)[u \<degree> a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r" .
+          with rnf Cons eq show ?thesis
+            by (simp add: map_compose [symmetric] o_def) iprover
+        qed
       next
-	assume neq: "x \<noteq> i"
-	from App have "listall ?R ts" by (iprover dest: listall_conj2)
-	with TrueI TrueI uNF uT argsT
-	have "\<exists>ts'. \<forall>j. Var j \<degree>\<degree> map (\<lambda>t. t[u/i]) ts \<rightarrow>\<^sub>\<beta>\<^sup>* Var j \<degree>\<degree> ts' \<and>
-	  NF (Var j \<degree>\<degree> ts')" (is "\<exists>ts'. ?ex ts'")
-	  by (rule norm_list [of "\<lambda>t. t", simplified])
-	then obtain ts' where NF: "?ex ts'" ..
-	from nat_le_dec show ?thesis
-	proof
-	  assume "i < x"
-	  with NF show ?thesis by simp iprover
-	next
-	  assume "\<not> (i < x)"
-	  with NF neq show ?thesis by (simp add: subst_Var) iprover
-	qed
+        assume neq: "x \<noteq> i"
+        from App have "listall ?R ts" by (iprover dest: listall_conj2)
+        with TrueI TrueI uNF uT argsT
+        have "\<exists>ts'. \<forall>j. Var j \<degree>\<degree> map (\<lambda>t. t[u/i]) ts \<rightarrow>\<^sub>\<beta>\<^sup>* Var j \<degree>\<degree> ts' \<and>
+          NF (Var j \<degree>\<degree> ts')" (is "\<exists>ts'. ?ex ts'")
+          by (rule norm_list [of "\<lambda>t. t", simplified])
+        then obtain ts' where NF: "?ex ts'" ..
+        from nat_le_dec show ?thesis
+        proof
+          assume "i < x"
+          with NF show ?thesis by simp iprover
+        next
+          assume "\<not> (i < x)"
+          with NF neq show ?thesis by (simp add: subst_Var) iprover
+        qed
       qed
     next
       case (Abs r e_ T'_ u_ i_)
@@ -194,7 +194,7 @@
       moreover have "e\<langle>0:R\<rangle> \<turnstile> lift u 0 : T" using uT by (rule lift_type)
       ultimately have "\<exists>t'. r[lift u 0/Suc i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" by (rule Abs)
       thus "\<exists>t'. Abs r[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'"
-	by simp (iprover intro: rtrancl_beta_Abs NF.Abs)
+        by simp (iprover intro: rtrancl_beta_Abs NF.Abs)
     }
   qed
 qed
@@ -239,11 +239,11 @@
     show "e\<langle>0:T \<Rightarrow> U\<rangle> \<turnstile> Var 0 \<degree> lift t' 0 : U"
     proof (rule typing.App)
       show "e\<langle>0:T \<Rightarrow> U\<rangle> \<turnstile> Var 0 : T \<Rightarrow> U"
-      	by (rule typing.Var) simp
+        by (rule typing.Var) simp
       from tred have "e \<turnstile> t' : T"
-      	by (rule subject_reduction') (rule rtyping_imp_typing, rule App.hyps)
+        by (rule subject_reduction') (rule rtyping_imp_typing, rule App.hyps)
       thus "e\<langle>0:T \<Rightarrow> U\<rangle> \<turnstile> lift t' 0 : T"
-      	by (rule lift_type)
+        by (rule lift_type)
     qed
     from sred show "e \<turnstile> s' : T \<Rightarrow> U"
       by (rule subject_reduction') (rule rtyping_imp_typing, rule App.hyps)
--- a/src/HOL/Library/Abstract_Rat.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Abstract_Rat.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -56,18 +56,18 @@
     moreover
     {assume b: "b > 0"
       from b have "?b' \<ge> 0" 
-	by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos])  
+        by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos])  
       with nz' have b': "?b' > 0" by arith 
       from b b' anz bnz nz' gp1 have ?thesis 
-	by (simp add: isnormNum_def normNum_def Let_def split_def fst_conv snd_conv)}
+        by (simp add: isnormNum_def normNum_def Let_def split_def fst_conv snd_conv)}
     moreover {assume b: "b < 0"
       {assume b': "?b' \<ge> 0" 
-	from gpos have th: "?g \<ge> 0" by arith
-	from mult_nonneg_nonneg[OF th b'] zdvd_mult_div_cancel[OF gdvd(2)]
-	have False using b by arith }
+        from gpos have th: "?g \<ge> 0" by arith
+        from mult_nonneg_nonneg[OF th b'] zdvd_mult_div_cancel[OF gdvd(2)]
+        have False using b by arith }
       hence b': "?b' < 0" by (presburger add: linorder_not_le[symmetric]) 
       from anz bnz nz' b b' gp1 have ?thesis 
-	by (simp add: isnormNum_def normNum_def Let_def split_def)}
+        by (simp add: isnormNum_def normNum_def Let_def split_def)}
     ultimately have ?thesis by blast
   }
   ultimately show ?thesis by blast
@@ -287,14 +287,14 @@
       hence "of_int b' * of_int a / (of_int b * of_int b') + of_int b * of_int a' / (of_int b * of_int b') = ?z"  by (simp add:add_divide_distrib) 
       hence th: "of_int a / of_int b + of_int a' / of_int b' = ?z" using bb' aa' by simp 
       from z aa' bb' have ?thesis 
-	by (simp add: th Nadd_def normNum_def INum_def split_def)}
+        by (simp add: th Nadd_def normNum_def INum_def split_def)}
     moreover {assume z: "a * b' + b * a' \<noteq> 0"
       let ?g = "gcd (a * b' + b * a') (b*b')"
       have gz: "?g \<noteq> 0" using z by simp
       have ?thesis using aa' bb' z gz
-	of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a * b' + b * a'" and y="b*b'"]]	of_int_div[where ?'a = 'a,
-	OF gz gcd_dvd2_int[where x="a * b' + b * a'" and y="b*b'"]]
-	by (simp add: x y Nadd_def INum_def normNum_def Let_def add_divide_distrib)}
+        of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a * b' + b * a'" and y="b*b'"]]  of_int_div[where ?'a = 'a,
+        OF gz gcd_dvd2_int[where x="a * b' + b * a'" and y="b*b'"]]
+        by (simp add: x y Nadd_def INum_def normNum_def Let_def add_divide_distrib)}
     ultimately have ?thesis using aa' bb' 
       by (simp add: Nadd_def INum_def normNum_def x y Let_def) }
   ultimately show ?thesis by blast
--- a/src/HOL/Library/AssocList.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/AssocList.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -233,9 +233,9 @@
       fix v
       assume "v \<in> ran (map_of (a#al))"
       then obtain x where "map_of (a#al) x = Some v"
-	by (auto simp add: ran_def)
+        by (auto simp add: ran_def)
       then show "v \<in> {snd a} \<union> ran (map_of al)"
-	by (auto split: split_if_asm simp add: ran_def)
+        by (auto split: split_if_asm simp add: ran_def)
     qed
   next
     show "{snd a} \<union> ran (map_of al) \<subseteq> ran (map_of (a # al))"
@@ -244,22 +244,22 @@
       assume v_in: "v \<in> {snd a} \<union> ran (map_of al)"
       show "v \<in> ran (map_of (a#al))"
       proof (cases "v=snd a")
-	case True
-	with v_in show ?thesis
-	  by (auto simp add: ran_def)
+        case True
+        with v_in show ?thesis
+          by (auto simp add: ran_def)
       next
-	case False
-	with v_in have "v \<in> ran (map_of al)" by auto
-	then obtain x where al_x: "map_of al x = Some v"
-	  by (auto simp add: ran_def)
-	from map_of_SomeD [OF this]
-	have "x \<in> fst ` set al"
-	  by (force simp add: image_def)
-	with Cons.prems have "x\<noteq>fst a"
-	  by - (rule ccontr,simp)
-	with al_x
-	show ?thesis
-	  by (auto simp add: ran_def)
+        case False
+        with v_in have "v \<in> ran (map_of al)" by auto
+        then obtain x where al_x: "map_of al x = Some v"
+          by (auto simp add: ran_def)
+        from map_of_SomeD [OF this]
+        have "x \<in> fst ` set al"
+          by (force simp add: image_def)
+        with Cons.prems have "x\<noteq>fst a"
+          by - (rule ccontr,simp)
+        with al_x
+        show ?thesis
+          by (auto simp add: ran_def)
       qed
     qed
   qed
@@ -537,17 +537,17 @@
       case True
       from True delete_notin_dom [of k xs]
       have "map_of (delete (fst x) xs) k = None"
-	by (simp add: map_of_eq_None_iff)
+        by (simp add: map_of_eq_None_iff)
       with hyp show ?thesis
-	using True None
-	by simp
+        using True None
+        by simp
     next
       case False
       from False have "map_of (delete (fst x) xs) k = map_of xs k"
-	by simp
+        by simp
       with hyp show ?thesis
-	using False None
-	by (simp add: map_comp_def)
+        using False None
+        by (simp add: map_comp_def)
     qed
   next
     case (Some v)
@@ -631,12 +631,12 @@
       case True
       with None hyp
       show ?thesis
-	by (simp add: delete_idem)
+        by (simp add: delete_idem)
     next
       case False
       from None False hyp
       show ?thesis
-	by (simp add: delete_twist)
+        by (simp add: delete_twist)
     qed
   next
     case (Some v)
@@ -736,30 +736,30 @@
       case True
       from Cons.hyps
       show ?thesis
-	using x_D True
-	by simp
+        using x_D True
+        by simp
     next
       case False
       note not_fst_a_x = this
       show ?thesis
       proof (cases "fst a \<in> D")
-	case True 
-	with not_fst_a_x 
-	have "delete x (restrict D (a#al)) = a#(delete x (restrict D al))"
-	  by (cases a) (simp add: restrict_eq)
-	also from not_fst_a_x True hyp have "\<dots> = restrict (D - {x}) (a # al)"
-	  by (cases a) (simp add: restrict_eq)
-	finally show ?thesis
-	  using x_D by simp
+        case True 
+        with not_fst_a_x 
+        have "delete x (restrict D (a#al)) = a#(delete x (restrict D al))"
+          by (cases a) (simp add: restrict_eq)
+        also from not_fst_a_x True hyp have "\<dots> = restrict (D - {x}) (a # al)"
+          by (cases a) (simp add: restrict_eq)
+        finally show ?thesis
+          using x_D by simp
       next
-	case False
-	hence "delete x (restrict D (a#al)) = delete x (restrict D al)"
-	  by (cases a) (simp add: restrict_eq)
-	moreover from False not_fst_a_x
-	have "restrict (D - {x}) (a # al) = restrict (D - {x}) al"
-	  by (cases a) (simp add: restrict_eq)
-	ultimately
-	show ?thesis using x_D hyp by simp
+        case False
+        hence "delete x (restrict D (a#al)) = delete x (restrict D al)"
+          by (cases a) (simp add: restrict_eq)
+        moreover from False not_fst_a_x
+        have "restrict (D - {x}) (a # al) = restrict (D - {x}) al"
+          by (cases a) (simp add: restrict_eq)
+        ultimately
+        show ?thesis using x_D hyp by simp
       qed
     qed
   next
--- a/src/HOL/Library/Binomial.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Binomial.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -398,7 +398,7 @@
   {assume kn: "k > n" 
     from kn binomial_eq_0[OF kn] have ?thesis 
       by (simp add: gbinomial_pochhammer field_simps
-	pochhammer_of_nat_eq_0_iff)}
+        pochhammer_of_nat_eq_0_iff)}
   moreover
   {assume "k=0" then have ?thesis by simp}
   moreover
@@ -420,7 +420,7 @@
     from eq[symmetric]
     have ?thesis using kn
       apply (simp add: binomial_fact[OF kn, where ?'a = 'a] 
-	gbinomial_pochhammer field_simps pochhammer_Suc_setprod)
+        gbinomial_pochhammer field_simps pochhammer_Suc_setprod)
       apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc)
       unfolding setprod_Un_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \<Rightarrow> 'a"] eq[unfolded h]
       unfolding mult_assoc[symmetric] 
--- a/src/HOL/Library/Coinductive_List.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Coinductive_List.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -460,7 +460,7 @@
     next
       case (Some p)
       with h h' MN have "M = CONS (fst p) (h (snd p))"
-	and "N = CONS (fst p) (h' (snd p))"
+        and "N = CONS (fst p) (h' (snd p))"
         by (simp_all split: prod.split)
       then have ?EqCONS by (auto iff: Id_on_iff)
       then show ?thesis ..
--- a/src/HOL/Library/Commutative_Ring.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Commutative_Ring.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -288,8 +288,8 @@
             by (simp only: numerals even_nat_div_two_times_two [OF `even w`])
           have "Ipol ls P * Ipol ls P = Ipol ls P ^ Suc (Suc 0)"
             by (simp add: power_Suc)
-	  then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2"
-	    by (simp add: numerals)
+          then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2"
+            by (simp add: numerals)
           with Suc show ?thesis
             by (auto simp add: power_mult [symmetric, of _ 2 _] two_times mul_ci sqr_ci
                      simp del: power_Suc)
--- a/src/HOL/Library/Continuity.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Continuity.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -67,11 +67,11 @@
     have "chain(%i. (F ^^ i) bot)"
     proof -
       { fix i have "(F ^^ i) bot \<le> (F ^^ (Suc i)) bot"
-	proof (induct i)
-	  case 0 show ?case by simp
-	next
-	  case Suc thus ?case using monoD[OF mono Suc] by auto
-	qed }
+        proof (induct i)
+          case 0 show ?case by simp
+        next
+          case Suc thus ?case using monoD[OF mono Suc] by auto
+        qed }
       thus ?thesis by(auto simp add:chain_def)
     qed
     hence "F ?U = (SUP i. (F ^^ (i+1)) bot)" using `continuous F` by (simp add:continuous_def)
--- a/src/HOL/Library/Convex_Euclidean_Space.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Convex_Euclidean_Space.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -5,7 +5,7 @@
 header {* Convex sets, functions and related things. *}
 
 theory Convex_Euclidean_Space
-  imports Topology_Euclidean_Space
+imports Topology_Euclidean_Space
 begin
 
 
@@ -173,33 +173,33 @@
       case (Suc n) fix s::"'a set" and u::"'a \<Rightarrow> real"
       assume IA:"\<And>u s.  \<lbrakk>2 < card s; \<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V; finite s;
                s \<noteq> {}; s \<subseteq> V; setsum u s = 1; n \<equiv> card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" and
-	as:"Suc n \<equiv> card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
+        as:"Suc n \<equiv> card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
            "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = 1"
       have "\<exists>x\<in>s. u x \<noteq> 1" proof(rule_tac ccontr)
-	assume " \<not> (\<exists>x\<in>s. u x \<noteq> 1)" hence "setsum u s = real_of_nat (card s)" unfolding card_eq_setsum by auto
-	thus False using as(7) and `card s > 2` by (metis Numeral1_eq1_nat less_0_number_of less_int_code(15)
-	  less_nat_number_of not_less_iff_gr_or_eq of_nat_1 of_nat_eq_iff pos2 rel_simps(4)) qed
+        assume " \<not> (\<exists>x\<in>s. u x \<noteq> 1)" hence "setsum u s = real_of_nat (card s)" unfolding card_eq_setsum by auto
+        thus False using as(7) and `card s > 2` by (metis Numeral1_eq1_nat less_0_number_of less_int_code(15)
+          less_nat_number_of not_less_iff_gr_or_eq of_nat_1 of_nat_eq_iff pos2 rel_simps(4)) qed
       then obtain x where x:"x\<in>s" "u x \<noteq> 1" by auto
 
       have c:"card (s - {x}) = card s - 1" apply(rule card_Diff_singleton) using `x\<in>s` as(4) by auto
       have *:"s = insert x (s - {x})" "finite (s - {x})" using `x\<in>s` and as(4) by auto
       have **:"setsum u (s - {x}) = 1 - u x"
-	using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[THEN sym] as(7)] by auto
+        using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[THEN sym] as(7)] by auto
       have ***:"inverse (1 - u x) * setsum u (s - {x}) = 1" unfolding ** using `u x \<noteq> 1` by auto
       have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V" proof(cases "card (s - {x}) > 2")
-	case True hence "s - {x} \<noteq> {}" "card (s - {x}) = n" unfolding c and as(1)[symmetric] proof(rule_tac ccontr) 
-	  assume "\<not> s - {x} \<noteq> {}" hence "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp 
-	  thus False using True by auto qed auto
-	thus ?thesis apply(rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
-	unfolding setsum_right_distrib[THEN sym] using as and *** and True by auto
+        case True hence "s - {x} \<noteq> {}" "card (s - {x}) = n" unfolding c and as(1)[symmetric] proof(rule_tac ccontr) 
+          assume "\<not> s - {x} \<noteq> {}" hence "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp 
+          thus False using True by auto qed auto
+        thus ?thesis apply(rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
+        unfolding setsum_right_distrib[THEN sym] using as and *** and True by auto
       next case False hence "card (s - {x}) = Suc (Suc 0)" using as(2) and c by auto
-	then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b" unfolding card_Suc_eq by auto
-	thus ?thesis using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
-	  using *** *(2) and `s \<subseteq> V` unfolding setsum_right_distrib by(auto simp add: setsum_clauses(2)) qed
+        then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b" unfolding card_Suc_eq by auto
+        thus ?thesis using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
+          using *** *(2) and `s \<subseteq> V` unfolding setsum_right_distrib by(auto simp add: setsum_clauses(2)) qed
       thus "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" unfolding scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric]
- 	 apply(subst *) unfolding setsum_clauses(2)[OF *(2)]
-	 using as(3)[THEN bspec[where x=x], THEN bspec[where x="(inverse (1 - u x)) *\<^sub>R (\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa)"], 
-	 THEN spec[where x="u x"], THEN spec[where x="1 - u x"]] and rev_subsetD[OF `x\<in>s` `s\<subseteq>V`] and `u x \<noteq> 1` by auto
+         apply(subst *) unfolding setsum_clauses(2)[OF *(2)]
+         using as(3)[THEN bspec[where x=x], THEN bspec[where x="(inverse (1 - u x)) *\<^sub>R (\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa)"], 
+         THEN spec[where x="u x"], THEN spec[where x="1 - u x"]] and rev_subsetD[OF `x\<in>s` `s\<subseteq>V`] and `u x \<noteq> 1` by auto
     qed auto
   next assume "card s = 1" then obtain a where "s={a}" by(auto simp add: card_Suc_eq)
     thus ?thesis using as(4,5) by simp
@@ -272,20 +272,20 @@
     have *:"\<And>x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)" by auto
     have ?lhs proof(cases "a\<in>s")
       case True thus ?thesis
-	apply(rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
-	unfolding setsum_clauses(2)[OF `?as`]  apply simp
-	unfolding scaleR_left_distrib and setsum_addf 
-	unfolding vu and * and scaleR_zero_left
-	by (auto simp add: setsum_delta[OF `?as`])
+        apply(rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
+        unfolding setsum_clauses(2)[OF `?as`]  apply simp
+        unfolding scaleR_left_distrib and setsum_addf 
+        unfolding vu and * and scaleR_zero_left
+        by (auto simp add: setsum_delta[OF `?as`])
     next
       case False 
       hence **:"\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)"
                "\<And>x. x \<in> s \<Longrightarrow> u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto
       from False show ?thesis
-	apply(rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
-	unfolding setsum_clauses(2)[OF `?as`] and * using vu
-	using setsum_cong2[of s "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF **(2)]
-	using setsum_cong2[of s u "\<lambda>x. if x = a then v else u x", OF **(1)] by auto  
+        apply(rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
+        unfolding setsum_clauses(2)[OF `?as`] and * using vu
+        using setsum_cong2[of s "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF **(2)]
+        using setsum_cong2[of s u "\<lambda>x. if x = a then v else u x", OF **(1)] by auto  
     qed }
   ultimately show "?lhs = ?rhs" by blast
 qed
@@ -473,22 +473,22 @@
     assume as:"finite f" "x \<notin> f" "insert x f \<subseteq> s" "\<forall>x\<in>insert x f. 0 \<le> u x" "setsum u (insert x f) = (1::real)"
     show "(\<Sum>x\<in>insert x f. u x *\<^sub>R x) \<in> s" proof(cases "u x = 1")
       case True hence "setsum (\<lambda>x. u x *\<^sub>R x) f = 0" apply(rule_tac setsum_0') apply(rule ccontr) unfolding ball_simps apply(erule bexE) proof-
-	fix y assume y:"y \<in> f" "u y *\<^sub>R y \<noteq> 0"
-	hence uy:"u y \<noteq> 0" by auto
-	hence "setsum (\<lambda>k. if k=y then u y else 0) f \<le> setsum u f" apply(rule_tac setsum_mono) using as(4) by auto
-	hence "setsum u f \<ge> u y" using y(1) and as(1) by(auto simp add: setsum_delta) 
-	hence "setsum u f > 0" using uy apply(rule_tac less_le_trans[of _ "u y"]) using as(4) and y(1) by auto
-	thus False using as(2,5) unfolding setsum_clauses(2)[OF as(1)] and True by auto qed
+        fix y assume y:"y \<in> f" "u y *\<^sub>R y \<noteq> 0"
+        hence uy:"u y \<noteq> 0" by auto
+        hence "setsum (\<lambda>k. if k=y then u y else 0) f \<le> setsum u f" apply(rule_tac setsum_mono) using as(4) by auto
+        hence "setsum u f \<ge> u y" using y(1) and as(1) by(auto simp add: setsum_delta) 
+        hence "setsum u f > 0" using uy apply(rule_tac less_le_trans[of _ "u y"]) using as(4) and y(1) by auto
+        thus False using as(2,5) unfolding setsum_clauses(2)[OF as(1)] and True by auto qed
       thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2,3) unfolding True by auto
     next
       have *:"setsum u f = setsum u (insert x f) - u x" using as(2) unfolding setsum_clauses(2)[OF as(1)] by auto
       have **:"u x \<le> 1" apply(rule ccontr) unfolding not_le using as(5)[unfolded setsum_clauses(2)[OF as(1)]] and as(2)
-	using setsum_nonneg[of f u] and as(4) by auto
+        using setsum_nonneg[of f u] and as(4) by auto
       case False hence "inverse (1 - u x) *\<^sub>R (\<Sum>x\<in>f. u x *\<^sub>R x) \<in> s" unfolding scaleR_right.setsum and scaleR_scaleR
-	apply(rule_tac ind[THEN spec, THEN mp]) apply rule defer apply rule apply rule apply(rule mult_nonneg_nonneg)
-	unfolding setsum_right_distrib[THEN sym] and * using as and ** by auto
+        apply(rule_tac ind[THEN spec, THEN mp]) apply rule defer apply rule apply rule apply(rule mult_nonneg_nonneg)
+        unfolding setsum_right_distrib[THEN sym] and * using as and ** by auto
       hence "u x *\<^sub>R x + (1 - u x) *\<^sub>R ((inverse (1 - u x)) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) f) \<in>s" 
-	apply(rule_tac asm(1)[THEN bspec, THEN bspec, THEN spec, THEN mp, THEN spec, THEN mp, THEN mp]) using as and ** False by auto 
+        apply(rule_tac asm(1)[THEN bspec, THEN bspec, THEN spec, THEN mp, THEN spec, THEN mp, THEN mp]) using as and ** False by auto 
       thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2) and False by auto qed
   qed auto thus "t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" by auto
 qed
@@ -586,13 +586,13 @@
     { fix x e::real assume as:"0 \<le> x" "x \<le> 1" "0 < e"
       { fix y have *:"(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2) = (y - x) *\<^sub>R x1 - (y - x) *\<^sub>R x2"
           by (simp add: algebra_simps)
-	assume "\<bar>y - x\<bar> < e / norm (x1 - x2)"
-	hence "norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e"
-	  unfolding * and scaleR_right_diff_distrib[THEN sym]
-	  unfolding less_divide_eq using n by auto  }
+        assume "\<bar>y - x\<bar> < e / norm (x1 - x2)"
+        hence "norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e"
+          unfolding * and scaleR_right_diff_distrib[THEN sym]
+          unfolding less_divide_eq using n by auto  }
       hence "\<exists>d>0. \<forall>y. \<bar>y - x\<bar> < d \<longrightarrow> norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e"
-	apply(rule_tac x="e / norm (x1 - x2)" in exI) using as
-	apply auto unfolding zero_less_divide_iff using n by simp  }  note * = this
+        apply(rule_tac x="e / norm (x1 - x2)" in exI) using as
+        apply auto unfolding zero_less_divide_iff using n by simp  }  note * = this
 
     have "\<exists>x\<ge>0. x \<le> 1 \<and> (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e1 \<and> (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e2"
       apply(rule connected_real_lemma) apply (simp add: `x1\<in>e1` `x2\<in>e2` dist_commute)+
@@ -837,7 +837,7 @@
     proof(cases "u * v1 + v * v2 = 0")
       have *:"\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps)
       case True hence **:"u * v1 = 0" "v * v2 = 0" apply- apply(rule_tac [!] ccontr)
-	using mult_nonneg_nonneg[OF `u\<ge>0` `v1\<ge>0`] mult_nonneg_nonneg[OF `v\<ge>0` `v2\<ge>0`] by auto
+        using mult_nonneg_nonneg[OF `u\<ge>0` `v1\<ge>0`] mult_nonneg_nonneg[OF `v\<ge>0` `v2\<ge>0`] by auto
       hence "u * u1 + v * u2 = 1" using as(3) obt1(3) obt2(3) by auto
       thus ?thesis unfolding obt1(5) obt2(5) * using assms hull_subset[of s convex] by(auto simp add: ** scaleR_right_distrib)
     next
@@ -845,12 +845,12 @@
       also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) 
       also have "\<dots> = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto
       case False have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" apply -
-	apply(rule add_nonneg_nonneg) prefer 4 apply(rule add_nonneg_nonneg) apply(rule_tac [!] mult_nonneg_nonneg)
-	using as(1,2) obt1(1,2) obt2(1,2) by auto 
+        apply(rule add_nonneg_nonneg) prefer 4 apply(rule add_nonneg_nonneg) apply(rule_tac [!] mult_nonneg_nonneg)
+        using as(1,2) obt1(1,2) obt2(1,2) by auto 
       thus ?thesis unfolding obt1(5) obt2(5) unfolding * and ** using False
-	apply(rule_tac x="((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) defer
-	apply(rule convex_convex_hull[of s, unfolded convex_def, rule_format]) using obt1(4) obt2(4)
-	unfolding add_divide_distrib[THEN sym] and real_0_le_divide_iff
+        apply(rule_tac x="((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) defer
+        apply(rule convex_convex_hull[of s, unfolded convex_def, rule_format]) using obt1(4) obt2(4)
+        unfolding add_divide_distrib[THEN sym] and real_0_le_divide_iff
         by (auto simp add: scaleR_left_distrib scaleR_right_distrib)
     qed note * = this
     have u1:"u1 \<le> 1" apply(rule ccontr) unfolding obt1(3)[THEN sym] and not_le using obt1(2) by auto
@@ -904,7 +904,7 @@
     next def j \<equiv> "i - k1"
       case False with i have "j \<in> {1..k2}" unfolding j_def by auto
       thus ?thesis unfolding j_def[symmetric] using False
-	using mult_nonneg_nonneg[of v "u2 j"] and uv(2) y(1)[THEN bspec[where x=j]] by auto qed
+        using mult_nonneg_nonneg[of v "u2 j"] and uv(2) y(1)[THEN bspec[where x=j]] by auto qed
   qed(auto simp add: not_le x(2,3) y(2,3) uv(3))
 qed
 
@@ -960,8 +960,8 @@
     have fin':"\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
     { fix j assume "j\<in>{1..k}"
       hence "y j \<in> p" "0 \<le> setsum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
-	using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp
-	apply(rule setsum_nonneg) using obt(1) by auto } 
+        using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp
+        apply(rule setsum_nonneg) using obt(1) by auto } 
     moreover
     have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v}) = 1"  
       unfolding setsum_image_gen[OF fin, THEN sym] using obt(2) by auto
@@ -1192,7 +1192,7 @@
       assume as:"\<forall>x\<in>s. 0 \<le> w x"
       hence "setsum w (s - {v}) \<ge> 0" apply(rule_tac setsum_nonneg) by auto
       hence "setsum w s > 0" unfolding setsum_diff1'[OF obt(1) `v\<in>s`]
-	using as[THEN bspec[where x=v]] and `v\<in>s` using `w v \<noteq> 0` by auto
+        using as[THEN bspec[where x=v]] and `v\<in>s` using `w v \<noteq> 0` by auto
       thus False using wv(1) by auto
     qed hence "i\<noteq>{}" unfolding i_def by auto
 
@@ -1201,12 +1201,12 @@
     have t:"\<forall>v\<in>s. u v + t * w v \<ge> 0" proof
       fix v assume "v\<in>s" hence v:"0\<le>u v" using obt(4)[THEN bspec[where x=v]] by auto
       show"0 \<le> u v + t * w v" proof(cases "w v < 0")
-	case False thus ?thesis apply(rule_tac add_nonneg_nonneg) 
-	  using v apply simp apply(rule mult_nonneg_nonneg) using `t\<ge>0` by auto next
-	case True hence "t \<le> u v / (- w v)" using `v\<in>s`
-	  unfolding t_def i_def apply(rule_tac Min_le) using obt(1) by auto 
-	thus ?thesis unfolding real_0_le_add_iff
-	  using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[THEN sym]]] by auto
+        case False thus ?thesis apply(rule_tac add_nonneg_nonneg) 
+          using v apply simp apply(rule mult_nonneg_nonneg) using `t\<ge>0` by auto next
+        case True hence "t \<le> u v / (- w v)" using `v\<in>s`
+          unfolding t_def i_def apply(rule_tac Min_le) using obt(1) by auto 
+        thus ?thesis unfolding real_0_le_add_iff
+          using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[THEN sym]]] by auto
       qed qed
 
     obtain a where "a\<in>s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
@@ -1349,59 +1349,59 @@
     case (Suc n)
     show ?case proof(cases "n=0")
       case True have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s"
-	unfolding expand_set_eq and mem_Collect_eq proof(rule, rule)
-	fix x assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-	then obtain t where t:"finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" by auto
-	show "x\<in>s" proof(cases "card t = 0")
-	  case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by(simp add: convex_hull_empty)
-	next
-	  case False hence "card t = Suc 0" using t(3) `n=0` by auto
-	  then obtain a where "t = {a}" unfolding card_Suc_eq by auto
-	  thus ?thesis using t(2,4) by (simp add: convex_hull_singleton)
-	qed
+        unfolding expand_set_eq and mem_Collect_eq proof(rule, rule)
+        fix x assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
+        then obtain t where t:"finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" by auto
+        show "x\<in>s" proof(cases "card t = 0")
+          case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by(simp add: convex_hull_empty)
+        next
+          case False hence "card t = Suc 0" using t(3) `n=0` by auto
+          then obtain a where "t = {a}" unfolding card_Suc_eq by auto
+          thus ?thesis using t(2,4) by (simp add: convex_hull_singleton)
+        qed
       next
-	fix x assume "x\<in>s"
-	thus "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-	  apply(rule_tac x="{x}" in exI) unfolding convex_hull_singleton by auto 
+        fix x assume "x\<in>s"
+        thus "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
+          apply(rule_tac x="{x}" in exI) unfolding convex_hull_singleton by auto 
       qed thus ?thesis using assms by simp
     next
       case False have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
-	{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 
-	0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> {x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> x \<in> convex hull t}}"
-	unfolding expand_set_eq and mem_Collect_eq proof(rule,rule)
-	fix x assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
+        { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 
+        0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> {x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> x \<in> convex hull t}}"
+        unfolding expand_set_eq and mem_Collect_eq proof(rule,rule)
+        fix x assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
           0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)"
-	then obtain u v c t where obt:"x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
+        then obtain u v c t where obt:"x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
           "0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n"  "v \<in> convex hull t" by auto
-	moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u t"
-	  apply(rule mem_convex) using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex]
-	  using obt(7) and hull_mono[of t "insert u t"] by auto
-	ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-	  apply(rule_tac x="insert u t" in exI) by (auto simp add: card_insert_if)
+        moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u t"
+          apply(rule mem_convex) using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex]
+          using obt(7) and hull_mono[of t "insert u t"] by auto
+        ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
+          apply(rule_tac x="insert u t" in exI) by (auto simp add: card_insert_if)
       next
-	fix x assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-	then obtain t where t:"finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" by auto
-	let ?P = "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
+        fix x assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
+        then obtain t where t:"finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" by auto
+        let ?P = "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
           0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)"
-	show ?P proof(cases "card t = Suc n")
-	  case False hence "card t \<le> n" using t(3) by auto
-	  thus ?P apply(rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using `w\<in>s` and t
-	    by(auto intro!: exI[where x=t])
-	next
-	  case True then obtain a u where au:"t = insert a u" "a\<notin>u" apply(drule_tac card_eq_SucD) by auto
-	  show ?P proof(cases "u={}")
-	    case True hence "x=a" using t(4)[unfolded au] by auto
-	    show ?P unfolding `x=a` apply(rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI)
-	      using t and `n\<noteq>0` unfolding au by(auto intro!: exI[where x="{a}"] simp add: convex_hull_singleton)
-	  next
-	    case False obtain ux vx b where obt:"ux\<ge>0" "vx\<ge>0" "ux + vx = 1" "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
-	      using t(4)[unfolded au convex_hull_insert[OF False]] by auto
-	    have *:"1 - vx = ux" using obt(3) by auto
-	    show ?P apply(rule_tac x=a in exI, rule_tac x=b in exI, rule_tac x=vx in exI)
-	      using obt and t(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)]
-	      by(auto intro!: exI[where x=u])
-	  qed
-	qed
+        show ?P proof(cases "card t = Suc n")
+          case False hence "card t \<le> n" using t(3) by auto
+          thus ?P apply(rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using `w\<in>s` and t
+            by(auto intro!: exI[where x=t])
+        next
+          case True then obtain a u where au:"t = insert a u" "a\<notin>u" apply(drule_tac card_eq_SucD) by auto
+          show ?P proof(cases "u={}")
+            case True hence "x=a" using t(4)[unfolded au] by auto
+            show ?P unfolding `x=a` apply(rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI)
+              using t and `n\<noteq>0` unfolding au by(auto intro!: exI[where x="{a}"] simp add: convex_hull_singleton)
+          next
+            case False obtain ux vx b where obt:"ux\<ge>0" "vx\<ge>0" "ux + vx = 1" "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
+              using t(4)[unfolded au convex_hull_insert[OF False]] by auto
+            have *:"1 - vx = ux" using obt(3) by auto
+            show ?P apply(rule_tac x=a in exI, rule_tac x=b in exI, rule_tac x=vx in exI)
+              using obt and t(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)]
+              by(auto intro!: exI[where x=u])
+          qed
+        qed
       qed
       thus ?thesis using compact_convex_combinations[OF assms Suc] by simp 
     qed
@@ -1449,43 +1449,43 @@
     show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)"
     proof(cases "y\<in>convex hull s")
       case True then obtain z where "z\<in>convex hull s" "norm (y - a) < norm (z - a)"
-	using as(3)[THEN bspec[where x=y]] and y(2) by auto
+        using as(3)[THEN bspec[where x=y]] and y(2) by auto
       thus ?thesis apply(rule_tac x=z in bexI) unfolding convex_hull_insert[OF False] by auto
     next
       case False show ?thesis  using obt(3) proof(cases "u=0", case_tac[!] "v=0")
-	assume "u=0" "v\<noteq>0" hence "y = b" using obt by auto
-	thus ?thesis using False and obt(4) by auto
+        assume "u=0" "v\<noteq>0" hence "y = b" using obt by auto
+        thus ?thesis using False and obt(4) by auto
       next
-	assume "u\<noteq>0" "v=0" hence "y = x" using obt by auto
-	thus ?thesis using y(2) by auto
+        assume "u\<noteq>0" "v=0" hence "y = x" using obt by auto
+        thus ?thesis using y(2) by auto
       next
-	assume "u\<noteq>0" "v\<noteq>0"
-	then obtain w where w:"w>0" "w<u" "w<v" using real_lbound_gt_zero[of u v] and obt(1,2) by auto
-	have "x\<noteq>b" proof(rule ccontr) 
-	  assume "\<not> x\<noteq>b" hence "y=b" unfolding obt(5)
-	    using obt(3) by(auto simp add: scaleR_left_distrib[THEN sym])
-	  thus False using obt(4) and False by simp qed
-	hence *:"w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
-	show ?thesis using dist_increases_online[OF *, of a y]
- 	proof(erule_tac disjE)
-	  assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
-	  hence "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)"
-	    unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
-	  moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
-	    unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
-	    apply(rule_tac x="u + w" in exI) apply rule defer 
-	    apply(rule_tac x="v - w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
-	  ultimately show ?thesis by auto
-	next
-	  assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
-	  hence "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)"
-	    unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
-	  moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
-	    unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
-	    apply(rule_tac x="u - w" in exI) apply rule defer 
-	    apply(rule_tac x="v + w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
-	  ultimately show ?thesis by auto
-	qed
+        assume "u\<noteq>0" "v\<noteq>0"
+        then obtain w where w:"w>0" "w<u" "w<v" using real_lbound_gt_zero[of u v] and obt(1,2) by auto
+        have "x\<noteq>b" proof(rule ccontr) 
+          assume "\<not> x\<noteq>b" hence "y=b" unfolding obt(5)
+            using obt(3) by(auto simp add: scaleR_left_distrib[THEN sym])
+          thus False using obt(4) and False by simp qed
+        hence *:"w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
+        show ?thesis using dist_increases_online[OF *, of a y]
+        proof(erule_tac disjE)
+          assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
+          hence "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)"
+            unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
+          moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
+            unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
+            apply(rule_tac x="u + w" in exI) apply rule defer 
+            apply(rule_tac x="v - w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
+          ultimately show ?thesis by auto
+        next
+          assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
+          hence "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)"
+            unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
+          moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
+            unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
+            apply(rule_tac x="u - w" in exI) apply rule defer 
+            apply(rule_tac x="v + w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
+          ultimately show ?thesis by auto
+        qed
       qed auto
     qed
   qed auto
@@ -1704,8 +1704,8 @@
       assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z"
       then obtain u where "u>0" "u\<le>1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" by auto
       thus False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]]
-	using assms(1)[unfolded convex_alt, THEN bspec[where x=y]]
-	using `x\<in>s` `y\<in>s` by (auto simp add: dist_commute algebra_simps) qed
+        using assms(1)[unfolded convex_alt, THEN bspec[where x=y]]
+        using `x\<in>s` `y\<in>s` by (auto simp add: dist_commute algebra_simps) qed
     moreover have "0 < norm (y - z) ^ 2" using `y\<in>s` `z\<notin>s` by auto
     hence "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp
     ultimately show "inner (y - z) z + (norm (y - z))\<twosuperior> / 2 < inner (y - z) x"
@@ -2079,10 +2079,10 @@
     have "v=1" apply(rule ccontr) unfolding neq_iff apply(erule disjE) proof-
       assume "v<1" thus False using v(3)[THEN spec[where x=1]] using x and fs by auto next
       assume "v>1" thus False using assms(3)[THEN bspec[where x="v *\<^sub>R x"], THEN spec[where x="inverse v"]]
-	using v and x and fs unfolding inverse_less_1_iff by auto qed
+        using v and x and fs unfolding inverse_less_1_iff by auto qed
     show "u *\<^sub>R x \<in> s \<longleftrightarrow> u \<le> 1" apply rule  using v(3)[unfolded `v=1`, THEN spec[where x=u]] proof-
       assume "u\<le>1" thus "u *\<^sub>R x \<in> s" apply(cases "u=1")
-	using assms(3)[THEN bspec[where x=x], THEN spec[where x=u]] using `0\<le>u` and x and fs by auto qed auto qed
+        using assms(3)[THEN bspec[where x=x], THEN spec[where x=u]] using `0\<le>u` and x and fs by auto qed auto qed
 
   have "\<exists>surf. homeomorphism (frontier s) sphere pi surf"
     apply(rule homeomorphism_compact) apply(rule compact_frontier[OF assms(1)])
@@ -2117,10 +2117,10 @@
     have "norm x *\<^sub>R surf (pi x) \<in> s" proof(cases "x=0 \<or> norm x = 1") 
       case False hence "pi x \<in> sphere" "norm x < 1" using pi(1)[of x] as by(auto simp add: dist_norm)
       thus ?thesis apply(rule_tac assms(3)[rule_format, THEN DiffD1])
-	apply(rule_tac fs[unfolded subset_eq, rule_format])
-	unfolding surf(5)[THEN sym] by auto
+        apply(rule_tac fs[unfolded subset_eq, rule_format])
+        unfolding surf(5)[THEN sym] by auto
     next case True thus ?thesis apply rule defer unfolding pi_def apply(rule fs[unfolded subset_eq, rule_format])
-	unfolding  surf(5)[unfolded sphere_def, THEN sym] using `0\<in>s` by auto qed } note hom = this
+        unfolding  surf(5)[unfolded sphere_def, THEN sym] using `0\<in>s` by auto qed } note hom = this
 
   { fix x assume "x\<in>s"
     hence "x \<in> (\<lambda>x. norm x *\<^sub>R surf (pi x)) ` cball 0 1" proof(cases "x=0")
@@ -2131,19 +2131,19 @@
       hence "pi (surf (pi x)) = pi x" apply(rule_tac surf(4)[rule_format]) by assumption
       hence **:"norm x *\<^sub>R (?a *\<^sub>R surf (pi x)) = x" apply(rule_tac scaleR_left_imp_eq[OF invn]) unfolding pi_def using invn by auto
       hence *:"?a * norm x > 0" and"?a > 0" "?a \<noteq> 0" using surf(5) `0\<notin>frontier s` apply -
-	apply(rule_tac mult_pos_pos) using False[unfolded zero_less_norm_iff[THEN sym]] by auto
+        apply(rule_tac mult_pos_pos) using False[unfolded zero_less_norm_iff[THEN sym]] by auto
       have "norm (surf (pi x)) \<noteq> 0" using ** False by auto
       hence "norm x = norm ((?a * norm x) *\<^sub>R surf (pi x))"
-	unfolding norm_scaleR abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`] by auto
+        unfolding norm_scaleR abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`] by auto
       moreover have "pi x = pi ((inverse (norm (surf (pi x))) * norm x) *\<^sub>R surf (pi x))" 
-	unfolding pi(2)[OF *] surf(4)[rule_format, OF pix] ..
+        unfolding pi(2)[OF *] surf(4)[rule_format, OF pix] ..
       moreover have "surf (pi x) \<in> frontier s" using surf(5) pix by auto
       hence "dist 0 (inverse (norm (surf (pi x))) *\<^sub>R x) \<le> 1" unfolding dist_norm
-	using ** and * using front_smul[THEN bspec[where x="surf (pi x)"], THEN spec[where x="norm x * ?a"]]
-	using False `x\<in>s` by(auto simp add:field_simps)
+        using ** and * using front_smul[THEN bspec[where x="surf (pi x)"], THEN spec[where x="norm x * ?a"]]
+        using False `x\<in>s` by(auto simp add:field_simps)
       ultimately show ?thesis unfolding image_iff apply(rule_tac x="inverse (norm (surf(pi x))) *\<^sub>R x" in bexI)
-	apply(subst injpi[THEN sym]) unfolding abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`]
-	unfolding pi(2)[OF `?a > 0`] by auto
+        apply(subst injpi[THEN sym]) unfolding abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`]
+        unfolding pi(2)[OF `?a > 0`] by auto
     qed } note hom2 = this
 
   show ?thesis apply(subst homeomorphic_sym) apply(rule homeomorphic_compact[where f="\<lambda>x. norm x *\<^sub>R surf (pi x)"])
@@ -2152,39 +2152,39 @@
     fix x::"real^'n" assume as:"x \<in> cball 0 1"
     thus "continuous (at x) (\<lambda>x. norm x *\<^sub>R surf (pi x))" proof(cases "x=0")
       case False thus ?thesis apply(rule_tac continuous_mul, rule_tac continuous_at_norm)
-	using cont_surfpi unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def by auto
+        using cont_surfpi unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def by auto
     next guess a using UNIV_witness[where 'a = 'n] ..
       obtain B where B:"\<forall>x\<in>s. norm x \<le> B" using compact_imp_bounded[OF assms(1)] unfolding bounded_iff by auto
       hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis a" in ballE) defer apply(erule_tac x="basis a" in ballE)
-	unfolding Ball_def mem_cball dist_norm by (auto simp add: norm_basis[unfolded One_nat_def])
+        unfolding Ball_def mem_cball dist_norm by (auto simp add: norm_basis[unfolded One_nat_def])
       case True show ?thesis unfolding True continuous_at Lim_at apply(rule,rule) apply(rule_tac x="e / B" in exI)
-	apply(rule) apply(rule divide_pos_pos) prefer 3 apply(rule,rule,erule conjE)
-	unfolding norm_0 scaleR_zero_left dist_norm diff_0_right norm_scaleR abs_norm_cancel proof-
-	fix e and x::"real^'n" assume as:"norm x < e / B" "0 < norm x" "0<e"
-	hence "surf (pi x) \<in> frontier s" using pi(1)[of x] unfolding surf(5)[THEN sym] by auto
-	hence "norm (surf (pi x)) \<le> B" using B fs by auto
-	hence "norm x * norm (surf (pi x)) \<le> norm x * B" using as(2) by auto
-	also have "\<dots> < e / B * B" apply(rule mult_strict_right_mono) using as(1) `B>0` by auto
-	also have "\<dots> = e" using `B>0` by auto
-	finally show "norm x * norm (surf (pi x)) < e" by assumption
+        apply(rule) apply(rule divide_pos_pos) prefer 3 apply(rule,rule,erule conjE)
+        unfolding norm_0 scaleR_zero_left dist_norm diff_0_right norm_scaleR abs_norm_cancel proof-
+        fix e and x::"real^'n" assume as:"norm x < e / B" "0 < norm x" "0<e"
+        hence "surf (pi x) \<in> frontier s" using pi(1)[of x] unfolding surf(5)[THEN sym] by auto
+        hence "norm (surf (pi x)) \<le> B" using B fs by auto
+        hence "norm x * norm (surf (pi x)) \<le> norm x * B" using as(2) by auto
+        also have "\<dots> < e / B * B" apply(rule mult_strict_right_mono) using as(1) `B>0` by auto
+        also have "\<dots> = e" using `B>0` by auto
+        finally show "norm x * norm (surf (pi x)) < e" by assumption
       qed(insert `B>0`, auto) qed
   next { fix x assume as:"surf (pi x) = 0"
       have "x = 0" proof(rule ccontr)
-	assume "x\<noteq>0" hence "pi x \<in> sphere" using pi(1) by auto
-	hence "surf (pi x) \<in> frontier s" using surf(5) by auto
-	thus False using `0\<notin>frontier s` unfolding as by simp qed
+        assume "x\<noteq>0" hence "pi x \<in> sphere" using pi(1) by auto
+        hence "surf (pi x) \<in> frontier s" using surf(5) by auto
+        thus False using `0\<notin>frontier s` unfolding as by simp qed
     } note surf_0 = this
     show "inj_on (\<lambda>x. norm x *\<^sub>R surf (pi x)) (cball 0 1)" unfolding inj_on_def proof(rule,rule,rule)
       fix x y assume as:"x \<in> cball 0 1" "y \<in> cball 0 1" "norm x *\<^sub>R surf (pi x) = norm y *\<^sub>R surf (pi y)"
       thus "x=y" proof(cases "x=0 \<or> y=0") 
-	case True thus ?thesis using as by(auto elim: surf_0) next
-	case False
-	hence "pi (surf (pi x)) = pi (surf (pi y))" using as(3)
-	  using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"] by auto
-	moreover have "pi x \<in> sphere" "pi y \<in> sphere" using pi(1) False by auto
-	ultimately have *:"pi x = pi y" using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]] by auto 
-	moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0)
-	ultimately show ?thesis using injpi by auto qed qed
+        case True thus ?thesis using as by(auto elim: surf_0) next
+        case False
+        hence "pi (surf (pi x)) = pi (surf (pi y))" using as(3)
+          using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"] by auto
+        moreover have "pi x \<in> sphere" "pi y \<in> sphere" using pi(1) False by auto
+        ultimately have *:"pi x = pi y" using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]] by auto 
+        moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0)
+        ultimately show ?thesis using injpi by auto qed qed
   qed auto qed
 
 lemma homeomorphic_convex_compact_lemma: fixes s::"(real^'n::finite) set"
@@ -2396,33 +2396,33 @@
       have "xi \<in> (\<lambda>i. x$i) ` {i. x$i \<noteq> 0}" unfolding xi_def apply(rule Min_in) using False by auto
       then obtain i where i':"x$i = xi" "x$i \<noteq> 0" by auto
       have i:"\<And>j. x$j > 0 \<Longrightarrow> x$i \<le> x$j"
-	unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff
-	defer apply(rule_tac x=j in bexI) using i' by auto
+        unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff
+        defer apply(rule_tac x=j in bexI) using i' by auto
       have i01:"x$i \<le> 1" "x$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] using i'(2) `x$i \<noteq> 0`
-	by(auto simp add: Cart_lambda_beta) 
+        by(auto simp add: Cart_lambda_beta) 
       show ?thesis proof(cases "x$i=1")
-	case True have "\<forall>j\<in>{i. x$i \<noteq> 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof-
-	  fix j assume "x $ j \<noteq> 0" "x $ j \<noteq> 1"
-	  hence j:"x$j \<in> {0<..<1}" using Suc(2) by(auto simp add: vector_less_eq_def elim!:allE[where x=j])
-	  hence "x$j \<in> op $ x ` {i. x $ i \<noteq> 0}" by auto 
-	  hence "x$j \<ge> x$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto
-	  thus False using True Suc(2) j by(auto simp add: vector_less_eq_def elim!:ballE[where x=j]) qed
+        case True have "\<forall>j\<in>{i. x$i \<noteq> 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof-
+          fix j assume "x $ j \<noteq> 0" "x $ j \<noteq> 1"
+          hence j:"x$j \<in> {0<..<1}" using Suc(2) by(auto simp add: vector_less_eq_def elim!:allE[where x=j])
+          hence "x$j \<in> op $ x ` {i. x $ i \<noteq> 0}" by auto 
+          hence "x$j \<ge> x$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto
+          thus False using True Suc(2) j by(auto simp add: vector_less_eq_def elim!:ballE[where x=j]) qed
         thus "x\<in>convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
-	  by(auto simp add: Cart_lambda_beta)
+          by(auto simp add: Cart_lambda_beta)
       next let ?y = "\<lambda>j. if x$j = 0 then 0 else (x$j - x$i) / (1 - x$i)"
-	case False hence *:"x = x$i *\<^sub>R (\<chi> j. if x$j = 0 then 0 else 1) + (1 - x$i) *\<^sub>R (\<chi> j. ?y j)" unfolding Cart_eq
-	  by(auto simp add: Cart_lambda_beta vector_add_component vector_smult_component vector_minus_component field_simps)
-	{ fix j have "x$j \<noteq> 0 \<Longrightarrow> 0 \<le> (x $ j - x $ i) / (1 - x $ i)" "(x $ j - x $ i) / (1 - x $ i) \<le> 1"
-	    apply(rule_tac divide_nonneg_pos) using i(1)[of j] using False i01
-	    using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps Cart_lambda_beta) 
-	  hence "0 \<le> ?y j \<and> ?y j \<le> 1" by auto }
-	moreover have "i\<in>{j. x$j \<noteq> 0} - {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0}" using i01 by(auto simp add: Cart_lambda_beta)
-	hence "{j. x$j \<noteq> 0} \<noteq> {j. ((\<chi> j. ?y j)::real^'n::finite) $ j \<noteq> 0}" by auto
-	hence **:"{j. ((\<chi> j. ?y j)::real^'n::finite) $ j \<noteq> 0} \<subset> {j. x$j \<noteq> 0}" apply - apply rule by(auto simp add: Cart_lambda_beta)  
-	have "card {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0} \<le> n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto
-	ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format])
-	  apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1))
-	  unfolding mem_interval using i01 Suc(3) by (auto simp add: Cart_lambda_beta)
+        case False hence *:"x = x$i *\<^sub>R (\<chi> j. if x$j = 0 then 0 else 1) + (1 - x$i) *\<^sub>R (\<chi> j. ?y j)" unfolding Cart_eq
+          by(auto simp add: Cart_lambda_beta vector_add_component vector_smult_component vector_minus_component field_simps)
+        { fix j have "x$j \<noteq> 0 \<Longrightarrow> 0 \<le> (x $ j - x $ i) / (1 - x $ i)" "(x $ j - x $ i) / (1 - x $ i) \<le> 1"
+            apply(rule_tac divide_nonneg_pos) using i(1)[of j] using False i01
+            using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps Cart_lambda_beta) 
+          hence "0 \<le> ?y j \<and> ?y j \<le> 1" by auto }
+        moreover have "i\<in>{j. x$j \<noteq> 0} - {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0}" using i01 by(auto simp add: Cart_lambda_beta)
+        hence "{j. x$j \<noteq> 0} \<noteq> {j. ((\<chi> j. ?y j)::real^'n::finite) $ j \<noteq> 0}" by auto
+        hence **:"{j. ((\<chi> j. ?y j)::real^'n::finite) $ j \<noteq> 0} \<subset> {j. x$j \<noteq> 0}" apply - apply rule by(auto simp add: Cart_lambda_beta)  
+        have "card {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0} \<le> n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto
+        ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format])
+          apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1))
+          unfolding mem_interval using i01 Suc(3) by (auto simp add: Cart_lambda_beta)
       qed qed qed } note * = this
   show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule 
     apply(rule_tac n2="CARD('n)" in *) prefer 3 apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule
@@ -2448,10 +2448,10 @@
     unfolding image_iff defer apply(erule bexE) proof-
     fix y assume as:"y\<in>{x - ?d .. x + ?d}"
     { fix i::'n have "x $ i \<le> d + y $ i" "y $ i \<le> d + x $ i" using as[unfolded mem_interval, THEN spec[where x=i]]
-	by(auto simp add: vector_component)
+        by(auto simp add: vector_component)
       hence "1 \<ge> inverse d * (x $ i - y $ i)" "1 \<ge> inverse d * (y $ i - x $ i)"
-	apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[THEN sym]
-	using assms by(auto simp add: field_simps right_inverse) 
+        apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[THEN sym]
+        using assms by(auto simp add: field_simps right_inverse) 
       hence "inverse d * (x $ i * 2) \<le> 2 + inverse d * (y $ i * 2)"
             "inverse d * (y $ i * 2) \<le> 2 + inverse d * (x $ i * 2)" by(auto simp add:field_simps) }
     hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> {0..1}" unfolding mem_interval using assms
@@ -2487,34 +2487,34 @@
       case False def t \<equiv> "k / norm (y - x)"
       have "2 < t" "0<t" unfolding t_def using as False and `k>0` by(auto simp add:field_simps)
       have "y\<in>s" apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm
-	apply(rule order_trans[of _ "2 * norm (x - y)"]) using as by(auto simp add: field_simps norm_minus_commute) 
+        apply(rule order_trans[of _ "2 * norm (x - y)"]) using as by(auto simp add: field_simps norm_minus_commute) 
       { def w \<equiv> "x + t *\<^sub>R (y - x)"
-	have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
-	  unfolding t_def using `k>0` by auto
-	have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp add: algebra_simps)
-	also have "\<dots> = 0"  using `t>0` by(auto simp add:field_simps)
-	finally have w:"(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
-	have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
-	hence "(f w - f x) / t < e"
-	  using B(2)[OF `w\<in>s`] and B(2)[OF `x\<in>s`] using `t>0` by(auto simp add:field_simps) 
-	hence th1:"f y - f x < e" apply- apply(rule le_less_trans) defer apply assumption
-	  using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
-	  using `0<t` `2<t` and `x\<in>s` `w\<in>s` by(auto simp add:field_simps) }
+        have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
+          unfolding t_def using `k>0` by auto
+        have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp add: algebra_simps)
+        also have "\<dots> = 0"  using `t>0` by(auto simp add:field_simps)
+        finally have w:"(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
+        have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
+        hence "(f w - f x) / t < e"
+          using B(2)[OF `w\<in>s`] and B(2)[OF `x\<in>s`] using `t>0` by(auto simp add:field_simps) 
+        hence th1:"f y - f x < e" apply- apply(rule le_less_trans) defer apply assumption
+          using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
+          using `0<t` `2<t` and `x\<in>s` `w\<in>s` by(auto simp add:field_simps) }
       moreover 
       { def w \<equiv> "x - t *\<^sub>R (y - x)"
-	have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
-	  unfolding t_def using `k>0` by auto
-	have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp add: algebra_simps)
-	also have "\<dots>=x" using `t>0` by (auto simp add:field_simps)
-	finally have w:"(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
-	have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
-	hence *:"(f w - f y) / t < e" using B(2)[OF `w\<in>s`] and B(2)[OF `y\<in>s`] using `t>0` by(auto simp add:field_simps) 
-	have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" 
-	  using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
-	  using `0<t` `2<t` and `y\<in>s` `w\<in>s` by (auto simp add:field_simps)
-	also have "\<dots> = (f w + t * f y) / (1 + t)" using `t>0` unfolding real_divide_def by (auto simp add:field_simps)
-	also have "\<dots> < e + f y" using `t>0` * `e>0` by(auto simp add:field_simps)
-	finally have "f x - f y < e" by auto }
+        have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
+          unfolding t_def using `k>0` by auto
+        have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp add: algebra_simps)
+        also have "\<dots>=x" using `t>0` by (auto simp add:field_simps)
+        finally have w:"(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
+        have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
+        hence *:"(f w - f y) / t < e" using B(2)[OF `w\<in>s`] and B(2)[OF `y\<in>s`] using `t>0` by(auto simp add:field_simps) 
+        have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" 
+          using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
+          using `0<t` `2<t` and `y\<in>s` `w\<in>s` by (auto simp add:field_simps)
+        also have "\<dots> = (f w + t * f y) / (1 + t)" using `t>0` unfolding real_divide_def by (auto simp add:field_simps)
+        also have "\<dots> < e + f y" using `t>0` * `e>0` by(auto simp add:field_simps)
+        finally have "f x - f y < e" by auto }
       ultimately show ?thesis by auto 
     qed(insert `0<e`, auto) 
   qed(insert `0<e` `0<k` `0<B`, auto simp add:field_simps intro!:mult_pos_pos) qed
@@ -2567,7 +2567,7 @@
   hence "\<forall>y\<in>cball x d. abs (f y) \<le> k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof
     fix y assume y:"y\<in>cball x d"
     { fix i::'n have "x $ i - d \<le> y $ i"  "y $ i \<le> x $ i + d" 
-	using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by(auto simp add: vector_component)  }
+        using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by(auto simp add: vector_component)  }
     thus "f y \<le> k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm 
       by(auto simp add: vector_component_simps) qed
   hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous)
@@ -2671,22 +2671,22 @@
     apply rule apply(erule exE, (erule conjE)+) apply(subst dist_triangle_eq) proof-
       fix u assume as:"x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" 
       hence *:"a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
-	unfolding as(1) by(auto simp add:algebra_simps)
+        unfolding as(1) by(auto simp add:algebra_simps)
       show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
-	unfolding norm_minus_commute[of x a] * Cart_eq using as(2,3)
-	by(auto simp add: vector_component_simps field_simps)
+        unfolding norm_minus_commute[of x a] * Cart_eq using as(2,3)
+        by(auto simp add: vector_component_simps field_simps)
     next assume as:"dist a b = dist a x + dist x b"
       have "norm (a - x) / norm (a - b) \<le> 1" unfolding divide_le_eq_1_pos[OF Fal2] unfolding as[unfolded dist_norm] norm_ge_zero by auto 
       thus "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" apply(rule_tac x="dist a x / dist a b" in exI)
-	unfolding dist_norm Cart_eq apply- apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4 proof rule
-	  fix i::'n have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i =
-	    ((norm (a - b) - norm (a - x)) * (a $ i) + norm (a - x) * (b $ i)) / norm (a - b)"
-	    using Fal by(auto simp add:vector_component_simps field_simps)
-	  also have "\<dots> = x$i" apply(rule divide_eq_imp[OF Fal])
-	    unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq Cart_eq,rule_format, of i]
-	    by(auto simp add:field_simps vector_component_simps)
-	  finally show "x $ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i" by auto
-	qed(insert Fal2, auto) qed qed
+        unfolding dist_norm Cart_eq apply- apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4 proof rule
+          fix i::'n have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i =
+            ((norm (a - b) - norm (a - x)) * (a $ i) + norm (a - x) * (b $ i)) / norm (a - b)"
+            using Fal by(auto simp add:vector_component_simps field_simps)
+          also have "\<dots> = x$i" apply(rule divide_eq_imp[OF Fal])
+            unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq Cart_eq,rule_format, of i]
+            by(auto simp add:field_simps vector_component_simps)
+          finally show "x $ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i" by auto
+        qed(insert Fal2, auto) qed qed
 
 lemma between_midpoint: fixes a::"real^'n::finite" shows
   "between (a,b) (midpoint a b)" (is ?t1) 
@@ -2730,12 +2730,12 @@
     case False hence x:"x islimpt s" using assms(3)[unfolded closure_def] by auto
     show ?thesis proof(cases "e=1")
       case True obtain y where "y\<in>s" "y \<noteq> x" "dist y x < 1"
-	using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
+        using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
       thus ?thesis apply(rule_tac x=y in bexI) unfolding True using `d>0` by auto next
       case False hence "0 < e * d / (1 - e)" and *:"1 - e > 0"
-	using `e\<le>1` `e>0` `d>0` by(auto intro!:mult_pos_pos divide_pos_pos)
+        using `e\<le>1` `e>0` `d>0` by(auto intro!:mult_pos_pos divide_pos_pos)
       then obtain y where "y\<in>s" "y \<noteq> x" "dist y x < e * d / (1 - e)"
-	using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
+        using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
       thus ?thesis apply(rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] by auto qed qed
   then obtain y where "y\<in>s" and y:"norm (y - x) * (1 - e) < e * d" by auto
   def z \<equiv> "c + ((1 - e) / e) *\<^sub>R (x - y)"
@@ -2805,12 +2805,12 @@
     fix y assume y:"dist x y < min (Min (op $ x ` UNIV)) ?d"
     have "setsum (op $ y) UNIV \<le> setsum (\<lambda>i. x$i + ?d) UNIV" proof(rule setsum_mono)
       fix i::'n have "abs (y$i - x$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i]
-	using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add:vector_component_simps norm_minus_commute)
+        using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add:vector_component_simps norm_minus_commute)
       thus "y $ i \<le> x $ i + ?d" by auto qed
     also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant card_enum real_eq_of_nat using dimindex_ge_1 by(auto simp add: Suc_le_eq)
     finally show "(\<forall>i. 0 \<le> y $ i) \<and> setsum (op $ y) UNIV \<le> 1" apply- proof(rule,rule)
       fix i::'n have "norm (x - y) < x$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
-	using Min_gr_iff[of "op $ x ` dimset x"] dimindex_ge_1 by auto
+        using Min_gr_iff[of "op $ x ` dimset x"] dimindex_ge_1 by auto
       thus "0 \<le> y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by(auto simp add: vector_component_simps)
     qed auto qed auto qed
 
@@ -3110,10 +3110,10 @@
   { fix x assume as:"g 1 = g 0" "x \<in> {0..1::real^1}" " \<forall>y\<in>{0..1} \<inter> {x. \<not> a $ 1 + x $ 1 \<le> 1}. g x \<noteq> g (a + y - 1)" 
     hence "\<exists>y\<in>{0..1} \<inter> {x. a $ 1 + x $ 1 \<le> 1}. g x = g (a + y)" proof(cases "a \<le> x")
       case False thus ?thesis apply(rule_tac x="1 + x - a" in bexI)
-	using as(1,2) and as(3)[THEN bspec[where x="1 + x - a"]] and assms(1)
-	by(auto simp add:vector_component_simps field_simps atomize_not) next
+        using as(1,2) and as(3)[THEN bspec[where x="1 + x - a"]] and assms(1)
+        by(auto simp add:vector_component_simps field_simps atomize_not) next
       case True thus ?thesis using as(1-2) and assms(1) apply(rule_tac x="x - a" in bexI)
-	by(auto simp add:vector_component_simps field_simps) qed }
+        by(auto simp add:vector_component_simps field_simps) qed }
   thus ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def 
     by(auto simp add:vector_component_simps image_iff) qed
 
@@ -3316,24 +3316,24 @@
       by(auto elim!: ballE simp add: not_less le_Suc_eq)
     fix k assume "k \<in> {2..CARD('n)}" thus "path_connected (?A k)" proof(induct k)
       case (Suc k) show ?case proof(cases "k = 1")
-	case False from Suc have d:"k \<in> {1..CARD('n)}" "Suc k \<in> {1..CARD('n)}" by auto
-	hence "\<psi> k \<noteq> \<psi> (Suc k)" using \<psi>[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=k]] by auto
-	hence **:"?basis k + ?basis (Suc k) \<in> {x. 0 < inner (?basis (Suc k)) x} \<inter> (?A k)" 
+        case False from Suc have d:"k \<in> {1..CARD('n)}" "Suc k \<in> {1..CARD('n)}" by auto
+        hence "\<psi> k \<noteq> \<psi> (Suc k)" using \<psi>[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=k]] by auto
+        hence **:"?basis k + ?basis (Suc k) \<in> {x. 0 < inner (?basis (Suc k)) x} \<inter> (?A k)" 
           "?basis k - ?basis (Suc k) \<in> {x. 0 > inner (?basis (Suc k)) x} \<inter> ({x. 0 < inner (?basis (Suc k)) x} \<union> (?A k))" using d
-	  by(auto simp add: inner_basis vector_component_simps intro!:bexI[where x=k])
-	show ?thesis unfolding * Un_assoc apply(rule path_connected_Un) defer apply(rule path_connected_Un) 
-	  prefer 5 apply(rule_tac[1-2] convex_imp_path_connected, rule convex_halfspace_lt, rule convex_halfspace_gt)
-	  apply(rule Suc(1)) apply(rule_tac[2-3] ccontr) using d ** False by auto
+          by(auto simp add: inner_basis vector_component_simps intro!:bexI[where x=k])
+        show ?thesis unfolding * Un_assoc apply(rule path_connected_Un) defer apply(rule path_connected_Un) 
+          prefer 5 apply(rule_tac[1-2] convex_imp_path_connected, rule convex_halfspace_lt, rule convex_halfspace_gt)
+          apply(rule Suc(1)) apply(rule_tac[2-3] ccontr) using d ** False by auto
       next case True hence d:"1\<in>{1..CARD('n)}" "2\<in>{1..CARD('n)}" using Suc(2) by auto
-	have ***:"Suc 1 = 2" by auto
-	have **:"\<And>s t P Q. s \<union> t \<union> {x. P x \<or> Q x} = (s \<union> {x. P x}) \<union> (t \<union> {x. Q x})" by auto
-	have "\<psi> 2 \<noteq> \<psi> (Suc 0)" apply(rule ccontr) using \<psi>[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=2]] using assms by auto
-	thus ?thesis unfolding * True unfolding ** neq_iff bex_disj_distrib apply -
-	  apply(rule path_connected_Un, rule_tac[1-2] path_connected_Un) defer 3 apply(rule_tac[1-4] convex_imp_path_connected) 
-	  apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I)
-	  apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I)
-	  apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I)
-	  using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps inner_basis)
+        have ***:"Suc 1 = 2" by auto
+        have **:"\<And>s t P Q. s \<union> t \<union> {x. P x \<or> Q x} = (s \<union> {x. P x}) \<union> (t \<union> {x. Q x})" by auto
+        have "\<psi> 2 \<noteq> \<psi> (Suc 0)" apply(rule ccontr) using \<psi>[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=2]] using assms by auto
+        thus ?thesis unfolding * True unfolding ** neq_iff bex_disj_distrib apply -
+          apply(rule path_connected_Un, rule_tac[1-2] path_connected_Un) defer 3 apply(rule_tac[1-4] convex_imp_path_connected) 
+          apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I)
+          apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I)
+          apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I)
+          using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps inner_basis)
   qed qed auto qed note lem = this
 
   have ***:"\<And>x::real^'n. (\<exists>i\<in>{1..CARD('n)}. inner (basis (\<psi> i)) x \<noteq> 0) \<longleftrightarrow> (\<exists>i. inner (basis i) x \<noteq> 0)"
--- a/src/HOL/Library/Determinants.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Determinants.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -139,9 +139,9 @@
     also have "\<dots> = setprod (\<lambda>i. ?di A i (p i)) ?U"
     proof-
       {fix i assume i: "i \<in> ?U"
-	from i permutes_inv_o[OF pU] permutes_in_image[OF pU]
-	have "((\<lambda>i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)"
-	  unfolding transp_def by (simp add: expand_fun_eq)}
+        from i permutes_inv_o[OF pU] permutes_in_image[OF pU]
+        have "((\<lambda>i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)"
+          unfolding transp_def by (simp add: expand_fun_eq)}
       then show "setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\<lambda>i. ?di A i (p i)) ?U" by (auto intro: setprod_cong)
     qed
     finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" using sth
@@ -280,7 +280,7 @@
   fixes A :: "'a::ordered_idom^'n^'n::finite"
   assumes ij: "i \<noteq> j"
   and r: "row i A = row j A"
-  shows	"det A = 0"
+  shows "det A = 0"
 proof-
   have tha: "\<And>(a::'a) b. a = b ==> b = - a ==> a = 0"
     by simp
@@ -298,7 +298,7 @@
   fixes A :: "'a::ordered_idom^'n^'n::finite"
   assumes ij: "i \<noteq> j"
   and r: "column i A = column j A"
-  shows	"det A = 0"
+  shows "det A = 0"
 apply (subst det_transp[symmetric])
 apply (rule det_identical_rows[OF ij])
 by (metis row_transp r)
@@ -631,23 +631,23 @@
     let ?B = "(\<chi> i. B$f i) :: 'a^'n^'n"
     {assume fni: "\<not> inj_on f ?U"
       then obtain i j where ij: "f i = f j" "i \<noteq> j"
-	unfolding inj_on_def by blast
+        unfolding inj_on_def by blast
       from ij
       have rth: "row i ?B = row j ?B" by (vector row_def)
       from det_identical_rows[OF ij(2) rth]
       have "det (\<chi> i. A$i$f i *s B$f i) = 0"
-	unfolding det_rows_mul by simp}
+        unfolding det_rows_mul by simp}
     moreover
     {assume fi: "inj_on f ?U"
       from f fi have fith: "\<And>i j. f i = f j \<Longrightarrow> i = j"
-	unfolding inj_on_def by metis
+        unfolding inj_on_def by metis
       note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]]
 
       {fix y
-	from fs f have "\<exists>x. f x = y" by blast
-	then obtain x where x: "f x = y" by blast
-	{fix z assume z: "f z = y" from fith x z have "z = x" by metis}
-	with x have "\<exists>!x. f x = y" by blast}
+        from fs f have "\<exists>x. f x = y" by blast
+        then obtain x where x: "f x = y" by blast
+        {fix z assume z: "f z = y" from fith x z have "z = x" by metis}
+        with x have "\<exists>!x. f x = y" by blast}
       with f(3) have "det (\<chi> i. A$i$f i *s B$f i) = 0" by blast}
     ultimately have "det (\<chi> i. A$i$f i *s B$f i) = 0" by blast}
   hence zth: "\<forall> f\<in> ?F - ?PU. det (\<chi> i. A$i$f i *s B$f i) = 0" by simp
@@ -665,23 +665,23 @@
       fix q assume qU: "q \<in> ?PU"
       hence q: "q permutes ?U" by blast
       from p q have pp: "permutation p" and pq: "permutation q"
-	unfolding permutation_permutes by auto
+        unfolding permutation_permutes by auto
       have th00: "of_int (sign p) * of_int (sign p) = (1::'a)"
-	"\<And>a. of_int (sign p) * (of_int (sign p) * a) = a"
-	unfolding mult_assoc[symmetric]	unfolding of_int_mult[symmetric]
-	by (simp_all add: sign_idempotent)
+        "\<And>a. of_int (sign p) * (of_int (sign p) * a) = a"
+        unfolding mult_assoc[symmetric] unfolding of_int_mult[symmetric]
+        by (simp_all add: sign_idempotent)
       have ths: "?s q = ?s p * ?s (q o inv p)"
-	using pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
-	by (simp add:  th00 mult_ac sign_idempotent sign_compose)
+        using pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
+        by (simp add:  th00 mult_ac sign_idempotent sign_compose)
       have th001: "setprod (\<lambda>i. B$i$ q (inv p i)) ?U = setprod ((\<lambda>i. B$i$ q (inv p i)) o p) ?U"
-	by (rule setprod_permute[OF p])
+        by (rule setprod_permute[OF p])
       have thp: "setprod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = setprod (\<lambda>i. A$i$p i) ?U * setprod (\<lambda>i. B$i$ q (inv p i)) ?U"
-	unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p]
-	apply (rule setprod_cong[OF refl])
-	using permutes_in_image[OF q] by vector
+        unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p]
+        apply (rule setprod_cong[OF refl])
+        using permutes_in_image[OF q] by vector
       show "?s q * setprod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\<lambda>i. B$i$(q o inv p) i) ?U)"
-	using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
-	by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose)
+        using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
+        by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose)
     qed
   }
   then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B"
@@ -875,11 +875,11 @@
     {fix i j
       let ?A = "transp ?mf ** ?mf"
       have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)"
-	"\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
-	by simp_all
+        "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
+        by simp_all
       from fd[rule_format, of "basis i" "basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
       have "?A$i$j = ?m1 $ i $ j"
-	by (simp add: dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def)}
+        by (simp add: dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def)}
     hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector
     with lf have ?rhs by blast}
   moreover
@@ -995,29 +995,29 @@
     moreover
     {assume "x = 0" "y \<noteq> 0"
       then have "dist (?g x) (?g y) = dist x y"
-	apply (simp add: dist_norm norm_mul)
-	apply (rule f1[rule_format])
-	by(simp add: norm_mul field_simps)}
+        apply (simp add: dist_norm norm_mul)
+        apply (rule f1[rule_format])
+        by(simp add: norm_mul field_simps)}
     moreover
     {assume "x \<noteq> 0" "y = 0"
       then have "dist (?g x) (?g y) = dist x y"
-	apply (simp add: dist_norm norm_mul)
-	apply (rule f1[rule_format])
-	by(simp add: norm_mul field_simps)}
+        apply (simp add: dist_norm norm_mul)
+        apply (rule f1[rule_format])
+        by(simp add: norm_mul field_simps)}
     moreover
     {assume z: "x \<noteq> 0" "y \<noteq> 0"
       have th00: "x = norm x *s (inverse (norm x) *s x)" "y = norm y *s (inverse (norm y) *s y)" "norm x *s f ((inverse (norm x) *s x)) = norm x *s f (inverse (norm x) *s x)"
-	"norm y *s f (inverse (norm y) *s y) = norm y *s f (inverse (norm y) *s y)"
-	"norm (inverse (norm x) *s x) = 1"
-	"norm (f (inverse (norm x) *s x)) = 1"
-	"norm (inverse (norm y) *s y) = 1"
-	"norm (f (inverse (norm y) *s y)) = 1"
-	"norm (f (inverse (norm x) *s x) - f (inverse (norm y) *s y)) =
-	norm (inverse (norm x) *s x - inverse (norm y) *s y)"
-	using z
-	by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_norm])
+        "norm y *s f (inverse (norm y) *s y) = norm y *s f (inverse (norm y) *s y)"
+        "norm (inverse (norm x) *s x) = 1"
+        "norm (f (inverse (norm x) *s x)) = 1"
+        "norm (inverse (norm y) *s y) = 1"
+        "norm (f (inverse (norm y) *s y)) = 1"
+        "norm (f (inverse (norm x) *s x) - f (inverse (norm y) *s y)) =
+        norm (inverse (norm x) *s x - inverse (norm y) *s y)"
+        using z
+        by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_norm])
       from z th0[OF th00] have "dist (?g x) (?g y) = dist x y"
-	by (simp add: dist_norm)}
+        by (simp add: dist_norm)}
     ultimately have "dist (?g x) (?g y) = dist x y" by blast}
   note thd = this
     show ?thesis
--- a/src/HOL/Library/Euclidean_Space.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Euclidean_Space.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,5 @@
-(* Title:      Library/Euclidean_Space
-   Author:     Amine Chaieb, University of Cambridge
+(*  Title:      Library/Euclidean_Space
+    Author:     Amine Chaieb, University of Cambridge
 *)
 
 header {* (Real) Vectors in Euclidean space, and elementary linear algebra.*}
@@ -917,11 +917,11 @@
       from le2 fa fb e12 alb have la: "l \<noteq> a" by metis
       hence lap: "l - a > 0" using alb by arith
       from e2[rule_format, OF le2] obtain e where
-	e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e2" by metis
+        e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e2" by metis
       from dst[OF alb e(1)] obtain d where
-	d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
+        d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
       have "\<exists>d'. d' < d \<and> d' >0 \<and> l - d' > a" using lap d(1)
-	apply ferrack by arith
+        apply ferrack by arith
       then obtain d' where d': "d' > 0" "d' < d" "l - d' > a" by metis
       from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e2" by metis
       from th0[rule_format, of "l - d'"] d' have "f (l - d') \<in> e2" by auto
@@ -933,16 +933,16 @@
     from le1 fa fb e12 alb have lb: "l \<noteq> b" by metis
       hence blp: "b - l > 0" using alb by arith
       from e1[rule_format, OF le1] obtain e where
-	e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e1" by metis
+        e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e1" by metis
       from dst[OF alb e(1)] obtain d where
-	d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
+        d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
       have "\<exists>d'. d' < d \<and> d' >0" using d(1) by dlo
       then obtain d' where d': "d' > 0" "d' < d" by metis
       from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e1" by auto
       hence "\<forall>y. l \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" using d' by auto
       with ale1 have "\<forall>y. a \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" by auto
       with l d' have False
-	by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) }
+        by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) }
     ultimately show ?thesis using alb by metis
 qed
 
@@ -1768,7 +1768,7 @@
     {assume C: "B < 0"
       have "norm (1::real ^ 'n) > 0" by (simp add: zero_less_norm_iff)
       with C have "B * norm (1:: real ^ 'n) < 0"
-	by (simp add: zero_compare_simps)
+        by (simp add: zero_compare_simps)
       with B[rule_format, of 1] norm_ge_zero[of "f 1"] have False by simp
     }
     then have Bp: "B \<ge> 0" by ferrack
@@ -1969,14 +1969,14 @@
     let ?w = "(\<chi> i. (f (basis i) \<bullet> y)) :: 'a ^ 'n"
     {fix x
       have "f x \<bullet> y = f (setsum (\<lambda>i. (x$i) *s basis i) ?N) \<bullet> y"
-	by (simp only: basis_expansion)
+        by (simp only: basis_expansion)
       also have "\<dots> = (setsum (\<lambda>i. (x$i) *s f (basis i)) ?N) \<bullet> y"
-	unfolding linear_setsum[OF lf fN]
-	by (simp add: linear_cmul[OF lf])
+        unfolding linear_setsum[OF lf fN]
+        by (simp add: linear_cmul[OF lf])
       finally have "f x \<bullet> y = x \<bullet> ?w"
-	apply (simp only: )
-	apply (simp add: dot_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps)
-	done}
+        apply (simp only: )
+        apply (simp add: dot_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps)
+        done}
   }
   then show ?thesis unfolding adjoint_def
     some_eq_ex[of "\<lambda>f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y"]
@@ -2551,17 +2551,17 @@
       by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero])
     {fix x :: "real ^'n"
       {assume "x = 0"
-	then have "norm (f x) \<le> b * norm x" by (simp add: linear_0[OF lf] bp)}
+        then have "norm (f x) \<le> b * norm x" by (simp add: linear_0[OF lf] bp)}
       moreover
       {assume x0: "x \<noteq> 0"
-	hence n0: "norm x \<noteq> 0" by (metis norm_eq_zero)
-	let ?c = "1/ norm x"
-	have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul)
-	with H have "norm (f(?c*s x)) \<le> b" by blast
-	hence "?c * norm (f x) \<le> b"
-	  by (simp add: linear_cmul[OF lf] norm_mul)
-	hence "norm (f x) \<le> b * norm x"
-	  using n0 norm_ge_zero[of x] by (auto simp add: field_simps)}
+        hence n0: "norm x \<noteq> 0" by (metis norm_eq_zero)
+        let ?c = "1/ norm x"
+        have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul)
+        with H have "norm (f(?c*s x)) \<le> b" by blast
+        hence "?c * norm (f x) \<le> b"
+          by (simp add: linear_cmul[OF lf] norm_mul)
+        hence "norm (f x) \<le> b * norm x"
+          using n0 norm_ge_zero[of x] by (auto simp add: field_simps)}
       ultimately have "norm (f x) \<le> b * norm x" by blast}
     then have ?rhs by blast}
   ultimately show ?thesis by blast
@@ -2580,15 +2580,15 @@
       unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def)
     {from rsup[OF Se b, unfolded onorm_def[symmetric]]
       show "norm (f x) <= onorm f * norm x"
-	apply -
-	apply (rule spec[where x = x])
-	unfolding norm_bound_generalize[OF lf, symmetric]
-	by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
+        apply -
+        apply (rule spec[where x = x])
+        unfolding norm_bound_generalize[OF lf, symmetric]
+        by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
     {
       show "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"
-	using rsup[OF Se b, unfolded onorm_def[symmetric]]
-	unfolding norm_bound_generalize[OF lf, symmetric]
-	by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
+        using rsup[OF Se b, unfolded onorm_def[symmetric]]
+        unfolding norm_bound_generalize[OF lf, symmetric]
+        by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
   }
 qed
 
@@ -3279,37 +3279,37 @@
 
     have "span_induct_alt_help S x"
       proof(rule span_induct[where x=x and S=S])
-	show "x \<in> span S" using x .
+        show "x \<in> span S" using x .
       next
-	fix x assume xS : "x \<in> S"
-	  from span_induct_alt_help_S[OF xS span_induct_alt_help_0, of 1]
-	  show "span_induct_alt_help S x" by simp
-	next
-	have "span_induct_alt_help S 0" by (rule span_induct_alt_help_0)
-	moreover
-	{fix x y assume h: "span_induct_alt_help S x" "span_induct_alt_help S y"
-	  from h
-	  have "span_induct_alt_help S (x + y)"
-	    apply (induct rule: span_induct_alt_help.induct)
-	    apply simp
-	    unfolding add_assoc
-	    apply (rule span_induct_alt_help_S)
-	    apply assumption
-	    apply simp
-	    done}
-	moreover
-	{fix c x assume xt: "span_induct_alt_help S x"
-	  then have "span_induct_alt_help S (c*s x)"
-	    apply (induct rule: span_induct_alt_help.induct)
-	    apply (simp add: span_induct_alt_help_0)
-	    apply (simp add: vector_smult_assoc vector_add_ldistrib)
-	    apply (rule span_induct_alt_help_S)
-	    apply assumption
-	    apply simp
-	    done
-	}
-	ultimately show "subspace (span_induct_alt_help S)"
-	  unfolding subspace_def mem_def Ball_def by blast
+        fix x assume xS : "x \<in> S"
+          from span_induct_alt_help_S[OF xS span_induct_alt_help_0, of 1]
+          show "span_induct_alt_help S x" by simp
+        next
+        have "span_induct_alt_help S 0" by (rule span_induct_alt_help_0)
+        moreover
+        {fix x y assume h: "span_induct_alt_help S x" "span_induct_alt_help S y"
+          from h
+          have "span_induct_alt_help S (x + y)"
+            apply (induct rule: span_induct_alt_help.induct)
+            apply simp
+            unfolding add_assoc
+            apply (rule span_induct_alt_help_S)
+            apply assumption
+            apply simp
+            done}
+        moreover
+        {fix c x assume xt: "span_induct_alt_help S x"
+          then have "span_induct_alt_help S (c*s x)"
+            apply (induct rule: span_induct_alt_help.induct)
+            apply (simp add: span_induct_alt_help_0)
+            apply (simp add: vector_smult_assoc vector_add_ldistrib)
+            apply (rule span_induct_alt_help_S)
+            apply assumption
+            apply simp
+            done
+        }
+        ultimately show "subspace (span_induct_alt_help S)"
+          unfolding subspace_def mem_def Ball_def by blast
       qed}
   with th0 show ?thesis by blast
 qed
@@ -3387,16 +3387,16 @@
   {fix x assume xS: "x \<in> S"
     {assume ab: "x = b"
       then have "?P x"
-	apply simp
-	apply (rule exI[where x="1"], simp)
-	by (rule span_0)}
+        apply simp
+        apply (rule exI[where x="1"], simp)
+        by (rule span_0)}
     moreover
     {assume ab: "x \<noteq> b"
       then have "?P x"  using xS
-	apply -
-	apply (rule exI[where x=0])
-	apply (rule span_superset)
-	by simp}
+        apply -
+        apply (rule exI[where x=0])
+        apply (rule span_superset)
+        by simp}
     ultimately have "?P x" by blast}
   moreover have "subspace ?P"
     unfolding subspace_def
@@ -3557,16 +3557,16 @@
     from fS SP x have th0: "finite (insert x S)" "insert x S \<subseteq> P" by blast+
     {assume xS: "x \<in> S"
       have S1: "S = (S - {x}) \<union> {x}"
-	and Sss:"finite (S - {x})" "finite {x}" "(S -{x}) \<inter> {x} = {}" using xS fS by auto
+        and Sss:"finite (S - {x})" "finite {x}" "(S -{x}) \<inter> {x} = {}" using xS fS by auto
       have "setsum (\<lambda>v. ?u v *s v) ?S =(\<Sum>v\<in>S - {x}. u v *s v) + (u x + c) *s x"
-	using xS
-	by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]]
-	  setsum_clauses(2)[OF fS] cong del: if_weak_cong)
+        using xS
+        by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]]
+          setsum_clauses(2)[OF fS] cong del: if_weak_cong)
       also have "\<dots> = (\<Sum>v\<in>S. u v *s v) + c *s x"
-	apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])
-	by (vector ring_simps)
+        apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])
+        by (vector ring_simps)
       also have "\<dots> = c*s x + y"
-	by (simp add: add_commute u)
+        by (simp add: add_commute u)
       finally have "setsum (\<lambda>v. ?u v *s v) ?S = c*s x + y" .
     then have "?Q ?S ?u (c*s x + y)" using th0 by blast}
   moreover
@@ -3621,7 +3621,7 @@
     have "setsum (\<lambda>v. ?u v *s v) ?S = setsum (\<lambda>v. (- (inverse (u ?a))) *s (u v *s v)) S - ?u v *s v"
       using fS vS uv
       by (simp add: setsum_diff1 vector_smult_lneg divide_inverse
-	vector_smult_assoc field_simps)
+        vector_smult_assoc field_simps)
     also have "\<dots> = ?a"
       unfolding setsum_cmul u
       using uv by (simp add: vector_smult_lneg)
@@ -3740,29 +3740,29 @@
   {assume aS: "a \<notin> S"
     {assume i: ?lhs
       then have ?rhs using aS
-	apply simp
-	apply (rule conjI)
-	apply (rule independent_mono)
-	apply assumption
-	apply blast
-	by (simp add: dependent_def)}
+        apply simp
+        apply (rule conjI)
+        apply (rule independent_mono)
+        apply assumption
+        apply blast
+        by (simp add: dependent_def)}
     moreover
     {assume i: ?rhs
       have ?lhs using i aS
-	apply simp
-	apply (auto simp add: dependent_def)
-	apply (case_tac "aa = a", auto)
-	apply (subgoal_tac "insert a S - {aa} = insert a (S - {aa})")
-	apply simp
-	apply (subgoal_tac "a \<in> span (insert aa (S - {aa}))")
-	apply (subgoal_tac "insert aa (S - {aa}) = S")
-	apply simp
-	apply blast
-	apply (rule in_span_insert)
-	apply assumption
-	apply blast
-	apply blast
-	done}
+        apply simp
+        apply (auto simp add: dependent_def)
+        apply (case_tac "aa = a", auto)
+        apply (subgoal_tac "insert a S - {aa} = insert a (S - {aa})")
+        apply simp
+        apply (subgoal_tac "a \<in> span (insert aa (S - {aa}))")
+        apply (subgoal_tac "insert aa (S - {aa}) = S")
+        apply simp
+        apply blast
+        apply (rule in_span_insert)
+        apply assumption
+        apply blast
+        apply blast
+        done}
     ultimately have ?thesis by blast}
   ultimately show ?thesis by blast
 qed
@@ -3837,7 +3837,7 @@
     from st(2) obtain b where b: "b \<in> t" "b \<notin> s" by blast
       from b have "t - {b} - s \<subset> t - s" by blast
       then have cardlt: "card (t - {b} - s) < n" using n ft
- 	by (auto intro: psubset_card_mono)
+        by (auto intro: psubset_card_mono)
       from b ft have ct0: "card t \<noteq> 0" by auto
     {assume stb: "s \<subseteq> span(t -{b})"
       from ft have ftb: "finite (t -{b})" by auto
@@ -3852,7 +3852,7 @@
       from u(1) ft b have "u hassize (card t - 1)" by auto
       then
       have th2: "insert b u hassize card t"
-	using  card_insert_disjoint[OF fu bu] ct0 by (auto simp add: hassize_def)
+        using  card_insert_disjoint[OF fu bu] ct0 by (auto simp add: hassize_def)
       from u(4) have "s \<subseteq> span u" .
       also have "\<dots> \<subseteq> span (insert b u)" apply (rule span_mono) by blast
       finally have th3: "s \<subseteq> span (insert b u)" .      from th0 th1 th2 th3 have th: "?P ?w"  by blast
@@ -3863,22 +3863,22 @@
       have ab: "a \<noteq> b" using a b by blast
       have at: "a \<notin> t" using a ab span_superset[of a "t- {b}"] by auto
       have mlt: "card ((insert a (t - {b})) - s) < n"
-	using cardlt ft n  a b by auto
+        using cardlt ft n  a b by auto
       have ft': "finite (insert a (t - {b}))" using ft by auto
       {fix x assume xs: "x \<in> s"
-	have t: "t \<subseteq> (insert b (insert a (t -{b})))" using b by auto
-	from b(1) have "b \<in> span t" by (simp add: span_superset)
-	have bs: "b \<in> span (insert a (t - {b}))"
-	  by (metis in_span_delete a sp mem_def subset_eq)
-	from xs sp have "x \<in> span t" by blast
-	with span_mono[OF t]
-	have x: "x \<in> span (insert b (insert a (t - {b})))" ..
-	from span_trans[OF bs x] have "x \<in> span (insert a (t - {b}))"  .}
+        have t: "t \<subseteq> (insert b (insert a (t -{b})))" using b by auto
+        from b(1) have "b \<in> span t" by (simp add: span_superset)
+        have bs: "b \<in> span (insert a (t - {b}))"
+          by (metis in_span_delete a sp mem_def subset_eq)
+        from xs sp have "x \<in> span t" by blast
+        with span_mono[OF t]
+        have x: "x \<in> span (insert b (insert a (t - {b})))" ..
+        from span_trans[OF bs x] have "x \<in> span (insert a (t - {b}))"  .}
       then have sp': "s \<subseteq> span (insert a (t - {b}))" by blast
 
       from H[rule_format, OF mlt ft' s sp' refl] obtain u where
-	u: "u hassize card (insert a (t -{b}))" "s \<subseteq> u" "u \<subseteq> s \<union> insert a (t -{b})"
-	"s \<subseteq> span u" by blast
+        u: "u hassize card (insert a (t -{b}))" "s \<subseteq> u" "u \<subseteq> s \<union> insert a (t -{b})"
+        "s \<subseteq> span u" by blast
       from u a b ft at ct0 have "?P u" by (auto simp add: hassize_def)
       then have ?ths by blast }
     ultimately have ?ths by blast
@@ -4165,35 +4165,35 @@
       from ya have Cy: "C = insert y (C - {y})" by blast
       have fth: "finite (C - {y})" using C by simp
       have "orthogonal x y"
-	using xa ya
-	unfolding orthogonal_def xa dot_lsub dot_rsub diff_eq_0_iff_eq
-	apply simp
-	apply (subst Cy)
-	using C(1) fth
-	apply (simp only: setsum_clauses)
-	thm dot_ladd
-	apply (auto simp add: dot_ladd dot_radd dot_lmult dot_rmult dot_eq_0 dot_sym[of y a] dot_lsum[OF fth])
-	apply (rule setsum_0')
-	apply clarsimp
-	apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
-	by auto}
+        using xa ya
+        unfolding orthogonal_def xa dot_lsub dot_rsub diff_eq_0_iff_eq
+        apply simp
+        apply (subst Cy)
+        using C(1) fth
+        apply (simp only: setsum_clauses)
+        thm dot_ladd
+        apply (auto simp add: dot_ladd dot_radd dot_lmult dot_rmult dot_eq_0 dot_sym[of y a] dot_lsum[OF fth])
+        apply (rule setsum_0')
+        apply clarsimp
+        apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
+        by auto}
     moreover
     {assume xa: "x \<noteq> ?a" "x \<in> C" and ya: "y = ?a"
       from xa have Cx: "C = insert x (C - {x})" by blast
       have fth: "finite (C - {x})" using C by simp
       have "orthogonal x y"
-	using xa ya
-	unfolding orthogonal_def ya dot_rsub dot_lsub diff_eq_0_iff_eq
-	apply simp
-	apply (subst Cx)
-	using C(1) fth
-	apply (simp only: setsum_clauses)
-	apply (subst dot_sym[of x])
-	apply (auto simp add: dot_radd dot_rmult dot_eq_0 dot_sym[of x a] dot_rsum[OF fth])
-	apply (rule setsum_0')
-	apply clarsimp
-	apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
-	by auto}
+        using xa ya
+        unfolding orthogonal_def ya dot_rsub dot_lsub diff_eq_0_iff_eq
+        apply simp
+        apply (subst Cx)
+        using C(1) fth
+        apply (simp only: setsum_clauses)
+        apply (subst dot_sym[of x])
+        apply (auto simp add: dot_radd dot_rmult dot_eq_0 dot_sym[of x a] dot_rsum[OF fth])
+        apply (rule setsum_0')
+        apply clarsimp
+        apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
+        by auto}
     moreover
     {assume xa: "x \<in> C" and ya: "y \<in> C"
       have "orthogonal x y" using xa ya xy C(4) unfolding pairwise_def by blast}
@@ -4258,13 +4258,13 @@
       from x have B': "B = insert x (B - {x})" by blast
       have fth: "finite (B - {x})" using fB by simp
       have "?a \<bullet> x = 0"
-	apply (subst B') using fB fth
-	unfolding setsum_clauses(2)[OF fth]
-	apply simp
-	apply (clarsimp simp add: dot_lsub dot_ladd dot_lmult dot_lsum dot_eq_0)
-	apply (rule setsum_0', rule ballI)
-	unfolding dot_sym
-	by (auto simp add: x field_simps dot_eq_0 intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])}
+        apply (subst B') using fB fth
+        unfolding setsum_clauses(2)[OF fth]
+        apply simp
+        apply (clarsimp simp add: dot_lsub dot_ladd dot_lmult dot_lsum dot_eq_0)
+        apply (rule setsum_0', rule ballI)
+        unfolding dot_sym
+        by (auto simp add: x field_simps dot_eq_0 intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])}
     then show "\<forall>x \<in> B. ?a \<bullet> x = 0" by blast
   qed
   with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"])
@@ -4365,14 +4365,14 @@
       using z .
     {fix k assume k: "z - k *s a \<in> span b"
       have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a"
-	by (simp add: ring_simps vector_sadd_rdistrib[symmetric])
+        by (simp add: ring_simps vector_sadd_rdistrib[symmetric])
       from span_sub[OF th0 k]
       have khz: "(k - ?h z) *s a \<in> span b" by (simp add: eq)
       {assume "k \<noteq> ?h z" hence k0: "k - ?h z \<noteq> 0" by simp
-	from k0 span_mul[OF khz, of "1 /(k - ?h z)"]
-	have "a \<in> span b" by (simp add: vector_smult_assoc)
-	with "2.prems"(1) "2.hyps"(2) have False
-	  by (auto simp add: dependent_def)}
+        from k0 span_mul[OF khz, of "1 /(k - ?h z)"]
+        have "a \<in> span b" by (simp add: vector_smult_assoc)
+        with "2.prems"(1) "2.hyps"(2) have False
+          by (auto simp add: dependent_def)}
       then have "k = ?h z" by blast}
     with th0 have "z - ?h z *s a \<in> span b \<and> (\<forall>k. z - k *s a \<in> span b \<longrightarrow> k = ?h z)" by blast}
   note h = this
@@ -4404,25 +4404,25 @@
   {fix x assume x: "x \<in> (insert a b)"
     {assume xa: "x = a"
       have ha1: "1 = ?h a"
-	apply (rule conjunct2[OF h, rule_format])
-	apply (metis span_superset insertI1)
-	using conjunct1[OF h, OF span_superset, OF insertI1]
-	by (auto simp add: span_0)
+        apply (rule conjunct2[OF h, rule_format])
+        apply (metis span_superset insertI1)
+        using conjunct1[OF h, OF span_superset, OF insertI1]
+        by (auto simp add: span_0)
 
       from xa ha1[symmetric] have "?g x = f x"
-	apply simp
-	using g(2)[rule_format, OF span_0, of 0]
-	by simp}
+        apply simp
+        using g(2)[rule_format, OF span_0, of 0]
+        by simp}
     moreover
     {assume xb: "x \<in> b"
       have h0: "0 = ?h x"
-	apply (rule conjunct2[OF h, rule_format])
-	apply (metis  span_superset insertI1 xb x)
-	apply simp
-	apply (metis span_superset xb)
-	done
+        apply (rule conjunct2[OF h, rule_format])
+        apply (metis  span_superset insertI1 xb x)
+        apply simp
+        apply (metis span_superset xb)
+        done
       have "?g x = f x"
-	by (simp add: h0[symmetric] g(3)[rule_format, OF xb])}
+        by (simp add: h0[symmetric] g(3)[rule_format, OF xb])}
     ultimately have "?g x = f x" using x by blast }
   ultimately show ?case apply - apply (rule exI[where x="?g"]) by blast
 qed
@@ -4687,7 +4687,7 @@
   {fix B :: "real ^'m^'n"  assume AB: "A ** B = mat 1"
     {fix x :: "real ^ 'm"
       have "A *v (B *v x) = x"
-	by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)}
+        by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)}
     hence "surj (op *v A)" unfolding surj_def by metis }
   moreover
   {assume sf: "surj (op *v A)"
@@ -4697,7 +4697,7 @@
 
     have "A ** (matrix g) = mat 1"
       unfolding matrix_eq  matrix_vector_mul_lid
-	matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
+        matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
       using g(2) unfolding o_def stupid_ext[symmetric] id_def
       .
     hence "\<exists>B. A ** (B::real^'m^'n) = mat 1" by blast
@@ -4716,9 +4716,9 @@
       and i: "i \<in> ?U"
       let ?x = "\<chi> i. c i"
       have th0:"A *v ?x = 0"
-	using c
-	unfolding matrix_mult_vsum Cart_eq
-	by auto
+        using c
+        unfolding matrix_mult_vsum Cart_eq
+        by auto
       from k[rule_format, OF th0] i
       have "c i = 0" by (vector Cart_eq)}
     hence ?rhs by blast}
@@ -4749,53 +4749,53 @@
   have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> span (columns A))" by blast
   {assume h: ?lhs
     {fix x:: "real ^'n"
-	from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m"
-	  where y: "setsum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
-	have "x \<in> span (columns A)"
-	  unfolding y[symmetric]
-	  apply (rule span_setsum[OF fU])
-	  apply clarify
-	  apply (rule span_mul)
-	  apply (rule span_superset)
-	  unfolding columns_def
-	  by blast}
+        from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m"
+          where y: "setsum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
+        have "x \<in> span (columns A)"
+          unfolding y[symmetric]
+          apply (rule span_setsum[OF fU])
+          apply clarify
+          apply (rule span_mul)
+          apply (rule span_superset)
+          unfolding columns_def
+          by blast}
     then have ?rhs unfolding rhseq by blast}
   moreover
   {assume h:?rhs
     let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y"
     {fix y have "?P y"
       proof(rule span_induct_alt[of ?P "columns A"])
-	show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
-	  apply (rule exI[where x=0])
-	  by (simp add: zero_index vector_smult_lzero)
+        show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
+          apply (rule exI[where x=0])
+          by (simp add: zero_index vector_smult_lzero)
       next
-	fix c y1 y2 assume y1: "y1 \<in> columns A" and y2: "?P y2"
-	from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"
-	  unfolding columns_def by blast
-	from y2 obtain x:: "real ^'m" where
-	  x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
-	let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m"
-	show "?P (c*s y1 + y2)"
-	  proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] cond_value_iff right_distrib cond_application_beta cong del: if_weak_cong)
-	    fix j
-	    have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
+        fix c y1 y2 assume y1: "y1 \<in> columns A" and y2: "?P y2"
+        from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"
+          unfolding columns_def by blast
+        from y2 obtain x:: "real ^'m" where
+          x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
+        let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m"
+        show "?P (c*s y1 + y2)"
+          proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] cond_value_iff right_distrib cond_application_beta cong del: if_weak_cong)
+            fix j
+            have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
            else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1)
-	      by (simp add: ring_simps)
-	    have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
+              by (simp add: ring_simps)
+            have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
            else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
-	      apply (rule setsum_cong[OF refl])
-	      using th by blast
-	    also have "\<dots> = setsum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
-	      by (simp add: setsum_addf)
-	    also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
-	      unfolding setsum_delta[OF fU]
-	      using i(1) by simp
-	    finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
+              apply (rule setsum_cong[OF refl])
+              using th by blast
+            also have "\<dots> = setsum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
+              by (simp add: setsum_addf)
+            also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
+              unfolding setsum_delta[OF fU]
+              using i(1) by simp
+            finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
            else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
-	  qed
-	next
-	  show "y \<in> span (columns A)" unfolding h by blast
-	qed}
+          qed
+        next
+          show "y \<in> span (columns A)" unfolding h by blast
+        qed}
     then have ?lhs unfolding lhseq ..}
   ultimately show ?thesis by blast
 qed
@@ -4844,22 +4844,22 @@
     {fix x y assume x: "x \<in> S" and y: "y \<in> S" and f: "f x = f y"
       from x fS have S0: "card S \<noteq> 0" by auto
       {assume xy: "x \<noteq> y"
-	have th: "card S \<le> card (f ` (S - {y}))"
-	  unfolding c
-	  apply (rule card_mono)
-	  apply (rule finite_imageI)
-	  using fS apply simp
-	  using h xy x y f unfolding subset_eq image_iff
-	  apply auto
-	  apply (case_tac "xa = f x")
-	  apply (rule bexI[where x=x])
-	  apply auto
-	  done
-	also have " \<dots> \<le> card (S -{y})"
-	  apply (rule card_image_le)
-	  using fS by simp
-	also have "\<dots> \<le> card S - 1" using y fS by simp
-	finally have False  using S0 by arith }
+        have th: "card S \<le> card (f ` (S - {y}))"
+          unfolding c
+          apply (rule card_mono)
+          apply (rule finite_imageI)
+          using fS apply simp
+          using h xy x y f unfolding subset_eq image_iff
+          apply auto
+          apply (case_tac "xa = f x")
+          apply (rule bexI[where x=x])
+          apply auto
+          done
+        also have " \<dots> \<le> card (S -{y})"
+          apply (rule card_image_le)
+          using fS by simp
+        also have "\<dots> \<le> card S - 1" using y fS by simp
+        finally have False  using S0 by arith }
       then have "x = y" by blast}
     then have ?rhs unfolding inj_on_def by blast}
   moreover
@@ -5311,26 +5311,26 @@
       then obtain u where u: "\<forall> x\<in> {0,x,y}. \<forall>y\<in> {0,x,y}. \<exists>c. x - y = c *s u" unfolding collinear_def by blast
       from u[rule_format, of x 0] u[rule_format, of y 0]
       obtain cx and cy where
-	cx: "x = cx*s u" and cy: "y = cy*s u"
-	by auto
+        cx: "x = cx*s u" and cy: "y = cy*s u"
+        by auto
       from cx x have cx0: "cx \<noteq> 0" by auto
       from cy y have cy0: "cy \<noteq> 0" by auto
       let ?d = "cy / cx"
       from cx cy cx0 have "y = ?d *s x"
-	by (simp add: vector_smult_assoc)
+        by (simp add: vector_smult_assoc)
       hence ?rhs using x y by blast}
     moreover
     {assume h: "?rhs"
       then obtain c where c: "y = c*s x" using x y by blast
       have ?lhs unfolding collinear_def c
-	apply (rule exI[where x=x])
-	apply auto
-	apply (rule exI[where x="- 1"], simp only: vector_smult_lneg vector_smult_lid)
-	apply (rule exI[where x= "-c"], simp only: vector_smult_lneg)
-	apply (rule exI[where x=1], simp)
-	apply (rule exI[where x="1 - c"], simp add: vector_smult_lneg vector_sub_rdistrib)
-	apply (rule exI[where x="c - 1"], simp add: vector_smult_lneg vector_sub_rdistrib)
-	done}
+        apply (rule exI[where x=x])
+        apply auto
+        apply (rule exI[where x="- 1"], simp only: vector_smult_lneg vector_smult_lid)
+        apply (rule exI[where x= "-c"], simp only: vector_smult_lneg)
+        apply (rule exI[where x=1], simp)
+        apply (rule exI[where x="1 - c"], simp add: vector_smult_lneg vector_sub_rdistrib)
+        apply (rule exI[where x="c - 1"], simp add: vector_smult_lneg vector_sub_rdistrib)
+        done}
     ultimately have ?thesis by blast}
   ultimately show ?thesis by blast
 qed
--- a/src/HOL/Library/Float.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Float.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -243,7 +243,7 @@
       fix x y z :: real
       assume "y \<noteq> 0"
       then have "(x * inverse y = z) = (x = z * y)"
-	by auto
+        by auto
     }
     note inverse = this
     have eq': "real a * (pow2 (b - b')) = real a'"
@@ -551,9 +551,9 @@
     } moreover
     { have "x + 1 \<le> x - x mod 2 + 2"
       proof -
-	have "x mod 2 < 2" using `0 < x` by auto
- 	hence "x < x - x mod 2 +  2" unfolding algebra_simps by auto
-	thus ?thesis by auto
+        have "x mod 2 < 2" using `0 < x` by auto
+        hence "x < x - x mod 2 +  2" unfolding algebra_simps by auto
+        thus ?thesis by auto
       qed
       also have "x - x mod 2 + 2 = (x div 2 + 1) * 2" unfolding algebra_simps using `0 < x` zdiv_zmod_equality2[of x 2 0] by auto
       also have "\<dots> \<le> 2^nat (bitlen (x div 2)) * 2" using hyp[OF `0 < x div 2`, THEN conjunct2] by (rule mult_right_mono, auto)
@@ -839,7 +839,7 @@
     have "0 < ?X div y"
     proof -
       have "2^nat (bitlen x - 1) \<le> y" and "y < 2^nat (bitlen y)"
-	using bitlen_bounds[OF `0 < x`, THEN conjunct1] bitlen_bounds[OF `0 < y`, THEN conjunct2] `x < y` by auto
+        using bitlen_bounds[OF `0 < x`, THEN conjunct1] bitlen_bounds[OF `0 < y`, THEN conjunct2] `x < y` by auto
       hence "(2::int)^nat (bitlen x - 1) < 2^nat (bitlen y)" by (rule order_le_less_trans)
       hence "bitlen x \<le> bitlen y" by auto
       hence len_less: "nat (bitlen x - 1) \<le> nat (int (n - 1) + bitlen y)" by auto
@@ -847,16 +847,16 @@
       have "x \<noteq> 0" and "y \<noteq> 0" using `0 < x` `0 < y` by auto
 
       have exp_eq: "nat (int (n - 1) + bitlen y) - nat (bitlen x - 1) = ?l"
-	using `bitlen x \<le> bitlen y` bitlen_ge1[OF `x \<noteq> 0`] bitlen_ge1[OF `y \<noteq> 0`] `0 < n` by auto
+        using `bitlen x \<le> bitlen y` bitlen_ge1[OF `x \<noteq> 0`] bitlen_ge1[OF `y \<noteq> 0`] `0 < n` by auto
 
       have "y * 2^nat (bitlen x - 1) \<le> y * x" 
-	using bitlen_bounds[OF `0 < x`, THEN conjunct1] `0 < y`[THEN less_imp_le] by (rule mult_left_mono)
+        using bitlen_bounds[OF `0 < x`, THEN conjunct1] `0 < y`[THEN less_imp_le] by (rule mult_left_mono)
       also have "\<dots> \<le> 2^nat (bitlen y) * x" using bitlen_bounds[OF `0 < y`, THEN conjunct2, THEN less_imp_le] `0 \<le> x` by (rule mult_right_mono)
       also have "\<dots> \<le> x * 2^nat (int (n - 1) + bitlen y)" unfolding mult_commute[of x] by (rule mult_right_mono, auto simp add: `0 \<le> x`)
       finally have "real y * 2^nat (bitlen x - 1) * inverse (2^nat (bitlen x - 1)) \<le> real x * 2^nat (int (n - 1) + bitlen y) * inverse (2^nat (bitlen x - 1))"
-	unfolding real_of_int_le_iff[symmetric] by auto
+        unfolding real_of_int_le_iff[symmetric] by auto
       hence "real y \<le> real x * (2^nat (int (n - 1) + bitlen y) / (2^nat (bitlen x - 1)))" 
-	unfolding real_mult_assoc real_divide_def by auto
+        unfolding real_mult_assoc real_divide_def by auto
       also have "\<dots> = real x * (2^(nat (int (n - 1) + bitlen y) - nat (bitlen x - 1)))" using power_diff[of "2::real", OF _ len_less] by auto
       finally have "y \<le> x * 2^?l" unfolding exp_eq unfolding real_of_int_le_iff[symmetric] by auto
       thus ?thesis using zdiv_greater_zero[OF `0 < y`] by auto
@@ -1194,19 +1194,19 @@
       case True
       have m: "m = m div ?p * ?p + 0" unfolding True[symmetric] using zdiv_zmod_equality2[where k=0, unfolded monoid_add_class.add_0_right, symmetric] .
       have "real (Float m e) = real (Float (m div ?p) (e + ?d))" unfolding real_of_float_simp arg_cong[OF m, of real]
-	by (auto simp add: pow2_add `0 < ?d` pow_d)
+        by (auto simp add: pow2_add `0 < ?d` pow_d)
       thus ?thesis
-	unfolding Float round_up.simps Let_def if_P[OF `m mod ?p = 0`] if_P[OF `0 < ?d`]
-	by auto
+        unfolding Float round_up.simps Let_def if_P[OF `m mod ?p = 0`] if_P[OF `0 < ?d`]
+        by auto
     next
       case False
       have "m = m div ?p * ?p + m mod ?p" unfolding zdiv_zmod_equality2[where k=0, unfolded monoid_add_class.add_0_right] ..
       also have "\<dots> \<le> (m div ?p + 1) * ?p" unfolding left_distrib zmult_1 by (rule add_left_mono, rule pos_mod_bound[OF `0 < ?p`, THEN less_imp_le])
       finally have "real (Float m e) \<le> real (Float (m div ?p + 1) (e + ?d))" unfolding real_of_float_simp add_commute[of e]
-	unfolding pow2_add mult_assoc[symmetric] real_of_int_le_iff[of m, symmetric]
-	by (auto intro!: mult_mono simp add: pow2_add `0 < ?d` pow_d)
+        unfolding pow2_add mult_assoc[symmetric] real_of_int_le_iff[of m, symmetric]
+        by (auto intro!: mult_mono simp add: pow2_add `0 < ?d` pow_d)
       thus ?thesis
-	unfolding Float round_up.simps Let_def if_not_P[OF `\<not> m mod ?p = 0`] if_P[OF `0 < ?d`] .
+        unfolding Float round_up.simps Let_def if_not_P[OF `\<not> m mod ?p = 0`] if_P[OF `0 < ?d`] .
     qed
   next
     case False
@@ -1264,7 +1264,7 @@
     have "real (m div 2^(nat ?l)) * pow2 ?l \<le> real m"
     proof -
       have "real (m div 2^(nat ?l)) * pow2 ?l = real (2^(nat ?l) * (m div 2^(nat ?l)))" unfolding real_of_int_mult real_of_int_power[symmetric] real_number_of unfolding pow2_int[symmetric] 
-	using `?l > 0` by auto
+        using `?l > 0` by auto
       also have "\<dots> \<le> real (2^(nat ?l) * (m div 2^(nat ?l)) + m mod 2^(nat ?l))" unfolding real_of_int_add by auto
       also have "\<dots> = real m" unfolding zmod_zdiv_equality[symmetric] ..
       finally show ?thesis by auto
--- a/src/HOL/Library/Formal_Power_Series.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -561,9 +561,9 @@
     have th1: "\<And>n. (2::real)^n >0" by auto
     {assume h: "dist a b > dist a c + dist b c"
       then have gt: "dist a b > dist a c" "dist a b > dist b c"
-	using pos by auto
+        using pos by auto
       from gt have gtn: "nab < nbc" "nab < nac"
-	unfolding dab dbc dac by (auto simp add: th1)
+        unfolding dab dbc dac by (auto simp add: th1)
       from nac'(1)[OF gtn(2)] nbc'(1)[OF gtn(1)]
       have "a$nab = b$nab" by simp
       with nab'(2) have False  by simp}
@@ -615,32 +615,32 @@
       from reals_power_lt_ex[OF rp th0] 
       obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0" by blast
       {fix n::nat
-	assume nn0: "n \<ge> n0"
-	then have thnn0: "(1/2)^n <= (1/2 :: real)^n0"
-	  by (auto intro: power_decreasing)
-	{assume "?s n = a" then have "dist (?s n) a < r" 
-	    unfolding dist_eq_0_iff[of "?s n" a, symmetric]
-	    using rp by (simp del: dist_eq_0_iff)}
-	moreover
-	{assume neq: "?s n \<noteq> a"
-	  from fps_eq_least_unique[OF neq] 
-	  obtain k where k: "leastP (\<lambda>i. ?s n $ i \<noteq> a$i) k" by blast
-	  have th0: "\<And>(a::'a fps) b. a \<noteq> b \<longleftrightarrow> (\<exists>n. a$n \<noteq> b$n)"
-	    by (simp add: fps_eq_iff)
-	  from neq have dth: "dist (?s n) a = (1/2)^k"
-	    unfolding th0 dist_fps_def
-	    unfolding the1_equality[OF fps_eq_least_unique[OF neq], OF k]
-	    by (auto simp add: inverse_eq_divide power_divide)
-
-	  from k have kn: "k > n"
-	    by (simp add: leastP_def setge_def fps_sum_rep_nth split:split_if_asm)
-	  then have "dist (?s n) a < (1/2)^n" unfolding dth
-	    by (auto intro: power_strict_decreasing)
-	  also have "\<dots> <= (1/2)^n0" using nn0
-	    by (auto intro: power_decreasing)
-	  also have "\<dots> < r" using n0 by simp
-	  finally have "dist (?s n) a < r" .}
-	ultimately have "dist (?s n) a < r" by blast}
+        assume nn0: "n \<ge> n0"
+        then have thnn0: "(1/2)^n <= (1/2 :: real)^n0"
+          by (auto intro: power_decreasing)
+        {assume "?s n = a" then have "dist (?s n) a < r" 
+            unfolding dist_eq_0_iff[of "?s n" a, symmetric]
+            using rp by (simp del: dist_eq_0_iff)}
+        moreover
+        {assume neq: "?s n \<noteq> a"
+          from fps_eq_least_unique[OF neq] 
+          obtain k where k: "leastP (\<lambda>i. ?s n $ i \<noteq> a$i) k" by blast
+          have th0: "\<And>(a::'a fps) b. a \<noteq> b \<longleftrightarrow> (\<exists>n. a$n \<noteq> b$n)"
+            by (simp add: fps_eq_iff)
+          from neq have dth: "dist (?s n) a = (1/2)^k"
+            unfolding th0 dist_fps_def
+            unfolding the1_equality[OF fps_eq_least_unique[OF neq], OF k]
+            by (auto simp add: inverse_eq_divide power_divide)
+
+          from k have kn: "k > n"
+            by (simp add: leastP_def setge_def fps_sum_rep_nth split:split_if_asm)
+          then have "dist (?s n) a < (1/2)^n" unfolding dth
+            by (auto intro: power_strict_decreasing)
+          also have "\<dots> <= (1/2)^n0" using nn0
+            by (auto intro: power_decreasing)
+          also have "\<dots> < r" using n0 by simp
+          finally have "dist (?s n) a < r" .}
+        ultimately have "dist (?s n) a < r" by blast}
       then have "\<exists>n0. \<forall> n \<ge> n0. dist (?s n) a < r " by blast}
     then show ?thesis  unfolding  LIMSEQ_def by blast
   qed
@@ -972,19 +972,19 @@
   {fix l assume k: "k = Suc l"
     {fix m assume mk: "m < k"
       {assume "m=0" hence "a^k $ m = 0" using startsby_zero_power[of a k] k a0
-	  by simp}
+          by simp}
       moreover
       {assume m0: "m \<noteq> 0"
-	have "a ^k $ m = (a^l * a) $m" by (simp add: k power_Suc mult_commute)
-	also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))" by (simp add: fps_mult_nth)
-	also have "\<dots> = 0" apply (rule setsum_0')
-	  apply auto
-	  apply (case_tac "aa = m")
-	  using a0
-	  apply simp
-	  apply (rule H[rule_format])
-	  using a0 k mk by auto
-	finally have "a^k $ m = 0" .}
+        have "a ^k $ m = (a^l * a) $m" by (simp add: k power_Suc mult_commute)
+        also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))" by (simp add: fps_mult_nth)
+        also have "\<dots> = 0" apply (rule setsum_0')
+          apply auto
+          apply (case_tac "aa = m")
+          using a0
+          apply simp
+          apply (rule H[rule_format])
+          using a0 k mk by auto
+        finally have "a^k $ m = 0" .}
     ultimately have "a^k $ m = 0" by blast}
     hence ?ths by blast}
   ultimately show ?ths by (cases k, auto)
@@ -1028,7 +1028,7 @@
     moreover
     {assume n: "n > 0"
       from startsby_zero_power[OF a0 n] eq a0 n have ?thesis
-	by (simp add: fps_inverse_def)}
+        by (simp add: fps_inverse_def)}
     ultimately have ?thesis by blast}
   moreover
   {assume a0: "a$0 \<noteq> 0"
@@ -1186,13 +1186,13 @@
       note th = Suc.hyps[symmetric]
       have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: ring_simps)
       also  have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
-	using th
-	unfolding fps_sub_nth by simp
+        using th
+        unfolding fps_sub_nth by simp
       also have "\<dots> = (if n < Suc (Suc k) then 0 else a n)"
-	unfolding X_power_mult_right_nth
-	apply (auto simp add: not_less fps_const_def)
-	apply (rule cong[of a a, OF refl])
-	by arith
+        unfolding X_power_mult_right_nth
+        apply (auto simp add: not_less fps_const_def)
+        apply (rule cong[of a a, OF refl])
+        by arith
       finally show ?case by simp
     qed
     finally have "?lhs $ n = ?rhs $ n"  .}
@@ -1238,16 +1238,16 @@
   have th0: "\<And>i. (1 - (X::'a fps)) $ i = (if i = 0 then 1 else if i = 1 then - 1 else 0)" by simp
   {fix n:: nat
     {assume "n=0" hence "a$n = ((1 - ?X) * ?sa) $ n"
-	by (simp add: fps_mult_nth)}
+        by (simp add: fps_mult_nth)}
     moreover
     {assume n0: "n \<noteq> 0"
       then have u: "{0} \<union> ({1} \<union> {2..n}) = {0..n}" "{1}\<union>{2..n} = {1..n}"
-	"{0..n - 1}\<union>{n} = {0..n}"
-	by (auto simp: expand_set_eq)
+        "{0..n - 1}\<union>{n} = {0..n}"
+        by (auto simp: expand_set_eq)
       have d: "{0} \<inter> ({1} \<union> {2..n}) = {}" "{1} \<inter> {2..n} = {}"
-	"{0..n - 1}\<inter>{n} ={}" using n0 by simp_all
+        "{0..n - 1}\<inter>{n} ={}" using n0 by simp_all
       have f: "finite {0}" "finite {1}" "finite {2 .. n}"
-	"finite {0 .. n - 1}" "finite {n}" by simp_all
+        "finite {0 .. n - 1}" "finite {n}" by simp_all
     have "((1 - ?X) * ?sa) $ n = setsum (\<lambda>i. (1 - ?X)$ i * ?sa $ (n - i)) {0 .. n}"
       by (simp add: fps_mult_nth)
     also have "\<dots> = a$n" unfolding th0
@@ -1526,29 +1526,29 @@
   {assume h: ?lhs
     {fix n have "b$n = c$n"
       proof(induct n rule: nat_less_induct)
-	fix n assume H: "\<forall>m<n. b$m = c$m"
-	{assume n0: "n=0"
-	  from h have "(b oo a)$n = (c oo a)$n" by simp
-	  hence "b$n = c$n" using n0 by (simp add: fps_compose_nth)}
-	moreover
-	{fix n1 assume n1: "n = Suc n1"
-	  have f: "finite {0 .. n1}" "finite {n}" by simp_all
-	  have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using n1 by auto
-	  have d: "{0 .. n1} \<inter> {n} = {}" using n1 by auto
-	  have seq: "(\<Sum>i = 0..n1. b $ i * a ^ i $ n) = (\<Sum>i = 0..n1. c $ i * a ^ i $ n)"
-	    apply (rule setsum_cong2)
-	    using H n1 by auto
-	  have th0: "(b oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n"
-	    unfolding fps_compose_nth setsum_Un_disjoint[OF f d, unfolded eq] seq
-	    using startsby_zero_power_nth_same[OF a0]
-	    by simp
-	  have th1: "(c oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n"
-	    unfolding fps_compose_nth setsum_Un_disjoint[OF f d, unfolded eq]
-	    using startsby_zero_power_nth_same[OF a0]
-	    by simp
-	  from h[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1
-	  have "b$n = c$n" by auto}
-	ultimately show "b$n = c$n" by (cases n, auto)
+        fix n assume H: "\<forall>m<n. b$m = c$m"
+        {assume n0: "n=0"
+          from h have "(b oo a)$n = (c oo a)$n" by simp
+          hence "b$n = c$n" using n0 by (simp add: fps_compose_nth)}
+        moreover
+        {fix n1 assume n1: "n = Suc n1"
+          have f: "finite {0 .. n1}" "finite {n}" by simp_all
+          have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using n1 by auto
+          have d: "{0 .. n1} \<inter> {n} = {}" using n1 by auto
+          have seq: "(\<Sum>i = 0..n1. b $ i * a ^ i $ n) = (\<Sum>i = 0..n1. c $ i * a ^ i $ n)"
+            apply (rule setsum_cong2)
+            using H n1 by auto
+          have th0: "(b oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n"
+            unfolding fps_compose_nth setsum_Un_disjoint[OF f d, unfolded eq] seq
+            using startsby_zero_power_nth_same[OF a0]
+            by simp
+          have th1: "(c oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n"
+            unfolding fps_compose_nth setsum_Un_disjoint[OF f d, unfolded eq]
+            using startsby_zero_power_nth_same[OF a0]
+            by simp
+          from h[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1
+          have "b$n = c$n" by auto}
+        ultimately show "b$n = c$n" by (cases n, auto)
       qed}
     then have ?rhs by (simp add: fps_eq_iff)}
   ultimately show ?thesis by blast
@@ -1580,11 +1580,11 @@
       have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})" using i by auto
       from xs have "Suc n = foldl op + 0 xs" by (simp add: natpermute_def)
       also have "\<dots> = setsum (nth xs) {0..<Suc k}" unfolding foldl_add_setsum using xs
-	by (simp add: natpermute_def)
+        by (simp add: natpermute_def)
       also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
-	unfolding eqs  setsum_Un_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
-	unfolding setsum_Un_disjoint[OF fths(2) fths(3) d(2)]
-	by simp
+        unfolding eqs  setsum_Un_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
+        unfolding setsum_Un_disjoint[OF fths(2) fths(3) d(2)]
+        by simp
       finally have False using c' by simp}
     then show "((r,Suc k,a,xs!i), r, Suc k, a, Suc n) \<in> ?R"
       apply auto by (metis not_less)}
@@ -1655,44 +1655,44 @@
     from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
     {fix z have "?r ^ Suc k $ z = a$z"
       proof(induct z rule: nat_less_induct)
-	fix n assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
-	{assume "n = 0" hence "?r ^ Suc k $ n = a $n"
-	    using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp}
-	moreover
-	{fix n1 assume n1: "n = Suc n1"
-	  have fK: "finite {0..k}" by simp
-	  have nz: "n \<noteq> 0" using n1 by arith
-	  let ?Pnk = "natpermute n (k + 1)"
-	  let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
-	  let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
-	  have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
-	  have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
-	  have f: "finite ?Pnkn" "finite ?Pnknn"
-	    using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
-	    by (metis natpermute_finite)+
-	  let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
-	  have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
-	  proof(rule setsum_cong2)
-	    fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
-	    let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
-	  from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
-	    unfolding natpermute_contain_maximal by auto
-	  have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
-	    apply (rule setprod_cong, simp)
-	    using i r0 by (simp del: replicate.simps)
-	  also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
-	    unfolding setprod_gen_delta[OF fK] using i r0 by simp
-	  finally show ?ths .
-	qed
-	then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
-	  by (simp add: natpermute_max_card[OF nz, simplified])
-	also have "\<dots> = a$n - setsum ?f ?Pnknn"
-	  unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
-	finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
-	have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
-	  unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
-	also have "\<dots> = a$n" unfolding fn by simp
-	finally have "?r ^ Suc k $ n = a $n" .}
+        fix n assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
+        {assume "n = 0" hence "?r ^ Suc k $ n = a $n"
+            using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp}
+        moreover
+        {fix n1 assume n1: "n = Suc n1"
+          have fK: "finite {0..k}" by simp
+          have nz: "n \<noteq> 0" using n1 by arith
+          let ?Pnk = "natpermute n (k + 1)"
+          let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
+          let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
+          have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
+          have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
+          have f: "finite ?Pnkn" "finite ?Pnknn"
+            using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
+            by (metis natpermute_finite)+
+          let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
+          have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
+          proof(rule setsum_cong2)
+            fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
+            let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
+          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
+            unfolding natpermute_contain_maximal by auto
+          have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
+            apply (rule setprod_cong, simp)
+            using i r0 by (simp del: replicate.simps)
+          also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
+            unfolding setprod_gen_delta[OF fK] using i r0 by simp
+          finally show ?ths .
+        qed
+        then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
+          by (simp add: natpermute_max_card[OF nz, simplified])
+        also have "\<dots> = a$n - setsum ?f ?Pnknn"
+          unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
+        finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
+        have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
+          unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
+        also have "\<dots> = a$n" unfolding fn by simp
+        finally have "?r ^ Suc k $ n = a $n" .}
       ultimately  show "?r ^ Suc k $ n = a $n" by (cases n, auto)
     qed }
   then have ?thesis using r0 by (simp add: fps_eq_iff)}
@@ -1717,42 +1717,42 @@
     proof(induct z rule: nat_less_induct)
       fix n assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
       {assume "n = 0" hence "?r ^ Suc k $ n = a $n"
-	  using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp}
+          using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp}
       moreover
       {fix n1 assume n1: "n = Suc n1"
-	have fK: "finite {0..k}" by simp
-	have nz: "n \<noteq> 0" using n1 by arith
-	let ?Pnk = "natpermute n (k + 1)"
-	let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
-	let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
-	have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
-	have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
-	have f: "finite ?Pnkn" "finite ?Pnknn"
-	  using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
-	  by (metis natpermute_finite)+
-	let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
-	have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
-	proof(rule setsum_cong2)
-	  fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
-	  let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
-	  from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
-	    unfolding natpermute_contain_maximal by auto
-	  have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
-	    apply (rule setprod_cong, simp)
-	    using i r0 by (simp del: replicate.simps)
-	  also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
-	    unfolding setprod_gen_delta[OF fK] using i r0 by simp
-	  finally show ?ths .
-	qed
-	then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
-	  by (simp add: natpermute_max_card[OF nz, simplified])
-	also have "\<dots> = a$n - setsum ?f ?Pnknn"
-	  unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
-	finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
-	have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
-	  unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
-	also have "\<dots> = a$n" unfolding fn by simp
-	finally have "?r ^ Suc k $ n = a $n" .}
+        have fK: "finite {0..k}" by simp
+        have nz: "n \<noteq> 0" using n1 by arith
+        let ?Pnk = "natpermute n (k + 1)"
+        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
+        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
+        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
+        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
+        have f: "finite ?Pnkn" "finite ?Pnknn"
+          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
+          by (metis natpermute_finite)+
+        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
+        have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
+        proof(rule setsum_cong2)
+          fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
+          let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
+          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
+            unfolding natpermute_contain_maximal by auto
+          have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
+            apply (rule setprod_cong, simp)
+            using i r0 by (simp del: replicate.simps)
+          also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
+            unfolding setprod_gen_delta[OF fK] using i r0 by simp
+          finally show ?ths .
+        qed
+        then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
+          by (simp add: natpermute_max_card[OF nz, simplified])
+        also have "\<dots> = a$n - setsum ?f ?Pnknn"
+          unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
+        finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
+        have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
+          unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
+        also have "\<dots> = a$n" unfolding fn by simp
+        finally have "?r ^ Suc k $ n = a $n" .}
       ultimately  show "?r ^ Suc k $ n = a $n" by (cases n, auto)
   qed }
   then show ?thesis by (simp add: fps_eq_iff)
@@ -1783,78 +1783,78 @@
     from a0 have a0r0: "a$0 = ?r$0" by simp
     {fix n have "a $ n = ?r $ n"
       proof(induct n rule: nat_less_induct)
-	fix n assume h: "\<forall>m<n. a$m = ?r $m"
-	{assume "n = 0" hence "a$n = ?r $n" using a0 by simp }
-	moreover
-	{fix n1 assume n1: "n = Suc n1"
-	  have fK: "finite {0..k}" by simp
-	have nz: "n \<noteq> 0" using n1 by arith
-	let ?Pnk = "natpermute n (Suc k)"
-	let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
-	let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
-	have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
-	have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
-	have f: "finite ?Pnkn" "finite ?Pnknn"
-	  using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
-	  by (metis natpermute_finite)+
-	let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
-	let ?g = "\<lambda>v. \<Prod>j\<in>{0..k}. a $ v ! j"
-	have "setsum ?g ?Pnkn = setsum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
-	proof(rule setsum_cong2)
-	  fix v assume v: "v \<in> {xs \<in> natpermute n (Suc k). n \<in> set xs}"
-	  let ?ths = "(\<Prod>j\<in>{0..k}. a $ v ! j) = a $ n * (?r$0)^k"
-	  from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
-	    unfolding Suc_eq_plus1 natpermute_contain_maximal by (auto simp del: replicate.simps)
-	  have "(\<Prod>j\<in>{0..k}. a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then a $ n else r (Suc k) (b$0))"
-	    apply (rule setprod_cong, simp)
-	    using i a0 by (simp del: replicate.simps)
-	  also have "\<dots> = a $ n * (?r $ 0)^k"
-	    unfolding  setprod_gen_delta[OF fK] using i by simp
-	  finally show ?ths .
-	qed
-	then have th0: "setsum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
-	  by (simp add: natpermute_max_card[OF nz, simplified])
-	have th1: "setsum ?g ?Pnknn = setsum ?f ?Pnknn"
-	proof (rule setsum_cong2, rule setprod_cong, simp)
-	  fix xs i assume xs: "xs \<in> ?Pnknn" and i: "i \<in> {0..k}"
-	  {assume c: "n \<le> xs ! i"
-	    from xs i have "xs !i \<noteq> n" by (auto simp add: in_set_conv_nth natpermute_def)
-	    with c have c': "n < xs!i" by arith
-	    have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}" by simp_all
-	    have d: "{0 ..< i} \<inter> ({i} \<union> {i+1 ..< Suc k}) = {}" "{i} \<inter> {i+1..< Suc k} = {}" by auto
-	    have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})" using i by auto
-	    from xs have "n = foldl op + 0 xs" by (simp add: natpermute_def)
-	    also have "\<dots> = setsum (nth xs) {0..<Suc k}" unfolding foldl_add_setsum using xs
-	      by (simp add: natpermute_def)
-	    also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
-	      unfolding eqs  setsum_Un_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
-	      unfolding setsum_Un_disjoint[OF fths(2) fths(3) d(2)]
-	      by simp
-	    finally have False using c' by simp}
-	  then have thn: "xs!i < n" by arith
-	  from h[rule_format, OF thn]
-	  show "a$(xs !i) = ?r$(xs!i)" .
-	qed
-	have th00: "\<And>(x::'a). of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x"
-	  by (simp add: field_simps del: of_nat_Suc)
-	from H have "b$n = a^Suc k $ n" by (simp add: fps_eq_iff)
-	also have "a ^ Suc k$n = setsum ?g ?Pnkn + setsum ?g ?Pnknn"
-	  unfolding fps_power_nth_Suc
-	  using setsum_Un_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
-	    unfolded eq, of ?g] by simp
-	also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + setsum ?f ?Pnknn" unfolding th0 th1 ..
-	finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - setsum ?f ?Pnknn" by simp
-	then have "a$n = (b$n - setsum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)"
-	  apply -
-	  apply (rule eq_divide_imp')
-	  using r00
-	  apply (simp del: of_nat_Suc)
-	  by (simp add: mult_ac)
-	then have "a$n = ?r $n"
-	  apply (simp del: of_nat_Suc)
-	  unfolding fps_radical_def n1
-	  by (simp add: field_simps n1 th00 del: of_nat_Suc)}
-	ultimately show "a$n = ?r $ n" by (cases n, auto)
+        fix n assume h: "\<forall>m<n. a$m = ?r $m"
+        {assume "n = 0" hence "a$n = ?r $n" using a0 by simp }
+        moreover
+        {fix n1 assume n1: "n = Suc n1"
+          have fK: "finite {0..k}" by simp
+        have nz: "n \<noteq> 0" using n1 by arith
+        let ?Pnk = "natpermute n (Suc k)"
+        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
+        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
+        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
+        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
+        have f: "finite ?Pnkn" "finite ?Pnknn"
+          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
+          by (metis natpermute_finite)+
+        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
+        let ?g = "\<lambda>v. \<Prod>j\<in>{0..k}. a $ v ! j"
+        have "setsum ?g ?Pnkn = setsum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
+        proof(rule setsum_cong2)
+          fix v assume v: "v \<in> {xs \<in> natpermute n (Suc k). n \<in> set xs}"
+          let ?ths = "(\<Prod>j\<in>{0..k}. a $ v ! j) = a $ n * (?r$0)^k"
+          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
+            unfolding Suc_eq_plus1 natpermute_contain_maximal by (auto simp del: replicate.simps)
+          have "(\<Prod>j\<in>{0..k}. a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then a $ n else r (Suc k) (b$0))"
+            apply (rule setprod_cong, simp)
+            using i a0 by (simp del: replicate.simps)
+          also have "\<dots> = a $ n * (?r $ 0)^k"
+            unfolding  setprod_gen_delta[OF fK] using i by simp
+          finally show ?ths .
+        qed
+        then have th0: "setsum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
+          by (simp add: natpermute_max_card[OF nz, simplified])
+        have th1: "setsum ?g ?Pnknn = setsum ?f ?Pnknn"
+        proof (rule setsum_cong2, rule setprod_cong, simp)
+          fix xs i assume xs: "xs \<in> ?Pnknn" and i: "i \<in> {0..k}"
+          {assume c: "n \<le> xs ! i"
+            from xs i have "xs !i \<noteq> n" by (auto simp add: in_set_conv_nth natpermute_def)
+            with c have c': "n < xs!i" by arith
+            have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}" by simp_all
+            have d: "{0 ..< i} \<inter> ({i} \<union> {i+1 ..< Suc k}) = {}" "{i} \<inter> {i+1..< Suc k} = {}" by auto
+            have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})" using i by auto
+            from xs have "n = foldl op + 0 xs" by (simp add: natpermute_def)
+            also have "\<dots> = setsum (nth xs) {0..<Suc k}" unfolding foldl_add_setsum using xs
+              by (simp add: natpermute_def)
+            also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
+              unfolding eqs  setsum_Un_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
+              unfolding setsum_Un_disjoint[OF fths(2) fths(3) d(2)]
+              by simp
+            finally have False using c' by simp}
+          then have thn: "xs!i < n" by arith
+          from h[rule_format, OF thn]
+          show "a$(xs !i) = ?r$(xs!i)" .
+        qed
+        have th00: "\<And>(x::'a). of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x"
+          by (simp add: field_simps del: of_nat_Suc)
+        from H have "b$n = a^Suc k $ n" by (simp add: fps_eq_iff)
+        also have "a ^ Suc k$n = setsum ?g ?Pnkn + setsum ?g ?Pnknn"
+          unfolding fps_power_nth_Suc
+          using setsum_Un_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
+            unfolded eq, of ?g] by simp
+        also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + setsum ?f ?Pnknn" unfolding th0 th1 ..
+        finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - setsum ?f ?Pnknn" by simp
+        then have "a$n = (b$n - setsum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)"
+          apply -
+          apply (rule eq_divide_imp')
+          using r00
+          apply (simp del: of_nat_Suc)
+          by (simp add: mult_ac)
+        then have "a$n = ?r $n"
+          apply (simp del: of_nat_Suc)
+          unfolding fps_radical_def n1
+          by (simp add: field_simps n1 th00 del: of_nat_Suc)}
+        ultimately show "a$n = ?r $ n" by (cases n, auto)
       qed}
     then have "a = ?r" by (simp add: fps_eq_iff)}
   ultimately show ?thesis by blast
@@ -2105,16 +2105,16 @@
     proof(induct n rule: nat_less_induct)
       fix n assume h: "\<forall>m<n. ?i$m = X$m"
       {assume "n=0" hence "?i $n = X$n" using a0
-	  by (simp add: fps_compose_nth fps_inv_def)}
+          by (simp add: fps_compose_nth fps_inv_def)}
       moreover
       {fix n1 assume n1: "n = Suc n1"
-	have "?i $ n = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
-	  by (simp add: fps_compose_nth n1 startsby_zero_power_nth_same[OF a0]
+        have "?i $ n = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
+          by (simp add: fps_compose_nth n1 startsby_zero_power_nth_same[OF a0]
                    del: power_Suc)
-	also have "\<dots> = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + (X$ Suc n1 - setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
-	  using a0 a1 n1 by (simp add: fps_inv_def)
-	also have "\<dots> = X$n" using n1 by simp
-	finally have "?i $ n = X$n" .}
+        also have "\<dots> = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + (X$ Suc n1 - setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
+          using a0 a1 n1 by (simp add: fps_inv_def)
+        also have "\<dots> = X$n" using n1 by simp
+        finally have "?i $ n = X$n" .}
       ultimately show "?i $ n = X$n" by (cases n, auto)
     qed}
   then show ?thesis by (simp add: fps_eq_iff)
@@ -2136,16 +2136,16 @@
     proof(induct n rule: nat_less_induct)
       fix n assume h: "\<forall>m<n. ?i$m = b$m"
       {assume "n=0" hence "?i $n = b$n" using a0
-	  by (simp add: fps_compose_nth fps_ginv_def)}
+          by (simp add: fps_compose_nth fps_ginv_def)}
       moreover
       {fix n1 assume n1: "n = Suc n1"
-	have "?i $ n = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1"
-	  by (simp add: fps_compose_nth n1 startsby_zero_power_nth_same[OF a0]
+        have "?i $ n = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1"
+          by (simp add: fps_compose_nth n1 startsby_zero_power_nth_same[OF a0]
                    del: power_Suc)
-	also have "\<dots> = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + (b$ Suc n1 - setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
-	  using a0 a1 n1 by (simp add: fps_ginv_def)
-	also have "\<dots> = b$n" using n1 by simp
-	finally have "?i $ n = b$n" .}
+        also have "\<dots> = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + (b$ Suc n1 - setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
+          using a0 a1 n1 by (simp add: fps_ginv_def)
+        also have "\<dots> = b$n" using n1 by simp
+        finally have "?i $ n = b$n" .}
       ultimately show "?i $ n = b$n" by (cases n, auto)
     qed}
   then show ?thesis by (simp add: fps_eq_iff)
@@ -2182,7 +2182,7 @@
     next
       fix x F assume fF: "finite F" and xF: "x \<notin> F" and h: "setsum f F oo a = setsum (\<lambda>i. f i oo a) F"
       show "setsum f (insert x F) oo a  = setsum (\<lambda>i. f i oo a) (insert x F)"
-	using fF xF h by (simp add: fps_compose_add_distrib)
+        using fF xF h by (simp add: fps_compose_add_distrib)
     qed}
   ultimately show ?thesis by blast
 qed
@@ -2419,10 +2419,10 @@
   {fix h assume h: "k = Suc h"
     {fix n
       {assume kn: "k>n" hence "?l $ n = ?r $n" using a0 startsby_zero_power_prefix[OF a0] h
-	  by (simp add: fps_compose_nth del: power_Suc)}
+          by (simp add: fps_compose_nth del: power_Suc)}
       moreover
       {assume kn: "k \<le> n"
-	hence "?l$n = ?r$n"
+        hence "?l$n = ?r$n"
           by (simp add: fps_compose_nth mult_delta_left setsum_delta)}
       moreover have "k >n \<or> k\<le> n"  by arith
       ultimately have "?l$n = ?r$n"  by blast}
@@ -2751,19 +2751,19 @@
       
       from lrn 
       have "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" 
-	apply (simp add: ring_simps del: of_nat_Suc)
-	by (cases n, simp_all add: field_simps del: of_nat_Suc)
+        apply (simp add: ring_simps del: of_nat_Suc)
+        by (cases n, simp_all add: field_simps del: of_nat_Suc)
     }
     note th0 = this
     {fix n have "a$n = (c gchoose n) * a$0"
       proof(induct n)
-	case 0 thus ?case by simp
+        case 0 thus ?case by simp
       next
-	case (Suc m)
-	thus ?case unfolding th0
-	  apply (simp add: field_simps del: of_nat_Suc)
-	  unfolding mult_assoc[symmetric] gbinomial_mult_1
-	  by (simp add: ring_simps)
+        case (Suc m)
+        thus ?case unfolding th0
+          apply (simp add: field_simps del: of_nat_Suc)
+          unfolding mult_assoc[symmetric] gbinomial_mult_1
+          by (simp add: ring_simps)
       qed}
     note th1 = this
     have ?rhs
@@ -2855,11 +2855,11 @@
   {fix k assume kn: "k \<in> {0..n}"
     {assume c:"pochhammer (b - of_nat n + 1) n = 0"
       then obtain j where j: "j < n" "b - of_nat n + 1 = - of_nat j"
-	unfolding pochhammer_eq_0_iff by blast
+        unfolding pochhammer_eq_0_iff by blast
       from j have "b = of_nat n - of_nat j - of_nat 1" 
-	by (simp add: algebra_simps)
+        by (simp add: algebra_simps)
       then have "b = of_nat (n - j - 1)" 
-	using j kn by (simp add: of_nat_diff)
+        using j kn by (simp add: of_nat_diff)
       with b have False using j by auto}
     then have nz: "pochhammer (1 + b - of_nat n) n \<noteq> 0" 
       by (auto simp add: algebra_simps)
@@ -2868,89 +2868,89 @@
       by (simp add: pochhammer_neq_0_mono)
     {assume k0: "k = 0 \<or> n =0" 
       then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" 
-	using kn
-	by (cases "k=0", simp_all add: gbinomial_pochhammer)}
+        using kn
+        by (cases "k=0", simp_all add: gbinomial_pochhammer)}
     moreover
     {assume n0: "n \<noteq> 0" and k0: "k \<noteq> 0" 
       then obtain m where m: "n = Suc m" by (cases n, auto)
       from k0 obtain h where h: "k = Suc h" by (cases k, auto)
       {assume kn: "k = n"
-	then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
-	  using kn pochhammer_minus'[where k=k and n=n and b=b]
-	  apply (simp add:  pochhammer_same)
-	  using bn0
-	  by (simp add: field_simps power_add[symmetric])}
+        then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
+          using kn pochhammer_minus'[where k=k and n=n and b=b]
+          apply (simp add:  pochhammer_same)
+          using bn0
+          by (simp add: field_simps power_add[symmetric])}
       moreover
       {assume nk: "k \<noteq> n"
-	have m1nk: "?m1 n = setprod (%i. - 1) {0..m}" 
-	  "?m1 k = setprod (%i. - 1) {0..h}"
-	  by (simp_all add: setprod_constant m h)
-	from kn nk have kn': "k < n" by simp
-	have bnz0: "pochhammer (b - of_nat n + 1) k \<noteq> 0"
-	  using bn0 kn 
-	  unfolding pochhammer_eq_0_iff
-	  apply auto
-	  apply (erule_tac x= "n - ka - 1" in allE)
-	  by (auto simp add: algebra_simps of_nat_diff)
-	have eq1: "setprod (%k. (1::'a) + of_nat m - of_nat k) {0 .. h} = setprod of_nat {Suc (m - h) .. Suc m}"	
-	  apply (rule strong_setprod_reindex_cong[where f="%k. Suc m - k "])
-	  using kn' h m
-	  apply (auto simp add: inj_on_def image_def)
-	  apply (rule_tac x="Suc m - x" in bexI)
-	  apply (simp_all add: of_nat_diff)
-	  done
-	
-	have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))"
-	  unfolding m1nk 
-	  
-	  unfolding m h pochhammer_Suc_setprod
-	  apply (simp add: field_simps del: fact_Suc id_def)
-	  unfolding fact_altdef_nat id_def
-	  unfolding of_nat_setprod
-	  unfolding setprod_timesf[symmetric]
-	  apply auto
-	  unfolding eq1
-	  apply (subst setprod_Un_disjoint[symmetric])
-	  apply (auto)
-	  apply (rule setprod_cong)
-	  apply auto
-	  done
-	have th20: "?m1 n * ?p b n = setprod (%i. b - of_nat i) {0..m}"
-	  unfolding m1nk 
-	  unfolding m h pochhammer_Suc_setprod
-	  unfolding setprod_timesf[symmetric]
-	  apply (rule setprod_cong)
-	  apply auto
-	  done
-	have th21:"pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {n - k .. n - 1}"
-	  unfolding h m 
-	  unfolding pochhammer_Suc_setprod
-	  apply (rule strong_setprod_reindex_cong[where f="%k. n - 1 - k"])
-	  using kn
-	  apply (auto simp add: inj_on_def m h image_def)
-	  apply (rule_tac x= "m - x" in bexI)
-	  by (auto simp add: of_nat_diff)
-	
-	have "?m1 n * ?p b n = pochhammer (b - of_nat n + 1) k * setprod (%i. b - of_nat i) {0.. n - k - 1}"
-	  unfolding th20 th21
-	  unfolding h m
-	  apply (subst setprod_Un_disjoint[symmetric])
-	  using kn' h m
-	  apply auto
-	  apply (rule setprod_cong)
-	  apply auto
-	  done
-	then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {0.. n - k - 1}" 
-	  using nz' by (simp add: field_simps)
-	have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) = ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)"
-	  using bnz0
-	  by (simp add: field_simps)
-	also have "\<dots> = b gchoose (n - k)" 
-	  unfolding th1 th2
-	  using kn' by (simp add: gbinomial_def)
-	finally have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" by simp}
+        have m1nk: "?m1 n = setprod (%i. - 1) {0..m}" 
+          "?m1 k = setprod (%i. - 1) {0..h}"
+          by (simp_all add: setprod_constant m h)
+        from kn nk have kn': "k < n" by simp
+        have bnz0: "pochhammer (b - of_nat n + 1) k \<noteq> 0"
+          using bn0 kn 
+          unfolding pochhammer_eq_0_iff
+          apply auto
+          apply (erule_tac x= "n - ka - 1" in allE)
+          by (auto simp add: algebra_simps of_nat_diff)
+        have eq1: "setprod (%k. (1::'a) + of_nat m - of_nat k) {0 .. h} = setprod of_nat {Suc (m - h) .. Suc m}"        
+          apply (rule strong_setprod_reindex_cong[where f="%k. Suc m - k "])
+          using kn' h m
+          apply (auto simp add: inj_on_def image_def)
+          apply (rule_tac x="Suc m - x" in bexI)
+          apply (simp_all add: of_nat_diff)
+          done
+        
+        have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))"
+          unfolding m1nk 
+          
+          unfolding m h pochhammer_Suc_setprod
+          apply (simp add: field_simps del: fact_Suc id_def)
+          unfolding fact_altdef_nat id_def
+          unfolding of_nat_setprod
+          unfolding setprod_timesf[symmetric]
+          apply auto
+          unfolding eq1
+          apply (subst setprod_Un_disjoint[symmetric])
+          apply (auto)
+          apply (rule setprod_cong)
+          apply auto
+          done
+        have th20: "?m1 n * ?p b n = setprod (%i. b - of_nat i) {0..m}"
+          unfolding m1nk 
+          unfolding m h pochhammer_Suc_setprod
+          unfolding setprod_timesf[symmetric]
+          apply (rule setprod_cong)
+          apply auto
+          done
+        have th21:"pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {n - k .. n - 1}"
+          unfolding h m 
+          unfolding pochhammer_Suc_setprod
+          apply (rule strong_setprod_reindex_cong[where f="%k. n - 1 - k"])
+          using kn
+          apply (auto simp add: inj_on_def m h image_def)
+          apply (rule_tac x= "m - x" in bexI)
+          by (auto simp add: of_nat_diff)
+        
+        have "?m1 n * ?p b n = pochhammer (b - of_nat n + 1) k * setprod (%i. b - of_nat i) {0.. n - k - 1}"
+          unfolding th20 th21
+          unfolding h m
+          apply (subst setprod_Un_disjoint[symmetric])
+          using kn' h m
+          apply auto
+          apply (rule setprod_cong)
+          apply auto
+          done
+        then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {0.. n - k - 1}" 
+          using nz' by (simp add: field_simps)
+        have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) = ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)"
+          using bnz0
+          by (simp add: field_simps)
+        also have "\<dots> = b gchoose (n - k)" 
+          unfolding th1 th2
+          using kn' by (simp add: gbinomial_def)
+        finally have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" by simp}
       ultimately have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
-	by (cases "k =n", auto)}
+        by (cases "k =n", auto)}
     ultimately have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" "pochhammer (1 + b - of_nat n) k \<noteq> 0 "
       using nz' 
       apply (cases "n=0", auto)
@@ -3010,14 +3010,14 @@
     {assume en: "even n"
       have "?lhs$n = of_nat (n+1) * (fps_sin c $ (n+1))" by simp
       also have "\<dots> = of_nat (n+1) * ((- 1)^(n div 2) * c^Suc n / of_nat (fact (Suc n)))"
-	using en by (simp add: fps_sin_def)
+        using en by (simp add: fps_sin_def)
       also have "\<dots> = (- 1)^(n div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
-	unfolding fact_Suc of_nat_mult
-	by (simp add: field_simps del: of_nat_add of_nat_Suc)
+        unfolding fact_Suc of_nat_mult
+        by (simp add: field_simps del: of_nat_add of_nat_Suc)
       also have "\<dots> = (- 1)^(n div 2) *c^Suc n / of_nat (fact n)"
-	by (simp add: field_simps del: of_nat_add of_nat_Suc)
+        by (simp add: field_simps del: of_nat_add of_nat_Suc)
       finally have "?lhs $n = ?rhs$n" using en
-	by (simp add: fps_cos_def ring_simps power_Suc )}
+        by (simp add: fps_cos_def ring_simps power_Suc )}
     then show "?lhs $ n = ?rhs $ n"
       by (cases "even n", simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
 qed
@@ -3034,19 +3034,19 @@
       from en have n0: "n \<noteq>0 " by presburger
       have "?lhs$n = of_nat (n+1) * (fps_cos c $ (n+1))" by simp
       also have "\<dots> = of_nat (n+1) * ((- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact (Suc n)))"
-	using en by (simp add: fps_cos_def)
+        using en by (simp add: fps_cos_def)
       also have "\<dots> = (- 1)^((n + 1) div 2)*c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
-	unfolding fact_Suc of_nat_mult
-	by (simp add: field_simps del: of_nat_add of_nat_Suc)
+        unfolding fact_Suc of_nat_mult
+        by (simp add: field_simps del: of_nat_add of_nat_Suc)
       also have "\<dots> = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)"
-	by (simp add: field_simps del: of_nat_add of_nat_Suc)
+        by (simp add: field_simps del: of_nat_add of_nat_Suc)
       also have "\<dots> = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)"
-	unfolding th0 unfolding th1[OF en] by simp
+        unfolding th0 unfolding th1[OF en] by simp
       finally have "?lhs $n = ?rhs$n" using en
-	by (simp add: fps_sin_def ring_simps power_Suc)}
+        by (simp add: fps_sin_def ring_simps power_Suc)}
     then show "?lhs $ n = ?rhs $ n"
       by (cases "even n", simp_all add: fps_deriv_def fps_sin_def
-	fps_cos_def)
+        fps_cos_def)
 qed
 
 lemma fps_sin_cos_sum_of_squares:
@@ -3189,18 +3189,18 @@
   {fix n::nat
     {assume en: "even n"
       from en obtain m where m: "n = 2*m" 
-	unfolding even_mult_two_ex by blast
+        unfolding even_mult_two_ex by blast
       
       have "?l $n = ?r$n" 
-	by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
-	  power_mult power_minus)}
+        by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
+          power_mult power_minus)}
     moreover
     {assume on: "odd n"
       from on obtain m where m: "n = 2*m + 1" 
-	unfolding odd_nat_equiv_def2 by (auto simp add: nat_mult_2)  
+        unfolding odd_nat_equiv_def2 by (auto simp add: nat_mult_2)  
       have "?l $n = ?r$n" 
-	by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
-	  power_mult power_minus)}
+        by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
+          power_mult power_minus)}
     ultimately have "?l $n = ?r$n"  by blast}
   then show ?thesis by (simp add: fps_eq_iff)
 qed
--- a/src/HOL/Library/FrechetDeriv.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/FrechetDeriv.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -196,11 +196,11 @@
     proof (intro exI allI)
       fix x
       have "norm (f (g x)) \<le> norm (g x) * Kf"
-	using f .
+        using f .
       also have "\<dots> \<le> (norm x * Kg) * Kf"
-	using g Kf [THEN order_less_imp_le] by (rule mult_right_mono)
+        using g Kf [THEN order_less_imp_le] by (rule mult_right_mono)
       also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)"
-	by (rule mult_assoc)
+        by (rule mult_assoc)
       finally show "norm (f (g x)) \<le> norm x * (Kg * Kf)" .
     qed
   qed
--- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -20,11 +20,11 @@
   {assume y0: "y = 0"
     {assume x0: "x \<ge> 0"
       then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
-	by (simp add: csqrt_def power2_eq_square)}
+        by (simp add: csqrt_def power2_eq_square)}
     moreover
     {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
       then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
-	by (simp add: csqrt_def power2_eq_square) }
+        by (simp add: csqrt_def power2_eq_square) }
     ultimately have ?thesis by blast}
   moreover
   {assume y0: "y\<noteq>0"
@@ -322,7 +322,7 @@
       hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
       from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
       have "cmod (s (?h n) - ?w) < e"
-	using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
+        using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
     hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
   with hs show ?thesis  by blast
 qed
@@ -358,13 +358,13 @@
       by (simp_all add: field_simps real_mult_order)
     show ?case
       proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
-	fix d w
-	assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
-	hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
-	from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
-	from H have th: "cmod (w-z) \<le> d" by simp
-	from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
-	show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
+        fix d w
+        assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
+        hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
+        from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
+        from H have th: "cmod (w-z) \<le> d" by simp
+        from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
+        show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
       qed
     qed
 qed
@@ -397,14 +397,14 @@
     {fix y
       from s[rule_format, of "-y"] have
     "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y"
-	unfolding minus_less_iff[of y ] equation_minus_iff by blast }
+        unfolding minus_less_iff[of y ] equation_minus_iff by blast }
     note s1 = this[unfolded minus_minus]
     from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m"
       by auto
     {fix n::nat
       from s1[rule_format, of "?m + 1/real (Suc n)"]
       have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
-	by simp}
+        by simp}
     hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
     from choice[OF th] obtain g where
       g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)"
@@ -416,32 +416,32 @@
       assume wr: "cmod w \<le> r"
       let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
       {assume e: "?e > 0"
-	hence e2: "?e/2 > 0" by simp
-	from poly_cont[OF e2, of z p] obtain d where
-	  d: "d>0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2" by blast
-	{fix w assume w: "cmod (w - z) < d"
-	  have "cmod(poly p w - poly p z) < ?e / 2"
-	    using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
-	note th1 = this
+        hence e2: "?e/2 > 0" by simp
+        from poly_cont[OF e2, of z p] obtain d where
+          d: "d>0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2" by blast
+        {fix w assume w: "cmod (w - z) < d"
+          have "cmod(poly p w - poly p z) < ?e / 2"
+            using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
+        note th1 = this
 
-	from fz(2)[rule_format, OF d(1)] obtain N1 where
-	  N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
-	from reals_Archimedean2[of "2/?e"] obtain N2::nat where
-	  N2: "2/?e < real N2" by blast
-	have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
-	  using N1[rule_format, of "N1 + N2"] th1 by simp
-	{fix a b e2 m :: real
-	have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
+        from fz(2)[rule_format, OF d(1)] obtain N1 where
+          N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
+        from reals_Archimedean2[of "2/?e"] obtain N2::nat where
+          N2: "2/?e < real N2" by blast
+        have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
+          using N1[rule_format, of "N1 + N2"] th1 by simp
+        {fix a b e2 m :: real
+        have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
           ==> False" by arith}
       note th0 = this
       have ath:
-	"\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
+        "\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
       from s1m[OF g(1)[rule_format]]
       have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
       from seq_suble[OF fz(1), of "N1+N2"]
       have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
       have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"
-	using N2 by auto
+        using N2 by auto
       from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
       from g(2)[rule_format, of "f (N1 + N2)"]
       have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
@@ -453,10 +453,10 @@
       with ath[OF th31 th32]
       have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith
       have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c"
-	by arith
+        by arith
       have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
 \<le> cmod (poly p (g (f (N1 + N2))) - poly p z)"
-	by (simp add: norm_triangle_ineq3)
+        by (simp add: norm_triangle_ineq3)
       from ath2[OF th22, of ?m]
       have thc2: "2*(?e/2) \<le> \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)" by simp
       from th0[OF th2 thc1 thc2] have False .}
@@ -502,10 +502,10 @@
       from h have z1: "cmod z \<ge> 1" by arith
       from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (pCons c cs) z"]]]
       have th1: "d \<le> cmod(z * poly (pCons c cs) z) - cmod a"
-	unfolding norm_mult by (simp add: algebra_simps)
+        unfolding norm_mult by (simp add: algebra_simps)
       from complex_mod_triangle_sub[of "z * poly (pCons c cs) z" a]
       have th2: "cmod(z * poly (pCons c cs) z) - cmod a \<le> cmod (poly (pCons a (pCons c cs)) z)"
-	by (simp add: diff_le_eq algebra_simps)
+        by (simp add: diff_le_eq algebra_simps)
       from th1 th2 have "d \<le> cmod (poly (pCons a (pCons c cs)) z)"  by arith}
     hence ?case by blast}
   moreover
@@ -516,11 +516,11 @@
       assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
       from c0 have "cmod c > 0" by simp
       from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)"
-	by (simp add: field_simps norm_mult)
+        by (simp add: field_simps norm_mult)
       have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
       from complex_mod_triangle_sub[of "z*c" a ]
       have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
-	by (simp add: algebra_simps)
+        by (simp add: algebra_simps)
       from ath[OF th1 th0] have "d \<le> cmod (poly (pCons a (pCons c cs)) z)"
         using cs0' by simp}
     then have ?case  by blast}
@@ -541,7 +541,7 @@
     {fix z assume z: "r \<le> cmod z"
       from v[of 0] r[OF z]
       have "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) z)"
-	by simp }
+        by simp }
     note v0 = this
     from v0 v ath[of r] have ?case by blast}
   moreover
@@ -644,11 +644,11 @@
     {assume h: "constant (poly q)"
       from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
       {fix x y
-	from th have "?p x = poly q (x - c)" by auto
-	also have "\<dots> = poly q (y - c)"
-	  using h unfolding constant_def by blast
-	also have "\<dots> = ?p y" using th by auto
-	finally have "?p x = ?p y" .}
+        from th have "?p x = poly q (x - c)" by auto
+        also have "\<dots> = poly q (y - c)"
+          using h unfolding constant_def by blast
+        also have "\<dots> = ?p y" using th by auto
+        finally have "?p x = ?p y" .}
       with nc have False unfolding constant_def by blast }
     hence qnc: "\<not> constant (poly q)" by blast
     from q(2) have pqc0: "?p c = poly q 0" by simp
@@ -664,19 +664,19 @@
       by (simp add: expand_poly_eq)
     {assume h: "\<And>x y. poly ?r x = poly ?r y"
       {fix x y
-	from qr[rule_format, of x]
-	have "poly q x = poly ?r x * ?a0" by auto
-	also have "\<dots> = poly ?r y * ?a0" using h by simp
-	also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
-	finally have "poly q x = poly q y" .}
+        from qr[rule_format, of x]
+        have "poly q x = poly ?r x * ?a0" by auto
+        also have "\<dots> = poly ?r y * ?a0" using h by simp
+        also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
+        finally have "poly q x = poly q y" .}
       with qnc have False unfolding constant_def by blast}
     hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
     from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
     {fix w
       have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
-	using qr[rule_format, of w] a00 by (simp add: divide_inverse mult_ac)
+        using qr[rule_format, of w] a00 by (simp add: divide_inverse mult_ac)
       also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
-	using a00 unfolding norm_divide by (simp add: field_simps)
+        using a00 unfolding norm_divide by (simp add: field_simps)
       finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
     note mrmq_eq = this
     from poly_decompose[OF rnc] obtain k a s where
@@ -685,72 +685,72 @@
     {assume "k + 1 = n"
       with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=0" by auto
       {fix w
-	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)"
-	  using kas(4)[rule_format, of w] s0 r01 by (simp add: algebra_simps)}
+        have "cmod (poly ?r w) = cmod (1 + a * w ^ k)"
+          using kas(4)[rule_format, of w] s0 r01 by (simp add: algebra_simps)}
       note hth = this [symmetric]
-	from reduce_poly_simple[OF kas(1,2)]
+        from reduce_poly_simple[OF kas(1,2)]
       have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
     moreover
     {assume kn: "k+1 \<noteq> n"
       from kn kas(3) q(1) n[symmetric] lgqr have k1n: "k + 1 < n" by simp
       have th01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))"
-	unfolding constant_def poly_pCons poly_monom
-	using kas(1) apply simp
-	by (rule exI[where x=0], rule exI[where x=1], simp)
+        unfolding constant_def poly_pCons poly_monom
+        using kas(1) apply simp
+        by (rule exI[where x=0], rule exI[where x=1], simp)
       from kas(1) kas(2) have th02: "k+1 = psize (pCons 1 (monom a (k - 1)))"
-	by (simp add: psize_def degree_monom_eq)
+        by (simp add: psize_def degree_monom_eq)
       from H[rule_format, OF k1n th01 th02]
       obtain w where w: "1 + w^k * a = 0"
-	unfolding poly_pCons poly_monom
-	using kas(2) by (cases k, auto simp add: algebra_simps)
+        unfolding poly_pCons poly_monom
+        using kas(2) by (cases k, auto simp add: algebra_simps)
       from poly_bound_exists[of "cmod w" s] obtain m where
-	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
+        m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
       have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
       from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
       then have wm1: "w^k * a = - 1" by simp
       have inv0: "0 < inverse (cmod w ^ (k + 1) * m)"
-	using norm_ge_zero[of w] w0 m(1)
-	  by (simp add: inverse_eq_divide zero_less_mult_iff)
+        using norm_ge_zero[of w] w0 m(1)
+          by (simp add: inverse_eq_divide zero_less_mult_iff)
       with real_down2[OF zero_less_one] obtain t where
-	t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
+        t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
       let ?ct = "complex_of_real t"
       let ?w = "?ct * w"
       have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" using kas(1) by (simp add: algebra_simps power_mult_distrib)
       also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
-	unfolding wm1 by (simp)
+        unfolding wm1 by (simp)
       finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)"
-	apply -
-	apply (rule cong[OF refl[of cmod]])
-	apply assumption
-	done
+        apply -
+        apply (rule cong[OF refl[of cmod]])
+        apply assumption
+        done
       with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"]
       have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)" unfolding norm_of_real by simp
       have ath: "\<And>x (t::real). 0\<le> x \<Longrightarrow> x < t \<Longrightarrow> t\<le>1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1" by arith
       have "t *cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
       then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult)
       from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
-	by (simp add: inverse_eq_divide field_simps)
+        by (simp add: inverse_eq_divide field_simps)
       with zero_less_power[OF t(1), of k]
       have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1"
-	apply - apply (rule mult_strict_left_mono) by simp_all
+        apply - apply (rule mult_strict_left_mono) by simp_all
       have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
-	by (simp add: algebra_simps power_mult_distrib norm_of_real norm_power norm_mult)
+        by (simp add: algebra_simps power_mult_distrib norm_of_real norm_power norm_mult)
       then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
-	using t(1,2) m(2)[rule_format, OF tw] w0
-	apply (simp only: )
-	apply auto
-	apply (rule mult_mono, simp_all add: norm_ge_zero)+
-	apply (simp add: zero_le_mult_iff zero_le_power)
-	done
+        using t(1,2) m(2)[rule_format, OF tw] w0
+        apply (simp only: )
+        apply auto
+        apply (rule mult_mono, simp_all add: norm_ge_zero)+
+        apply (simp add: zero_le_mult_iff zero_le_power)
+        done
       with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp
       from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1"
-	by auto
+        by auto
       from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
       have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" .
       from th11 th12
       have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith
       then have "cmod (poly ?r ?w) < 1"
-	unfolding kas(4)[rule_format, of ?w] r01 by simp
+        unfolding kas(4)[rule_format, of ?w] r01 by simp
       then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
     ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
     from cr0_contr cq0 q(2)
@@ -773,30 +773,30 @@
       from nc[unfolded constant_def, rule_format, of 0]
       have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto
       hence "cs = 0"
-	proof(induct cs)
-	  case (pCons d ds)
-	  {assume "d=0" hence ?case using pCons.prems pCons.hyps by simp}
-	  moreover
-	  {assume d0: "d\<noteq>0"
-	    from poly_bound_exists[of 1 ds] obtain m where
-	      m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
-	    have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
-	    from real_down2[OF dm zero_less_one] obtain x where
-	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
-	    let ?x = "complex_of_real x"
-	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
-	    from pCons.prems[rule_format, OF cx(1)]
-	    have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
-	    from m(2)[rule_format, OF cx(2)] x(1)
-	    have th0: "cmod (?x*poly ds ?x) \<le> x*m"
-	      by (simp add: norm_mult)
-	    from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
-	    with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
-	    with cth  have ?case by blast}
-	  ultimately show ?case by blast
-	qed simp}
+        proof(induct cs)
+          case (pCons d ds)
+          {assume "d=0" hence ?case using pCons.prems pCons.hyps by simp}
+          moreover
+          {assume d0: "d\<noteq>0"
+            from poly_bound_exists[of 1 ds] obtain m where
+              m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
+            have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
+            from real_down2[OF dm zero_less_one] obtain x where
+              x: "x > 0" "x < cmod d / m" "x < 1" by blast
+            let ?x = "complex_of_real x"
+            from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
+            from pCons.prems[rule_format, OF cx(1)]
+            have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
+            from m(2)[rule_format, OF cx(2)] x(1)
+            have th0: "cmod (?x*poly ds ?x) \<le> x*m"
+              by (simp add: norm_mult)
+            from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
+            with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
+            with cth  have ?case by blast}
+          ultimately show ?case by blast
+        qed simp}
       then have nc: "\<not> constant (poly (pCons c cs))" using pCons.prems c0
-	by blast
+        by blast
       from fundamental_theorem_of_algebra[OF nc] have ?case .}
   ultimately show ?case by blast
 qed simp
@@ -823,59 +823,59 @@
     {assume oa: "order a p \<noteq> 0"
       let ?op = "order a p"
       from pne have ap: "([:- a, 1:] ^ ?op) dvd p"
-	"\<not> [:- a, 1:] ^ (Suc ?op) dvd p" using order by blast+
+        "\<not> [:- a, 1:] ^ (Suc ?op) dvd p" using order by blast+
       note oop = order_degree[OF pne, unfolded dpn]
       {assume q0: "q = 0"
-	hence ?ths using n0
+        hence ?ths using n0
           by (simp add: power_0_left)}
       moreover
       {assume q0: "q \<noteq> 0"
-	from pq0[rule_format, OF a, unfolded poly_eq_0_iff_dvd]
-	obtain r where r: "q = [:- a, 1:] * r" by (rule dvdE)
-	from ap(1) obtain s where
-	  s: "p = [:- a, 1:] ^ ?op * s" by (rule dvdE)
-	have sne: "s \<noteq> 0"
-	  using s pne by auto
-	{assume ds0: "degree s = 0"
-	  from ds0 have "\<exists>k. s = [:k:]"
+        from pq0[rule_format, OF a, unfolded poly_eq_0_iff_dvd]
+        obtain r where r: "q = [:- a, 1:] * r" by (rule dvdE)
+        from ap(1) obtain s where
+          s: "p = [:- a, 1:] ^ ?op * s" by (rule dvdE)
+        have sne: "s \<noteq> 0"
+          using s pne by auto
+        {assume ds0: "degree s = 0"
+          from ds0 have "\<exists>k. s = [:k:]"
             by (cases s, simp split: if_splits)
-	  then obtain k where kpn: "s = [:k:]" by blast
+          then obtain k where kpn: "s = [:k:]" by blast
           from sne kpn have k: "k \<noteq> 0" by simp
-	  let ?w = "([:1/k:] * ([:-a,1:] ^ (n - ?op))) * (r ^ n)"
+          let ?w = "([:1/k:] * ([:-a,1:] ^ (n - ?op))) * (r ^ n)"
           from k oop [of a] have "q ^ n = p * ?w"
             apply -
             apply (subst r, subst s, subst kpn)
             apply (subst power_mult_distrib, simp)
             apply (subst power_add [symmetric], simp)
             done
-	  hence ?ths unfolding dvd_def by blast}
-	moreover
-	{assume ds0: "degree s \<noteq> 0"
-	  from ds0 sne dpn s oa
-	    have dsn: "degree s < n" apply auto
+          hence ?ths unfolding dvd_def by blast}
+        moreover
+        {assume ds0: "degree s \<noteq> 0"
+          from ds0 sne dpn s oa
+            have dsn: "degree s < n" apply auto
               apply (erule ssubst)
               apply (simp add: degree_mult_eq degree_linear_power)
               done
-	    {fix x assume h: "poly s x = 0"
-	      {assume xa: "x = a"
-		from h[unfolded xa poly_eq_0_iff_dvd] obtain u where
-		  u: "s = [:- a, 1:] * u" by (rule dvdE)
-		have "p = [:- a, 1:] ^ (Suc ?op) * u"
+            {fix x assume h: "poly s x = 0"
+              {assume xa: "x = a"
+                from h[unfolded xa poly_eq_0_iff_dvd] obtain u where
+                  u: "s = [:- a, 1:] * u" by (rule dvdE)
+                have "p = [:- a, 1:] ^ (Suc ?op) * u"
                   by (subst s, subst u, simp only: power_Suc mult_ac)
-		with ap(2)[unfolded dvd_def] have False by blast}
-	      note xa = this
-	      from h have "poly p x = 0" by (subst s, simp)
-	      with pq0 have "poly q x = 0" by blast
-	      with r xa have "poly r x = 0"
+                with ap(2)[unfolded dvd_def] have False by blast}
+              note xa = this
+              from h have "poly p x = 0" by (subst s, simp)
+              with pq0 have "poly q x = 0" by blast
+              with r xa have "poly r x = 0"
                 by (auto simp add: uminus_add_conv_diff)}
-	    note impth = this
-	    from IH[rule_format, OF dsn, of s r] impth ds0
-	    have "s dvd (r ^ (degree s))" by blast
-	    then obtain u where u: "r ^ (degree s) = s * u" ..
-	    hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
+            note impth = this
+            from IH[rule_format, OF dsn, of s r] impth ds0
+            have "s dvd (r ^ (degree s))" by blast
+            then obtain u where u: "r ^ (degree s) = s * u" ..
+            hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
               by (simp only: poly_mult[symmetric] poly_power[symmetric])
-	    let ?w = "(u * ([:-a,1:] ^ (n - ?op))) * (r ^ (n - degree s))"
-	    from oop[of a] dsn have "q ^ n = p * ?w"
+            let ?w = "(u * ([:-a,1:] ^ (n - ?op))) * (r ^ (n - degree s))"
+            from oop[of a] dsn have "q ^ n = p * ?w"
               apply -
               apply (subst s, subst r)
               apply (simp only: power_mult_distrib)
@@ -885,7 +885,7 @@
               apply (subst u [symmetric])
               apply (simp add: mult_ac power_add [symmetric])
               done
-	    hence ?ths unfolding dvd_def by blast}
+            hence ?ths unfolding dvd_def by blast}
       ultimately have ?ths by blast }
       ultimately have ?ths by blast}
     then have ?ths using a order_root pne by blast}
@@ -930,12 +930,12 @@
     {assume dp: "degree p \<noteq> 0"
       then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
       {assume "p dvd (q ^ (Suc n))"
-	then obtain u where u: "q ^ (Suc n) = p * u" ..
-	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
-	  hence "poly (q ^ (Suc n)) x \<noteq> 0" by simp
-	  hence False using u h(1) by (simp only: poly_mult) simp}}
-	with n nullstellensatz_lemma[of p q "degree p"] dp
-	have ?thesis by auto}
+        then obtain u where u: "q ^ (Suc n) = p * u" ..
+        {fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
+          hence "poly (q ^ (Suc n)) x \<noteq> 0" by simp
+          hence False using u h(1) by (simp only: poly_mult) simp}}
+        with n nullstellensatz_lemma[of p q "degree p"] dp
+        have ?thesis by auto}
     ultimately have ?thesis by blast}
   ultimately show ?thesis by blast
 qed
--- a/src/HOL/Library/Multiset.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Multiset.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -926,19 +926,19 @@
       case (Cons y ys')
       have x_in_ys: "x = y \<or> x \<in> set ys'"
       proof (cases "x = y")
-	case True then show ?thesis ..
+        case True then show ?thesis ..
       next
-	case False
-	from "1.prems" [symmetric] xCons Cons have "x :# multiset_of ys' + {#y#}" by simp
-	with False show ?thesis by (simp add: mem_set_multiset_eq)
+        case False
+        from "1.prems" [symmetric] xCons Cons have "x :# multiset_of ys' + {#y#}" by simp
+        with False show ?thesis by (simp add: mem_set_multiset_eq)
       qed
       from "1.hyps" have IH: "length xs' < length xs \<longrightarrow>
-	(\<forall>x. multiset_of xs' = multiset_of x \<longrightarrow> length xs' = length x)" by blast
+        (\<forall>x. multiset_of xs' = multiset_of x \<longrightarrow> length xs' = length x)" by blast
       from "1.prems" x_in_ys Cons xCons have "multiset_of xs' = multiset_of (remove1 x (y#ys'))"
-	apply -
-	apply (simp add: multiset_of_remove1, simp only: add_eq_conv_diff)
-	apply fastsimp
-	done
+        apply -
+        apply (simp add: multiset_of_remove1, simp only: add_eq_conv_diff)
+        apply fastsimp
+        done
       with IH xCons have IH': "length xs' = length (remove1 x (y#ys'))" by fastsimp
       from x_in_ys have "x \<noteq> y \<Longrightarrow> length ys' > 0" by auto
       with Cons xCons x_in_ys IH' show ?thesis by (auto simp add: length_remove1)
--- a/src/HOL/Library/Nat_Int_Bij.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Nat_Int_Bij.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -120,7 +120,7 @@
     have "2*z=2*(\<Sum>i\<le>r. i)+2*x"  using x_def a by simp arith
     also have "\<dots> = r * (r+1) + 2*x"   using gauss_sum_nat_upto by simp
     also have "\<dots> = (x+y)*(x+y+1)+2*x" using y_def b by simp
-    also { have "2 dvd ((x+y)*(x+y+1))"	using dvd2_a_x_suc_a by simp }
+    also { have "2 dvd ((x+y)*(x+y+1))" using dvd2_a_x_suc_a by simp }
     hence "\<dots> = 2 * nat2_to_nat(x,y)"
       using nat2_to_nat_def by (simp add: Let_def dvd_mult_div_cancel)
     finally have "z=nat2_to_nat (x, y)"  by simp
--- a/src/HOL/Library/Permutations.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Permutations.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -181,42 +181,42 @@
       from hassize_insert[OF xF H0] have Fs: "F hassize (n - 1)" .
       from H Fs have pFs: "?pF hassize fact (n - 1)" by blast
       hence pF'f: "finite ?pF'" using H0 unfolding hassize_def
-	apply (simp only: Collect_split Collect_mem_eq)
-	apply (rule finite_cartesian_product)
-	apply simp_all
-	done
+        apply (simp only: Collect_split Collect_mem_eq)
+        apply (rule finite_cartesian_product)
+        apply simp_all
+        done
 
       have ginj: "inj_on ?g ?pF'"
       proof-
-	{
-	  fix b p c q assume bp: "(b,p) \<in> ?pF'" and cq: "(c,q) \<in> ?pF'"
-	    and eq: "?g (b,p) = ?g (c,q)"
-	  from bp cq have ths: "b \<in> insert x F" "c \<in> insert x F" "x \<in> insert x F" "p permutes F" "q permutes F" by auto
-	  from ths(4) xF eq have "b = ?g (b,p) x" unfolding permutes_def
-	    by (auto simp add: swap_def fun_upd_def expand_fun_eq)
-	  also have "\<dots> = ?g (c,q) x" using ths(5) xF eq
-	    by (auto simp add: swap_def fun_upd_def expand_fun_eq)
-	  also have "\<dots> = c"using ths(5) xF unfolding permutes_def
-	    by (auto simp add: swap_def fun_upd_def expand_fun_eq)
-	  finally have bc: "b = c" .
-	  hence "Fun.swap x b id = Fun.swap x c id" by simp
-	  with eq have "Fun.swap x b id o p = Fun.swap x b id o q" by simp
-	  hence "Fun.swap x b id o (Fun.swap x b id o p) = Fun.swap x b id o (Fun.swap x b id o q)" by simp
-	  hence "p = q" by (simp add: o_assoc)
-	  with bc have "(b,p) = (c,q)" by simp }
-	thus ?thesis  unfolding inj_on_def by blast
+        {
+          fix b p c q assume bp: "(b,p) \<in> ?pF'" and cq: "(c,q) \<in> ?pF'"
+            and eq: "?g (b,p) = ?g (c,q)"
+          from bp cq have ths: "b \<in> insert x F" "c \<in> insert x F" "x \<in> insert x F" "p permutes F" "q permutes F" by auto
+          from ths(4) xF eq have "b = ?g (b,p) x" unfolding permutes_def
+            by (auto simp add: swap_def fun_upd_def expand_fun_eq)
+          also have "\<dots> = ?g (c,q) x" using ths(5) xF eq
+            by (auto simp add: swap_def fun_upd_def expand_fun_eq)
+          also have "\<dots> = c"using ths(5) xF unfolding permutes_def
+            by (auto simp add: swap_def fun_upd_def expand_fun_eq)
+          finally have bc: "b = c" .
+          hence "Fun.swap x b id = Fun.swap x c id" by simp
+          with eq have "Fun.swap x b id o p = Fun.swap x b id o q" by simp
+          hence "Fun.swap x b id o (Fun.swap x b id o p) = Fun.swap x b id o (Fun.swap x b id o q)" by simp
+          hence "p = q" by (simp add: o_assoc)
+          with bc have "(b,p) = (c,q)" by simp }
+        thus ?thesis  unfolding inj_on_def by blast
       qed
       from xF H0 have n0: "n \<noteq> 0 " by (auto simp add: hassize_def)
       hence "\<exists>m. n = Suc m" by arith
       then obtain m where n[simp]: "n = Suc m" by blast
       from pFs H0 have xFc: "card ?xF = fact n"
-	unfolding xfgpF' card_image[OF ginj] hassize_def
-	apply (simp only: Collect_split Collect_mem_eq card_cartesian_product)
-	by simp
+        unfolding xfgpF' card_image[OF ginj] hassize_def
+        apply (simp only: Collect_split Collect_mem_eq card_cartesian_product)
+        by simp
       from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" unfolding xfgpF' by simp
       have "?xF hassize fact n"
-	using xFf xFc
-	unfolding hassize_def  xFf by blast }
+        using xFf xFc
+        unfolding hassize_def  xFf by blast }
     thus "\<forall>n. (insert x F hassize n) \<longrightarrow> ({p. p permutes insert x F} hassize fact n)"
       by blast
   qed
@@ -689,17 +689,17 @@
       using p le
     proof(induct n arbitrary: S rule: less_induct)
       fix n S assume H: "\<And>m S. \<lbrakk>m < n; p permutes S; \<forall>i\<in>S. p i \<le> i\<rbrakk> \<Longrightarrow> p m = m"
-	"p permutes S" "\<forall>i \<in>S. p i \<le> i"
+        "p permutes S" "\<forall>i \<in>S. p i \<le> i"
       {assume "n \<notin> S"
-	with H(2) have "p n = n" unfolding permutes_def by metis}
+        with H(2) have "p n = n" unfolding permutes_def by metis}
       moreover
       {assume ns: "n \<in> S"
-	from H(3)  ns have "p n < n \<or> p n = n" by auto
-	moreover{assume h: "p n < n"
-	  from H h have "p (p n) = p n" by metis
-	  with permutes_inj[OF H(2)] have "p n = n" unfolding inj_on_def by blast
-	  with h have False by simp}
-	ultimately have "p n = n" by blast }
+        from H(3)  ns have "p n < n \<or> p n = n" by auto
+        moreover{assume h: "p n < n"
+          from H h have "p (p n) = p n" by metis
+          with permutes_inj[OF H(2)] have "p n = n" unfolding inj_on_def by blast
+          with h have False by simp}
+        ultimately have "p n = n" by blast }
       ultimately show "p n = n"  by blast
     qed}
   thus ?thesis by (auto simp add: expand_fun_eq)
@@ -840,13 +840,13 @@
       and p: "p permutes S" and q: "q permutes S"
       and eq: "Fun.swap a b id o p = Fun.swap a c id o q"
       from p q aS have pa: "p a = a" and qa: "q a = a"
-	unfolding permutes_def by metis+
+        unfolding permutes_def by metis+
       from eq have "(Fun.swap a b id o p) a  = (Fun.swap a c id o q) a" by simp
       hence bc: "b = c"
-	by (simp add: permutes_def pa qa o_def fun_upd_def swap_def id_def cong del: if_weak_cong split: split_if_asm)
+        by (simp add: permutes_def pa qa o_def fun_upd_def swap_def id_def cong del: if_weak_cong split: split_if_asm)
       from eq[unfolded bc] have "(\<lambda>p. Fun.swap a c id o p) (Fun.swap a c id o p) = (\<lambda>p. Fun.swap a c id o p) (Fun.swap a c id o q)" by simp
       hence "p = q" unfolding o_assoc swap_id_idempotent
-	by (simp add: o_def)
+        by (simp add: o_def)
       with bc have "b = c \<and> p = q" by blast
     }
 
--- a/src/HOL/Library/Ramsey.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Ramsey.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,5 @@
 (*  Title:      HOL/Library/Ramsey.thy
-    Author:     Tom Ridge. Converted to structured Isar by L C Paulson
+    Author:     Tom Ridge.  Converted to structured Isar by L C Paulson
 *)
 
 header "Ramsey's Theorem"
@@ -91,8 +91,8 @@
     from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \<in> YY" by blast
     let ?ramr = "{((y,Y,t),(y',Y',t')). y' \<in> Y & Y' \<subseteq> Y}"
     let ?propr = "%(y,Y,t).     
-		 y \<in> YY & y \<notin> Y & Y \<subseteq> YY & infinite Y & t < s
-		 & (\<forall>X. X\<subseteq>Y & finite X & card X = r --> (f o insert y) X = t)"
+                 y \<in> YY & y \<notin> Y & Y \<subseteq> YY & infinite Y & t < s
+                 & (\<forall>X. X\<subseteq>Y & finite X & card X = r --> (f o insert y) X = t)"
     have infYY': "infinite (YY-{yy})" using Suc.prems by auto
     have partf': "part r s (YY - {yy}) (f \<circ> insert yy)"
       by (simp add: o_def part_Suc_imp_part yy Suc.prems)
@@ -115,16 +115,16 @@
         with fields px yx' Suc.prems
         have partfx': "part r s (Yx - {yx'}) (f \<circ> insert yx')"
           by (simp add: o_def part_Suc_imp_part part_subset [where ?YY=YY]) 
-	from Suc.hyps [OF infYx' partfx']
-	obtain Y' and t'
-	where Y': "Y' \<subseteq> Yx - {yx'}"  "infinite Y'"  "t' < s"
-	       "\<forall>X. X\<subseteq>Y' \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yx') X = t'"
-	    by blast 
-	show ?thesis
-	proof
-	  show "?propr (yx',Y',t') & (x, (yx',Y',t')) \<in> ?ramr"
-  	    using fields Y' yx' px by blast
-	qed
+        from Suc.hyps [OF infYx' partfx']
+        obtain Y' and t'
+        where Y': "Y' \<subseteq> Yx - {yx'}"  "infinite Y'"  "t' < s"
+               "\<forall>X. X\<subseteq>Y' \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yx') X = t'"
+            by blast 
+        show ?thesis
+        proof
+          show "?propr (yx',Y',t') & (x, (yx',Y',t')) \<in> ?ramr"
+            using fields Y' yx' px by blast
+        qed
       qed
     qed
     from dependent_choice [OF transr propr0 proprstep]
@@ -191,7 +191,7 @@
                by (simp add: cardX ya)
              ultimately show ?thesis 
                using pg [of "LEAST x. x \<in> AA"] fields cardX
-	       by (clarsimp simp del:insert_Diff_single)
+               by (clarsimp simp del:insert_Diff_single)
            qed
            also have "... = s'" using AA AAleast fields by auto
            finally show ?thesis .
--- a/src/HOL/Library/Topology_Euclidean_Space.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Topology_Euclidean_Space.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,6 @@
-(* Title:      Topology
-   Author:     Amine Chaieb, University of Cambridge
-   Author:     Robert Himmelmann, TU Muenchen
+(*  Title:      HOL/Library/Topology_Euclidian_Space.thy
+    Author:     Amine Chaieb, University of Cambridge
+    Author:     Robert Himmelmann, TU Muenchen
 *)
 
 header {* Elementary topology in Euclidean space. *}
@@ -179,7 +179,7 @@
     moreover
     {assume S: "openin U S"
       hence "\<exists>T. openin U T \<and> S = T \<inter> V"
-	using openin_subset[OF S] UV by auto}
+        using openin_subset[OF S] UV by auto}
     ultimately have "(\<exists>T. openin U T \<and> S = T \<inter> V) \<longleftrightarrow> openin U S" by blast}
   then show ?thesis unfolding topology_eq openin_subtopology by blast
 qed
@@ -330,7 +330,7 @@
     have oT: "open ?T" by auto
     { fix x assume "x\<in>S"
       hence "x \<in> \<Union>{B. \<exists>x\<in>S. B = ball x (d x)}"
-	apply simp apply(rule_tac x="ball x(d x)" in exI) apply auto
+        apply simp apply(rule_tac x="ball x(d x)" in exI) apply auto
         by (rule d [THEN conjunct1])
       hence "x\<in> ?T \<inter> U" using SU and `x\<in>S` by auto  }
     moreover
@@ -404,7 +404,7 @@
     unfolding connected_def openin_open closedin_closed by auto
   {fix e2
     {fix e1 have "?P e2 e1 \<longleftrightarrow> (\<exists>t.  closed e2 \<and> t = S\<inter>e2 \<and> open e1 \<and> t = S\<inter>e1 \<and> t\<noteq>{} \<and> t\<noteq>S)"
-	by auto}
+        by auto}
     then have "(\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by metis}
   then have "\<forall>e2. (\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by blast
   then show ?thesis unfolding th0 th1 by simp
@@ -570,10 +570,10 @@
       have th:" \<And>b a (x::real). abs x <= b \<Longrightarrow> b <= a ==> ~(a + x < 0)" by arith
       have th': "\<And>x (y::real). x < 0 \<Longrightarrow> 0 <= y ==> abs x <= abs (y - x)" by arith
       have th1: "\<bar>x$i\<bar> \<le> \<bar>(x' - x)$i\<bar>" using x'(1) xi
-	apply (simp only: vector_component)
-	by (rule th') auto
+        apply (simp only: vector_component)
+        by (rule th') auto
       have th2: "\<bar>dist x x'\<bar> \<ge> \<bar>(x' - x)$i\<bar>" using  component_le_norm[of "x'-x" i]
-	apply (simp add: dist_norm) by norm
+        apply (simp add: dist_norm) by norm
       from th[OF th1 th2] x'(3) have False by (simp add: dist_commute) }
   then show ?thesis unfolding closed_limpt islimpt_approachable
     unfolding not_le[symmetric] by blast
@@ -722,19 +722,19 @@
       let ?exT = "\<lambda> y. (\<exists>T. open T \<and> y \<in> T \<and> T \<subseteq> UNIV - S)"
       assume "?lhs"
       hence *:"\<not> ?exT x"
-	unfolding interior_def
-	by simp
+        unfolding interior_def
+        by simp
       { assume "\<not> ?rhs"
-	hence False using *
-	  unfolding closure_def islimpt_def
-	  by blast
+        hence False using *
+          unfolding closure_def islimpt_def
+          by blast
       }
       thus "?rhs"
-	by blast
+        by blast
     next
       assume "?rhs" thus "?lhs"
-	unfolding closure_def interior_def islimpt_def
-	by blast
+        unfolding closure_def interior_def islimpt_def
+        by blast
     qed
   }
   thus ?thesis
@@ -773,8 +773,8 @@
     { fix x
       assume "x islimpt S"
       hence "x islimpt t" using *(1)
-	using islimpt_subset[of x, of S, of t]
-	by blast
+        using islimpt_subset[of x, of S, of t]
+        by blast
     }
     with * have "closure S \<subseteq> t"
       unfolding closure_def
@@ -902,16 +902,16 @@
     { assume "a\<in>S"
       have "\<exists>x\<in>S. dist a x < e" using `e>0` `a\<in>S` by(rule_tac x=a in bexI) auto
       moreover have "\<exists>x. x \<notin> S \<and> dist a x < e" using `?lhs` `a\<in>S`
-	unfolding frontier_closures closure_def islimpt_def using `e>0`
-	by (auto, erule_tac x="ball a e" in allE, auto)
+        unfolding frontier_closures closure_def islimpt_def using `e>0`
+        by (auto, erule_tac x="ball a e" in allE, auto)
       ultimately have ?rhse by auto
     }
     moreover
     { assume "a\<notin>S"
       hence ?rhse using `?lhs`
-	unfolding frontier_closures closure_def islimpt_def
-	using open_ball[of a e] `e > 0`
-	by (auto, erule_tac x = "ball a e" in allE, auto) (* FIXME: VERY slow! *)
+        unfolding frontier_closures closure_def islimpt_def
+        using open_ball[of a e] `e > 0`
+        by (auto, erule_tac x = "ball a e" in allE, auto) (* FIXME: VERY slow! *)
     }
     ultimately have ?rhse by auto
   }
@@ -1476,7 +1476,7 @@
       unfolding Limits.eventually_at by fast
     { fix x::"'a" assume "x \<noteq> a \<and> dist x a < d"
       hence "f x \<in> S" using d apply(erule_tac x="x-a" in allE)
-	by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute)
+        by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute)
     }
     hence "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<in> S" using d(1) by auto
     hence "eventually (\<lambda>x. f x \<in> S) (at a)" unfolding Limits.eventually_at .
@@ -1727,49 +1727,49 @@
     proof(cases "d \<le> dist x y")
       case True thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
       proof(cases "x=y")
-	case True hence False using `d \<le> dist x y` `d>0` by auto
-	thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by auto
+        case True hence False using `d \<le> dist x y` `d>0` by auto
+        thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by auto
       next
-	case False
-
-	have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x))
-	      = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
-	  unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[THEN sym] by auto
-	also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
-	  using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", THEN sym, of "y - x"]
-	  unfolding scaleR_minus_left scaleR_one
-	  by (auto simp add: norm_minus_commute)
-	also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
-	  unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]]
-	  unfolding real_add_mult_distrib using `x\<noteq>y`[unfolded dist_nz, unfolded dist_norm] by auto
-	also have "\<dots> \<le> e - d/2" using `d \<le> dist x y` and `d>0` and `?rhs` by(auto simp add: dist_norm)
-	finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using `d>0` by auto
-
-	moreover
-
-	have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
-	  using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff by (auto simp add: dist_commute)
-	moreover
-	have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" unfolding dist_norm apply simp unfolding norm_minus_cancel
-	  using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
-	  unfolding dist_norm by auto
-	ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by (rule_tac  x="y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto
+        case False
+
+        have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x))
+              = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
+          unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[THEN sym] by auto
+        also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
+          using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", THEN sym, of "y - x"]
+          unfolding scaleR_minus_left scaleR_one
+          by (auto simp add: norm_minus_commute)
+        also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
+          unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]]
+          unfolding real_add_mult_distrib using `x\<noteq>y`[unfolded dist_nz, unfolded dist_norm] by auto
+        also have "\<dots> \<le> e - d/2" using `d \<le> dist x y` and `d>0` and `?rhs` by(auto simp add: dist_norm)
+        finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using `d>0` by auto
+
+        moreover
+
+        have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
+          using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff by (auto simp add: dist_commute)
+        moreover
+        have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" unfolding dist_norm apply simp unfolding norm_minus_cancel
+          using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
+          unfolding dist_norm by auto
+        ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by (rule_tac  x="y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto
       qed
     next
       case False hence "d > dist x y" by auto
       show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
       proof(cases "x=y")
-	case True
-	obtain z where **: "z \<noteq> y" "dist z y < min e d"
+        case True
+        obtain z where **: "z \<noteq> y" "dist z y < min e d"
           using perfect_choose_dist[of "min e d" y]
-	  using `d > 0` `e>0` by auto
-	show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+          using `d > 0` `e>0` by auto
+        show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
           unfolding `x = y`
           using `z \<noteq> y` **
           by (rule_tac x=z in bexI, auto simp add: dist_commute)
       next
-	case False thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
-	  using `d>0` `d > dist x y` `?rhs` by(rule_tac x=x in bexI, auto)
+        case False thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+          using `d>0` `d > dist x y` `?rhs` by(rule_tac x=x in bexI, auto)
       qed
     qed  }
   thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto
@@ -1856,11 +1856,11 @@
     next
       case False
       have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" unfolding dist_norm
-	using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
+        using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
       hence *:"y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e" using d as(1)[unfolded subset_eq] by blast
       have "y - x \<noteq> 0" using `x \<noteq> y` by auto
       hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym]
-	using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto
+        using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto
 
       have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)"
         by (auto simp add: dist_norm algebra_simps)
@@ -2435,15 +2435,15 @@
     { fix e::real
       assume "e>0"
       with `?rhs` obtain N where N:"\<forall>n\<ge>N. dist (s n) (s N) < e/2"
-	by (erule_tac x="e/2" in allE) auto
+        by (erule_tac x="e/2" in allE) auto
       { fix n m
-	assume nm:"N \<le> m \<and> N \<le> n"
-	hence "dist (s m) (s n) < e" using N
-	  using dist_triangle_half_l[of "s m" "s N" "e" "s n"]
-	  by blast
+        assume nm:"N \<le> m \<and> N \<le> n"
+        hence "dist (s m) (s n) < e" using N
+          using dist_triangle_half_l[of "s m" "s N" "e" "s n"]
+          by blast
       }
       hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"
-	by blast
+        by blast
     }
     hence ?lhs
       unfolding cauchy_def
@@ -2488,10 +2488,10 @@
       from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2" unfolding cauchy_def using `e>0` apply (erule_tac x="e/2" in allE) by auto
       from lr(3)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2" using `e>0` by auto
       { fix n::nat assume n:"n \<ge> max N M"
-	have "dist ((f \<circ> r) n) l < e/2" using n M by auto
-	moreover have "r n \<ge> N" using lr'[of n] n by auto
-	hence "dist (f n) ((f \<circ> r) n) < e / 2" using N using n by auto
-	ultimately have "dist (f n) l < e" using dist_triangle_half_r[of "f (r n)" "f n" e l] by (auto simp add: dist_commute)  }
+        have "dist ((f \<circ> r) n) l < e/2" using n M by auto
+        moreover have "r n \<ge> N" using lr'[of n] n by auto
+        hence "dist (f n) ((f \<circ> r) n) < e / 2" using N using n by auto
+        ultimately have "dist (f n) l < e" using dist_triangle_half_r[of "f (r n)" "f n" e l] by (auto simp add: dist_commute)  }
       hence "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast  }
     hence "\<exists>l\<in>s. (f ---> l) sequentially" using `l\<in>s` unfolding Lim_sequentially by auto  }
   thus ?thesis unfolding complete_def by auto
@@ -2585,7 +2585,7 @@
       have "\<not> s \<subseteq> (\<Union>x\<in>x ` {0..<n}. ball x e)" using assm apply simp apply(erule_tac x="x ` {0 ..< n}" in allE) using as by auto
       then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)" unfolding subset_eq by auto
       have "Q (x n)" unfolding x_def and helper_1.simps[of s e n]
-	apply(rule someI2[where a=z]) unfolding x_def[symmetric] and Q_def using z by auto
+        apply(rule someI2[where a=z]) unfolding x_def[symmetric] and Q_def using z by auto
       thus "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)" unfolding Q_def by auto
     qed }
   hence "\<forall>n::nat. x n \<in> s" and x:"\<forall>n. \<forall>m < n. \<not> (dist (x m) (x n) < e)" by blast+
@@ -2715,12 +2715,12 @@
       case (Suc n)
       have *:"dist undefined (x n) + 1 < dist undefined (x (Suc n))"
         unfolding x_def and helper_2.simps
-	using beyond(2)[of "dist undefined (helper_2 beyond n) + 1"] by auto
+        using beyond(2)[of "dist undefined (helper_2 beyond n) + 1"] by auto
       thus ?case proof(cases "m < n")
-	case True thus ?thesis using Suc and * by auto
+        case True thus ?thesis using Suc and * by auto
       next
-	case False hence "m = n" using Suc(2) by auto
-	thus ?thesis using * by auto
+        case False hence "m = n" using Suc(2) by auto
+        thus ?thesis using * by auto
       qed
     qed  } note * = this
   { fix m n ::nat assume "m\<noteq>n"
@@ -3011,9 +3011,9 @@
       moreover
       have "r (max N n) \<ge> n" using lr(2) using subseq_bigger[of r "max N n"] by auto
       hence "(x \<circ> r) (max N n) \<in> s n"
-	using x apply(erule_tac x=n in allE)
-	using x apply(erule_tac x="r (max N n)" in allE)
-	using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto
+        using x apply(erule_tac x=n in allE)
+        using x apply(erule_tac x="r (max N n)" in allE)
+        using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto
       ultimately have "\<exists>y\<in>s n. dist y l < e" by auto
     }
     hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by blast
@@ -3086,9 +3086,9 @@
     then obtain N::nat where N:"\<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e / 2" using l[THEN spec[where x="e/2"]] by auto
     { fix n m::nat and x::"'b" assume "N \<le> m \<and> N \<le> n \<and> P x"
       hence "dist (s m x) (s n x) < e"
-	using N[THEN spec[where x=m], THEN spec[where x=x]]
-	using N[THEN spec[where x=n], THEN spec[where x=x]]
-	using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto  }
+        using N[THEN spec[where x=m], THEN spec[where x=x]]
+        using N[THEN spec[where x=n], THEN spec[where x=x]]
+        using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto  }
     hence "\<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e"  by auto  }
   thus ?rhs by auto
 next
@@ -3101,10 +3101,10 @@
       using `?rhs`[THEN spec[where x="e/2"]] by auto
     { fix x assume "P x"
       then obtain M where M:"\<forall>n\<ge>M. dist (s n x) (l x) < e/2"
-	using l[THEN spec[where x=x], unfolded Lim_sequentially] using `e>0` by(auto elim!: allE[where x="e/2"])
+        using l[THEN spec[where x=x], unfolded Lim_sequentially] using `e>0` by(auto elim!: allE[where x="e/2"])
       fix n::nat assume "n\<ge>N"
       hence "dist(s n x)(l x) < e"  using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]]
-	using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_commute)  }
+        using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_commute)  }
     hence "\<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e" by auto }
   thus ?lhs by auto
 qed
@@ -3172,7 +3172,7 @@
       using `?lhs`[unfolded continuous_within Lim_within] by auto
     { fix y assume "y\<in>f ` (ball x d \<inter> s)"
       hence "y \<in> ball (f x) e" using d(2) unfolding dist_nz[THEN sym]
-	apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto
+        apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto
     }
     hence "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute)  }
   thus ?rhs by auto
@@ -3307,9 +3307,9 @@
       hence "\<exists>N::nat. inverse (real (N + 1)) < d" using real_arch_inv[of d] by (auto, rule_tac x="n - 1" in exI)auto
       then obtain N::nat where N:"inverse (real (N + 1)) < d" by auto
       { fix n::nat assume n:"n\<ge>N"
-	hence "dist (x (inverse (real (n + 1)))) a < inverse (real (n + 1))" using x[THEN spec[where x="inverse (real (n + 1))"]] by auto
-	moreover have "inverse (real (n + 1)) < d" using N n by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
-	ultimately have "dist (x (inverse (real (n + 1)))) a < d" by auto
+        hence "dist (x (inverse (real (n + 1)))) a < inverse (real (n + 1))" using x[THEN spec[where x="inverse (real (n + 1))"]] by auto
+        moreover have "inverse (real (n + 1)) < d" using N n by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
+        ultimately have "dist (x (inverse (real (n + 1)))) a < d" by auto
       }
       hence "\<exists>N::nat. \<forall>n\<ge>N. dist (x (inverse (real (n + 1)))) a < d" by auto
     }
@@ -3345,12 +3345,12 @@
   { fix x y assume x:"\<forall>n. x n \<in> s" and y:"\<forall>n. y n \<in> s" and xy:"((\<lambda>n. x n - y n) ---> 0) sequentially"
     { fix e::real assume "e>0"
       then obtain d where "d>0" and d:"\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
-	using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
+        using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
       obtain N where N:"\<forall>n\<ge>N. norm (x n - y n - 0) < d" using xy[unfolded Lim_sequentially dist_norm] and `d>0` by auto
       { fix n assume "n\<ge>N"
-	hence "norm (f (x n) - f (y n) - 0) < e"
-	  using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y
-	  unfolding dist_commute and dist_norm by simp  }
+        hence "norm (f (x n) - f (y n) - 0) < e"
+          using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y
+          unfolding dist_commute and dist_norm by simp  }
       hence "\<exists>N. \<forall>n\<ge>N. norm (f (x n) - f (y n) - 0) < e"  by auto  }
     hence "((\<lambda>n. f(x n) - f(y n)) ---> 0) sequentially" unfolding Lim_sequentially and dist_norm by auto  }
   thus ?rhs by auto
@@ -3370,10 +3370,10 @@
     { fix e::real assume "e>0"
       then obtain N::nat where "N \<noteq> 0" and N:"0 < inverse (real N) \<and> inverse (real N) < e" unfolding real_arch_inv[of e]   by auto
       { fix n::nat assume "n\<ge>N"
-	hence "inverse (real n + 1) < inverse (real N)" using real_of_nat_ge_zero and `N\<noteq>0` by auto
-	also have "\<dots> < e" using N by auto
-	finally have "inverse (real n + 1) < e" by auto
-	hence "dist (x n - y n) 0 < e" unfolding 1 using xy0[THEN spec[where x=n]] by auto  }
+        hence "inverse (real n + 1) < inverse (real N)" using real_of_nat_ge_zero and `N\<noteq>0` by auto
+        also have "\<dots> < e" using N by auto
+        finally have "inverse (real n + 1) < e" by auto
+        hence "dist (x n - y n) 0 < e" unfolding 1 using xy0[THEN spec[where x=n]] by auto  }
       hence "\<exists>N. \<forall>n\<ge>N. dist (x n - y n) 0 < e" by auto  }
     hence "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n) - f (y n)) 0 < e" using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding Lim_sequentially by auto
     hence False unfolding 2 using fxy and `e>0` by auto  }
@@ -3598,9 +3598,9 @@
     moreover have "x \<in> ball x d" unfolding centre_in_ball using `d>0` by simp
     moreover
     { fix x' assume "x'\<in>ball x d" hence "f x' \<in> t"
-	using e[unfolded subset_eq Ball_def mem_ball, THEN spec[where x="f x'"]]    d[THEN spec[where x=x']]
-	unfolding mem_ball apply (auto simp add: dist_commute)
-	unfolding dist_nz[THEN sym] using as(2) by auto  }
+        using e[unfolded subset_eq Ball_def mem_ball, THEN spec[where x="f x'"]]    d[THEN spec[where x=x']]
+        unfolding mem_ball apply (auto simp add: dist_commute)
+        unfolding dist_nz[THEN sym] using as(2) by auto  }
     hence "\<forall>x'\<in>ball x d. f x' \<in> t" by auto
     ultimately have "\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x'\<in>s. f x' \<in> t)"
       apply(rule_tac x="ball x d" in exI) by simp  }
@@ -3613,7 +3613,7 @@
     then obtain d where "d>0" and d:"ball x d \<subseteq> s" unfolding open_contains_ball by auto
     { fix y assume "0 < dist y x \<and> dist y x < d"
       hence "dist (f y) (f x) < e" using d[unfolded subset_eq Ball_def mem_ball, THEN spec[where x=y]]
-	using s(3)[THEN bspec[where x=y], unfolded mem_ball] by (auto simp add: dist_commute)  }
+        using s(3)[THEN bspec[where x=y], unfolded mem_ball] by (auto simp add: dist_commute)  }
     hence "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" using `d>0` by auto  }
   thus ?lhs unfolding continuous_at Lim_at by auto
 qed
@@ -3638,7 +3638,7 @@
   { fix e::real and x assume "x\<in>s" "e>0"
     { fix xa x' assume "dist (f xa) (f x) < e" "xa \<in> s" "x' \<in> s" "dist (f xa) (f x') < e - dist (f xa) (f x)"
       hence "dist (f x') (f x) < e" using dist_triangle[of "f x'" "f x" "f xa"]
-	by (auto simp add: dist_commute)  }
+        by (auto simp add: dist_commute)  }
     hence "ball (f x) e \<inter> f ` s \<subseteq> f ` s \<and> (\<forall>xa\<in>ball (f x) e \<inter> f ` s. \<exists>ea>0. \<forall>x'\<in>f ` s. dist x' xa < ea \<longrightarrow> x' \<in> ball (f x) e \<inter> f ` s)" apply auto
       apply(rule_tac x="e - dist (f xa) (f x)" in exI) using `e>0` by (auto simp add: dist_commute)
     hence "\<forall>xa\<in>{xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}. \<exists>ea>0. \<forall>x'\<in>s. dist x' xa < ea \<longrightarrow> x' \<in> {xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}"
@@ -3838,8 +3838,8 @@
     moreover
     { fix y assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"
       hence "norm ((1 / c) *\<^sub>R y - x) < e" unfolding dist_norm
-	using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1)
-	  assms(1)[unfolded zero_less_abs_iff[THEN sym]] by (simp del:zero_less_abs_iff)
+        using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1)
+          assms(1)[unfolded zero_less_abs_iff[THEN sym]] by (simp del:zero_less_abs_iff)
       hence "y \<in> op *\<^sub>R c ` s" using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "op *\<^sub>R c"]  e[THEN spec[where x="(1 / c) *\<^sub>R y"]]  assms(1) unfolding dist_norm scaleR_scaleR by auto  }
     ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> op *\<^sub>R c ` s" apply(rule_tac x="e * abs c" in exI) by auto  }
   thus ?thesis unfolding open_dist by auto
@@ -3951,13 +3951,13 @@
       obtain z where "z\<in>s" and z:"ball x ea \<subseteq> ball z (d z (e / 2))" using ea[THEN bspec[where x=x]] and `x\<in>s` by auto
       hence "x\<in>ball z (d z (e / 2))" using `ea>0` unfolding subset_eq by auto
       hence "dist (f z) (f x) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `x\<in>s` and `z\<in>s`
-	by (auto  simp add: dist_commute)
+        by (auto  simp add: dist_commute)
       moreover have "y\<in>ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq]
-	by (auto simp add: dist_commute)
+        by (auto simp add: dist_commute)
       hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\<in>s` and `z\<in>s`
-	by (auto  simp add: dist_commute)
+        by (auto  simp add: dist_commute)
       ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"]
-	by (auto simp add: dist_commute)  }
+        by (auto simp add: dist_commute)  }
     then have "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `ea>0` by auto  }
   thus ?thesis unfolding uniformly_continuous_on_def by auto
 qed
@@ -4009,9 +4009,9 @@
     { fix y assume "y\<in>s" "dist y x < d"
       hence "dist (f n y) (f n x) < e / 3" using d[THEN bspec[where x=y]] by auto
       hence "norm (f n y - g x) < 2 * e / 3" using norm_triangle_lt[of "f n y - f n x" "f n x - g x" "2*e/3"]
-	using n(1)[THEN bspec[where x=x], OF `x\<in>s`] unfolding dist_norm unfolding ab_group_add_class.ab_diff_minus by auto
+        using n(1)[THEN bspec[where x=x], OF `x\<in>s`] unfolding dist_norm unfolding ab_group_add_class.ab_diff_minus by auto
       hence "dist (g y) (g x) < e" unfolding dist_norm using n(1)[THEN bspec[where x=y], OF `y\<in>s`]
-	unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff)  }
+        unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff)  }
     hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" using `d>0` by auto  }
   thus ?thesis unfolding continuous_on_def by auto
 qed
@@ -4307,8 +4307,8 @@
     { fix e::real assume "e>0"
       then obtain N::nat where N:"\<forall>n\<ge>N. dist (x n) l < e" using as(2)[unfolded Lim_sequentially, THEN spec[where x=e]] by auto
       { fix n::nat assume "n\<ge>N"
-	hence "dist (fstcart (x n)) (fstcart l) < e" "dist (sndcart (x n)) (sndcart l) < e"
-	  using N[THEN spec[where x=n]] dist_fstcart[of "x n" l] dist_sndcart[of "x n" l] by auto   }
+        hence "dist (fstcart (x n)) (fstcart l) < e" "dist (sndcart (x n)) (sndcart l) < e"
+          using N[THEN spec[where x=n]] dist_fstcart[of "x n" l] dist_sndcart[of "x n" l] by auto   }
       hence "\<exists>N. \<forall>n\<ge>N. dist (fstcart (x n)) (fstcart l) < e" "\<exists>N. \<forall>n\<ge>N. dist (sndcart (x n)) (sndcart l) < e" by auto  }
     ultimately have "fstcart l \<in> s" "sndcart l \<in> t"
       using assms(1)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. fstcart (x n)"], THEN spec[where x="fstcart l"]]
@@ -4483,16 +4483,16 @@
       }
       moreover
       { fix e::real assume "e>0"
-	hence "0 < e *\<bar>c\<bar>"  using `c\<noteq>0` mult_pos_pos[of e "abs c"] by auto
-	then obtain N where "\<forall>n\<ge>N. dist (x n) l < e * \<bar>c\<bar>"
+        hence "0 < e *\<bar>c\<bar>"  using `c\<noteq>0` mult_pos_pos[of e "abs c"] by auto
+        then obtain N where "\<forall>n\<ge>N. dist (x n) l < e * \<bar>c\<bar>"
           using as(2)[unfolded Lim_sequentially, THEN spec[where x="e * abs c"]] by auto
-	hence "\<exists>N. \<forall>n\<ge>N. dist (scaleR (1 / c) (x n)) (scaleR (1 / c) l) < e"
+        hence "\<exists>N. \<forall>n\<ge>N. dist (scaleR (1 / c) (x n)) (scaleR (1 / c) l) < e"
           unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym]
-	  using mult_imp_div_pos_less[of "abs c" _ e] `c\<noteq>0` by auto  }
+          using mult_imp_div_pos_less[of "abs c" _ e] `c\<noteq>0` by auto  }
       hence "((\<lambda>n. scaleR (1 / c) (x n)) ---> scaleR (1 / c) l) sequentially" unfolding Lim_sequentially by auto
       ultimately have "l \<in> scaleR c ` s"
         using assms[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. scaleR (1/c) (x n)"], THEN spec[where x="scaleR (1/c) l"]]
-	unfolding image_iff using `c\<noteq>0` apply(rule_tac x="scaleR (1 / c) l" in bexI) by auto  }
+        unfolding image_iff using `c\<noteq>0` apply(rule_tac x="scaleR (1 / c) l" in bexI) by auto  }
     thus ?thesis unfolding closed_sequential_limits by fast
   qed
 qed
@@ -4658,8 +4658,8 @@
     { fix i
       have "a$i < b$i" using as[THEN spec[where x=i]] by auto
       hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i"
-	unfolding vector_smult_component and vector_add_component
-	by (auto simp add: less_divide_eq_number_of1)  }
+        unfolding vector_smult_component and vector_add_component
+        by (auto simp add: less_divide_eq_number_of1)  }
     hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
   ultimately show ?th1 by blast
 
@@ -4673,8 +4673,8 @@
     { fix i
       have "a$i \<le> b$i" using as[THEN spec[where x=i]] by auto
       hence "a$i \<le> ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \<le> b$i"
-	unfolding vector_smult_component and vector_add_component
-	by (auto simp add: less_divide_eq_number_of1)  }
+        unfolding vector_smult_component and vector_add_component
+        by (auto simp add: less_divide_eq_number_of1)  }
     hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
   ultimately show ?th2 by blast
 qed
@@ -4735,30 +4735,30 @@
     { let ?x = "(\<chi> j. (if j=i then ((min (a$j) (d$j))+c$j)/2 else (c$j+d$j)/2))::real^'n"
       assume as2: "a$i > c$i"
       { fix j
-	have "c $ j < ?x $ j \<and> ?x $ j < d $ j" unfolding Cart_lambda_beta
-	  apply(cases "j=i") using as(2)[THEN spec[where x=j]]
-	  by (auto simp add: less_divide_eq_number_of1 as2)  }
+        have "c $ j < ?x $ j \<and> ?x $ j < d $ j" unfolding Cart_lambda_beta
+          apply(cases "j=i") using as(2)[THEN spec[where x=j]]
+          by (auto simp add: less_divide_eq_number_of1 as2)  }
       hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
       moreover
       have "?x\<notin>{a .. b}"
-	unfolding mem_interval apply auto apply(rule_tac x=i in exI)
-	using as(2)[THEN spec[where x=i]] and as2
-	by (auto simp add: less_divide_eq_number_of1)
+        unfolding mem_interval apply auto apply(rule_tac x=i in exI)
+        using as(2)[THEN spec[where x=i]] and as2
+        by (auto simp add: less_divide_eq_number_of1)
       ultimately have False using as by auto  }
     hence "a$i \<le> c$i" by(rule ccontr)auto
     moreover
     { let ?x = "(\<chi> j. (if j=i then ((max (b$j) (c$j))+d$j)/2 else (c$j+d$j)/2))::real^'n"
       assume as2: "b$i < d$i"
       { fix j
-	have "d $ j > ?x $ j \<and> ?x $ j > c $ j" unfolding Cart_lambda_beta
-	  apply(cases "j=i") using as(2)[THEN spec[where x=j]]
-	  by (auto simp add: less_divide_eq_number_of1 as2)  }
+        have "d $ j > ?x $ j \<and> ?x $ j > c $ j" unfolding Cart_lambda_beta
+          apply(cases "j=i") using as(2)[THEN spec[where x=j]]
+          by (auto simp add: less_divide_eq_number_of1 as2)  }
       hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
       moreover
       have "?x\<notin>{a .. b}"
-	unfolding mem_interval apply auto apply(rule_tac x=i in exI)
-	using as(2)[THEN spec[where x=i]] and as2
-	by (auto simp add: less_divide_eq_number_of1)
+        unfolding mem_interval apply auto apply(rule_tac x=i in exI)
+        using as(2)[THEN spec[where x=i]] and as2
+        by (auto simp add: less_divide_eq_number_of1)
       ultimately have False using as by auto  }
     hence "b$i \<ge> d$i" by(rule ccontr)auto
     ultimately
@@ -4802,8 +4802,8 @@
   { fix x assume x:"x\<in>{a<..<b}"
     { fix i
       have "\<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i"
-	using x[unfolded mem_interval, THEN spec[where x=i]]
-	using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto  }
+        using x[unfolded mem_interval, THEN spec[where x=i]]
+        using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto  }
 
     hence "\<forall>i. \<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i" by auto
     then obtain d where d:"\<forall>i. 0 < d i \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d i \<longrightarrow> a $ i < x' \<and> x' < b $ i)"
@@ -4815,10 +4815,10 @@
     moreover
     { fix x' assume as:"dist x' x < ?d"
       { fix i
-	have "\<bar>x'$i - x $ i\<bar> < d i"
-	  using norm_bound_component_lt[OF as[unfolded dist_norm], of i]
-	  unfolding vector_minus_component and Min_gr_iff[OF **] by auto
-	hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN spec[where x=i]] by auto  }
+        have "\<bar>x'$i - x $ i\<bar> < d i"
+          using norm_bound_component_lt[OF as[unfolded dist_norm], of i]
+          unfolding vector_minus_component and Min_gr_iff[OF **] by auto
+        hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN spec[where x=i]] by auto  }
       hence "a < x' \<and> x' < b" unfolding vector_less_def by auto  }
     ultimately have "\<exists>e>0. \<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a<..<b}" by (auto, rule_tac x="?d" in exI, simp)
   }
@@ -4831,13 +4831,13 @@
     { assume xa:"a$i > x$i"
       with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < a$i - x$i" by(erule_tac x="a$i - x$i" in allE)auto
       hence False unfolding mem_interval and dist_norm
-	using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xa by(auto elim!: allE[where x=i])
+        using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xa by(auto elim!: allE[where x=i])
     } hence "a$i \<le> x$i" by(rule ccontr)auto
     moreover
     { assume xb:"b$i < x$i"
       with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < x$i - b$i" by(erule_tac x="x$i - b$i" in allE)auto
       hence False unfolding mem_interval and dist_norm
-	using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xb by(auto elim!: allE[where x=i])
+        using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xb by(auto elim!: allE[where x=i])
     } hence "x$i \<le> b$i" by(rule ccontr)auto
     ultimately
     have "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" by auto }
@@ -4854,17 +4854,17 @@
     then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a..b}" unfolding open_dist and subset_eq by auto
     { fix i
       have "dist (x - (e / 2) *\<^sub>R basis i) x < e"
-	   "dist (x + (e / 2) *\<^sub>R basis i) x < e"
-	unfolding dist_norm apply auto
-	unfolding norm_minus_cancel using norm_basis[of i] and `e>0` by auto
+           "dist (x + (e / 2) *\<^sub>R basis i) x < e"
+        unfolding dist_norm apply auto
+        unfolding norm_minus_cancel using norm_basis[of i] and `e>0` by auto
       hence "a $ i \<le> (x - (e / 2) *\<^sub>R basis i) $ i"
                     "(x + (e / 2) *\<^sub>R basis i) $ i \<le> b $ i"
-	using e[THEN spec[where x="x - (e/2) *\<^sub>R basis i"]]
-	and   e[THEN spec[where x="x + (e/2) *\<^sub>R basis i"]]
-	unfolding mem_interval by (auto elim!: allE[where x=i])
+        using e[THEN spec[where x="x - (e/2) *\<^sub>R basis i"]]
+        and   e[THEN spec[where x="x + (e/2) *\<^sub>R basis i"]]
+        unfolding mem_interval by (auto elim!: allE[where x=i])
       hence "a $ i < x $ i" and "x $ i < b $ i"
-	unfolding vector_minus_component and vector_add_component
-	unfolding vector_smult_component and basis_component using `e>0` by auto   }
+        unfolding vector_minus_component and vector_add_component
+        unfolding vector_smult_component and basis_component using `e>0` by auto   }
     hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
   thus "?L \<subseteq> ?R" unfolding interior_def and subset_eq by auto
 qed
@@ -4943,22 +4943,22 @@
     { fix n assume fn:"f n < b \<longrightarrow> a < f n \<longrightarrow> f n = x" and xc:"x \<noteq> ?c"
       have *:"0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" unfolding inverse_le_1_iff by auto
       have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x =
-	x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
+        x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
         by (auto simp add: algebra_simps)
       hence "f n < b" and "a < f n" using open_closed_interval_convex[OF open_interval_midpoint[OF assms] as *] unfolding f_def by auto
       hence False using fn unfolding f_def using xc by(auto simp add: vector_mul_lcancel vector_ssub_ldistrib)  }
     moreover
     { assume "\<not> (f ---> x) sequentially"
       { fix e::real assume "e>0"
-	hence "\<exists>N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto
-	then obtain N::nat where "inverse (real (N + 1)) < e" by auto
-	hence "\<forall>n\<ge>N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
-	hence "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto  }
+        hence "\<exists>N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto
+        then obtain N::nat where "inverse (real (N + 1)) < e" by auto
+        hence "\<forall>n\<ge>N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
+        hence "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto  }
       hence "((\<lambda>n. inverse (real n + 1)) ---> 0) sequentially"
-	unfolding Lim_sequentially by(auto simp add: dist_norm)
+        unfolding Lim_sequentially by(auto simp add: dist_norm)
       hence "(f ---> x) sequentially" unfolding f_def
-	using Lim_add[OF Lim_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x]
-	using Lim_vmul[of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto  }
+        using Lim_add[OF Lim_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x]
+        using Lim_vmul[of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto  }
     ultimately have "x \<in> closure {a<..<b}"
       using as and open_interval_midpoint[OF assms] unfolding closure_def unfolding islimpt_sequential by(cases "x=?c")auto  }
   thus ?thesis using closure_minimal[OF interval_open_subset_closed closed_interval, of a b] by blast
@@ -5458,10 +5458,10 @@
     { fix n assume "n\<ge>N"
       hence "norm (f (x n - x N)) < e * d" using N[THEN spec[where x=n]] unfolding f.diff[THEN sym] by auto
       moreover have "e * norm (x n - x N) \<le> norm (f (x n - x N))"
-	using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]]
-	using normf[THEN bspec[where x="x n - x N"]] by auto
+        using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]]
+        using normf[THEN bspec[where x="x n - x N"]] by auto
       ultimately have "norm (x n - x N) < d" using `e>0`
-	using mult_left_less_imp_less[of e "norm (x n - x N)" d] by auto   }
+        using mult_left_less_imp_less[of e "norm (x n - x N)" d] by auto   }
     hence "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" by auto }
   thus ?thesis unfolding cauchy and dist_norm by auto
 qed
@@ -5534,8 +5534,8 @@
       have "\<forall>c. \<forall>x\<in>s. c *\<^sub>R x \<in> s" using s[unfolded subspace_def smult_conv_scaleR] by auto
       hence "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}" using `x\<in>s` and `x\<noteq>0` by auto
       thus "norm (f b) / norm b * norm x \<le> norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]]
-	unfolding f.scaleR and ba using `x\<noteq>0` `a\<noteq>0`
-	by (auto simp add: real_mult_commute pos_le_divide_eq pos_divide_le_eq)
+        unfolding f.scaleR and ba using `x\<noteq>0` `a\<noteq>0`
+        by (auto simp add: real_mult_commute pos_le_divide_eq pos_divide_le_eq)
     qed }
   ultimately
   show ?thesis by auto
@@ -5569,8 +5569,8 @@
     moreover
     { assume x:"x\<in>\<Inter>?Bs"
       { fix i assume i:"i \<in> ?D"
-	then obtain B where BB:"B \<in> ?Bs" and B:"B = {x::real^'n. inner (basis i) x = 0}" by auto
-	hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto  }
+        then obtain B where BB:"B \<in> ?Bs" and B:"B = {x::real^'n. inner (basis i) x = 0}" by auto
+        hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto  }
       hence "x\<in>?A" by auto }
     ultimately have "x\<in>?A \<longleftrightarrow> x\<in> \<Inter>?Bs" by auto }
   hence "?A = \<Inter> ?Bs" by auto
@@ -5601,19 +5601,19 @@
       def y \<equiv> "x - x$k *\<^sub>R basis k"
       have y:"x = y + (x$k) *\<^sub>R basis k" unfolding y_def by auto
       { fix i assume i':"i \<notin> F"
-	hence "y $ i = 0" unfolding y_def unfolding vector_minus_component
-	  and vector_smult_component and basis_component
-	  using *[THEN spec[where x=i]] by auto }
+        hence "y $ i = 0" unfolding y_def unfolding vector_minus_component
+          and vector_smult_component and basis_component
+          using *[THEN spec[where x=i]] by auto }
       hence "y \<in> span (basis ` (insert k F))" using insert(3)
-	using span_mono[of "?bas ` F" "?bas ` (insert k F)"]
-	using image_mono[OF **, of basis] by auto
+        using span_mono[of "?bas ` F" "?bas ` (insert k F)"]
+        using image_mono[OF **, of basis] by auto
       moreover
       have "basis k \<in> span (?bas ` (insert k F))" by(rule span_superset, auto)
       hence "x$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
         using span_mul [where 'a=real, unfolded smult_conv_scaleR] by auto
       ultimately
       have "y + x$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
-	using span_add by auto
+        using span_add by auto
       thus ?case using y by auto
     qed
   }
@@ -5758,14 +5758,14 @@
     hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
       unfolding image_iff Bex_def mem_interval vector_less_eq_def
       apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
-	intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
+        intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
       by(auto simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute diff_le_iff)
   } moreover
   { fix y assume "m *\<^sub>R b + c \<le> y" "y \<le> m *\<^sub>R a + c" "m < 0"
     hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
       unfolding image_iff Bex_def mem_interval vector_less_eq_def
       apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
-	intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
+        intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
       by(auto simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute diff_le_iff)
   }
   ultimately show ?thesis using False by auto
@@ -5802,9 +5802,9 @@
     next
       case (Suc m)
       hence "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d"
-	using `0 \<le> c` using mult_mono1_class.mult_mono1[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by auto
+        using `0 \<le> c` using mult_mono1_class.mult_mono1[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by auto
       thus ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s]
-	unfolding fzn and mult_le_cancel_left by auto
+        unfolding fzn and mult_le_cancel_left by auto
     qed
   } note cf_z = this
 
@@ -5815,15 +5815,15 @@
     next
       case (Suc k)
       have "(1 - c) * dist (z m) (z (m + Suc k)) \<le> (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))"
-	using dist_triangle and c by(auto simp add: dist_triangle)
+        using dist_triangle and c by(auto simp add: dist_triangle)
       also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"
-	using cf_z[of "m + k"] and c by auto
+        using cf_z[of "m + k"] and c by auto
       also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"
-	using Suc by (auto simp add: ring_simps)
+        using Suc by (auto simp add: ring_simps)
       also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
-	unfolding power_add by (auto simp add: ring_simps)
+        unfolding power_add by (auto simp add: ring_simps)
       also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
-	using c by (auto simp add: ring_simps)
+        using c by (auto simp add: ring_simps)
       finally show ?case by auto
     qed
   } note cf_z2 = this
@@ -5835,37 +5835,37 @@
       thus ?thesis using `e>0` by auto
     next
       case False hence "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"]
-	by (metis False d_def real_less_def)
+        by (metis False d_def real_less_def)
       hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0`
-	using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto
+        using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto
       then obtain N where N:"c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto
       { fix m n::nat assume "m>n" and as:"m\<ge>N" "n\<ge>N"
-	have *:"c ^ n \<le> c ^ N" using `n\<ge>N` and c using power_decreasing[OF `n\<ge>N`, of c] by auto
-	have "1 - c ^ (m - n) > 0" using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto
-	hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0"
-	  using real_mult_order[OF `d>0`, of "1 - c ^ (m - n)"]
-	  using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
-	  using `0 < 1 - c` by auto
-
-	have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"
-	  using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`]
-	  by (auto simp add: real_mult_commute dist_commute)
-	also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
-	  using mult_right_mono[OF * order_less_imp_le[OF **]]
-	  unfolding real_mult_assoc by auto
-	also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)"
-	  using mult_strict_right_mono[OF N **] unfolding real_mult_assoc by auto
-	also have "\<dots> = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto
-	also have "\<dots> \<le> e" using c and `1 - c ^ (m - n) > 0` and `e>0` using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto
-	finally have  "dist (z m) (z n) < e" by auto
+        have *:"c ^ n \<le> c ^ N" using `n\<ge>N` and c using power_decreasing[OF `n\<ge>N`, of c] by auto
+        have "1 - c ^ (m - n) > 0" using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto
+        hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0"
+          using real_mult_order[OF `d>0`, of "1 - c ^ (m - n)"]
+          using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
+          using `0 < 1 - c` by auto
+
+        have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"
+          using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`]
+          by (auto simp add: real_mult_commute dist_commute)
+        also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
+          using mult_right_mono[OF * order_less_imp_le[OF **]]
+          unfolding real_mult_assoc by auto
+        also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)"
+          using mult_strict_right_mono[OF N **] unfolding real_mult_assoc by auto
+        also have "\<dots> = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto
+        also have "\<dots> \<le> e" using c and `1 - c ^ (m - n) > 0` and `e>0` using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto
+        finally have  "dist (z m) (z n) < e" by auto
       } note * = this
       { fix m n::nat assume as:"N\<le>m" "N\<le>n"
-	hence "dist (z n) (z m) < e"
-	proof(cases "n = m")
-	  case True thus ?thesis using `e>0` by auto
-	next
-	  case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_commute)
-	qed }
+        hence "dist (z n) (z m) < e"
+        proof(cases "n = m")
+          case True thus ?thesis using `e>0` by auto
+        next
+          case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_commute)
+        qed }
       thus ?thesis by auto
     qed
   }
@@ -5943,11 +5943,11 @@
     next
       case (Suc n)
       thus ?case proof(cases "m\<le>n")
-	case True thus ?thesis using Suc(1)
-	  using dist'[OF fs fs, OF `w\<in>s` `z\<in>s`, of n n] by auto
+        case True thus ?thesis using Suc(1)
+          using dist'[OF fs fs, OF `w\<in>s` `z\<in>s`, of n n] by auto
       next
-	case False hence mn:"m = Suc n" using Suc(2) by simp
-	show ?thesis unfolding mn  by auto
+        case False hence mn:"m = Suc n" using Suc(2) by simp
+        show ?thesis unfolding mn  by auto
       qed
     qed } note distf = this
 
@@ -5969,23 +5969,23 @@
     have *:"\<And>fx fy (x::'a) y. dist fx fy \<le> dist x y \<Longrightarrow> \<not> (dist (fx - fy) (a - b) < dist a b - dist x y)" unfolding dist_norm by norm
     { fix x y :: 'a
       have "dist (-x) (-y) = dist x y" unfolding dist_norm
-	using norm_minus_cancel[of "x - y"] by (auto simp add: uminus_add_conv_diff) } note ** = this
+        using norm_minus_cancel[of "x - y"] by (auto simp add: uminus_add_conv_diff) } note ** = this
 
     { assume as:"dist a b > dist (f n x) (f n y)"
       then obtain Na Nb where "\<forall>m\<ge>Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2"
-	and "\<forall>m\<ge>Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2"
-	using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: less_divide_eq_number_of1)
+        and "\<forall>m\<ge>Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2"
+        using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: less_divide_eq_number_of1)
       hence "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) < dist a b - dist (f n x) (f n y)"
-	apply(erule_tac x="Na+Nb+n" in allE)
-	apply(erule_tac x="Na+Nb+n" in allE) apply simp
-	using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)"
+        apply(erule_tac x="Na+Nb+n" in allE)
+        apply(erule_tac x="Na+Nb+n" in allE) apply simp
+        using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)"
           "-b"  "- f (r (Na + Nb + n)) y"]
-	unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute)
+        unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute)
       moreover
       have "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) \<ge> dist a b - dist (f n x) (f n y)"
-	using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
-	using subseq_bigger[OF r, of "Na+Nb+n"]
-	using *[of "f (r (Na + Nb + n)) x" "f (r (Na + Nb + n)) y" "f n x" "f n y"] by auto
+        using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
+        using subseq_bigger[OF r, of "Na+Nb+n"]
+        using *[of "f (r (Na + Nb + n)) x" "f (r (Na + Nb + n)) y" "f n x" "f n y"] by auto
       ultimately have False by simp
     }
     hence "dist a b \<le> dist (f n x) (f n y)" by(rule ccontr)auto }
--- a/src/HOL/Library/Univ_Poly.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Univ_Poly.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,5 @@
-(*  Title:       Univ_Poly.thy
-    Author:      Amine Chaieb
+(*  Title:      HOL/Library/Univ_Poly.thy
+    Author:     Amine Chaieb
 *)
 
 header {* Univariate Polynomials *}
@@ -205,7 +205,7 @@
     obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
     have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)"
       using qr by(cases q, simp_all add: algebra_simps diff_def[symmetric]
-	minus_mult_left[symmetric] right_minus)
+        minus_mult_left[symmetric] right_minus)
     hence "\<exists>q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast}
   thus ?case by blast
 qed
@@ -220,7 +220,7 @@
   moreover
   {fix x xs assume p: "p = x#xs"
     {fix q assume "p = [-a, 1] *** q" hence "poly p a = 0"
-	by (simp add: poly_add poly_cmult minus_mult_left[symmetric])}
+        by (simp add: poly_add poly_cmult minus_mult_left[symmetric])}
     moreover
     {assume p0: "poly p a = 0"
       from poly_linear_rem[of x xs a] obtain q r
@@ -701,7 +701,7 @@
 apply (simp add: divides_def fun_eq poly_mult)
 apply (rule_tac x = "[]" in exI)
 apply (auto dest!: order2 [where a=a]
-	    intro: poly_exp_divides simp del: pexp_Suc)
+            intro: poly_exp_divides simp del: pexp_Suc)
 done
 qed
 
--- a/src/HOL/Library/Word.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Word.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,5 @@
 (*  Title:      HOL/Library/Word.thy
-    Author:     Sebastian Skalberg (TU Muenchen)
+    Author:     Sebastian Skalberg, TU Muenchen
 *)
 
 header {* Binary Words *}
@@ -654,7 +654,7 @@
           also have "... = \<one> # rev ys @ [y]"
             by simp
           finally show "nat_to_bv (2 * 2 ^ length ys + (bv_to_nat (rev ys) * 2 + bitval y)) =
-	      \<one> # rev ys @ [y]" .
+              \<one> # rev ys @ [y]" .
         qed
       qed
     qed
@@ -1103,7 +1103,7 @@
     proof (simp add: norm_signed_Cons,safe)
       assume "norm_unsigned w' = []"
       with weq and w0 show False
-	by (simp add: norm_empty_bv_to_nat_zero)
+        by (simp add: norm_empty_bv_to_nat_zero)
     next
       assume w'0: "norm_unsigned w' \<noteq> []"
       have "0 < bv_to_nat w'"
@@ -1997,7 +1997,7 @@
         have "\<forall>n. n < length ys --> ys ! (length ys - Suc n) = rev ys ! n"
           by simp
         hence "n < length ys --> ys ! (length ys - Suc n) = rev ys ! n" ..
-	from this and noty
+        from this and noty
         have "ys ! (length ys - Suc n) = rev ys ! n" ..
         thus "(y # ys) ! (length ys - n) = rev ys ! n"
           by (simp add: nth_Cons' noty linorder_not_less [symmetric])
--- a/src/HOL/Library/Zorn.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/Zorn.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,7 +1,8 @@
-(*  Title       : HOL/Library/Zorn.thy
-    Author      : Jacques D. Fleuriot, Tobias Nipkow
-    Description : Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF)
-                  The well-ordering theorem
+(*  Title:      HOL/Library/Zorn.thy
+    Author:     Jacques D. Fleuriot, Tobias Nipkow
+
+Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF).
+The well-ordering theorem.
 *)
 
 header {* Zorn's Lemma *}
@@ -289,7 +290,7 @@
       assume "a \<in> Field r" "?B a \<in> C" "b \<in> Field r" "?B b \<in> C"
       hence "?B a \<subseteq> ?B b \<or> ?B b \<subseteq> ?B a" using 2 by auto
       thus "(a, b) \<in> r \<or> (b, a) \<in> r" using `Preorder r` `a:Field r` `b:Field r`
-	by(simp add:subset_Image1_Image1_iff)
+        by (simp add:subset_Image1_Image1_iff)
     qed
     then obtain u where uA: "u:Field r" "\<forall>a\<in>?A. (a,u) : r" using u by auto
     have "\<forall>A\<in>C. A \<subseteq> r^-1 `` {u}" (is "?P u")
@@ -297,7 +298,7 @@
       fix a B assume aB: "B:C" "a:B"
       with 1 obtain x where "x:Field r" "B = r^-1 `` {x}" by auto
       thus "(a,u) : r" using uA aB `Preorder r`
-	by (auto simp add: preorder_on_def refl_on_def) (metis transD)
+        by (auto simp add: preorder_on_def refl_on_def) (metis transD)
     qed
     thus "EX u:Field r. ?P u" using `u:Field r` by blast
   qed
@@ -381,9 +382,9 @@
     next
       case (Suc i)
       moreover obtain s where "s\<in>R" and "(f(Suc(Suc i)), f(Suc i)) \<in> s"
-	using 1 by auto
+        using 1 by auto
       moreover hence "s initial_segment_of r \<or> r initial_segment_of s"
-	using assms(1) `r:R` by(simp add: Chain_def)
+        using assms(1) `r:R` by(simp add: Chain_def)
       ultimately show ?case by(simp add:init_seg_of_def) blast
     qed
   }
@@ -452,9 +453,9 @@
     proof
       assume "m={}"
       moreover have "Well_order {(x,x)}"
-	by(simp add:order_on_defs refl_on_def trans_def antisym_def total_on_def Field_def Domain_def Range_def)
+        by(simp add:order_on_defs refl_on_def trans_def antisym_def total_on_def Field_def Domain_def Range_def)
       ultimately show False using max
-	by (auto simp:I_def init_seg_of_def simp del:Field_insert)
+        by (auto simp:I_def init_seg_of_def simp del:Field_insert)
     qed
     hence "Field m \<noteq> {}" by(auto simp:Field_def)
     moreover have "wf(m-Id)" using `Well_order m`
@@ -476,10 +477,10 @@
     moreover have "wf(?m-Id)"
     proof-
       have "wf ?s" using `x \<notin> Field m`
-	by(auto simp add:wf_eq_minimal Field_def Domain_def Range_def) metis
+        by(auto simp add:wf_eq_minimal Field_def Domain_def Range_def) metis
       thus ?thesis using `wf(m-Id)` `x \<notin> Field m`
-	wf_subset[OF `wf ?s` Diff_subset]
-	by (fastsimp intro!: wf_Un simp add: Un_Diff Field_def)
+        wf_subset[OF `wf ?s` Diff_subset]
+        by (fastsimp intro!: wf_Un simp add: Un_Diff Field_def)
     qed
     ultimately have "Well_order ?m" by(simp add:order_on_defs)
 --{*We show that the extension is above m*}
--- a/src/HOL/Library/reflection.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Library/reflection.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -48,22 +48,22 @@
    val tys = map fastype_of fterms
    val vs = map Free (xs ~~ tys)
    val env = fterms ~~ vs
-		    (* FIXME!!!!*)
+                    (* FIXME!!!!*)
    fun replace_fterms (t as t1 $ t2) =
        (case AList.lookup (op aconv) env t of
-	    SOME v => v
-	  | NONE => replace_fterms t1 $ replace_fterms t2)
+            SOME v => v
+          | NONE => replace_fterms t1 $ replace_fterms t2)
      | replace_fterms t = (case AList.lookup (op aconv) env t of
-			       SOME v => v
-			     | NONE => t)
+                               SOME v => v
+                             | NONE => t)
 
    fun mk_def (Abs(x,xT,t),v) = HOLogic.mk_Trueprop ((HOLogic.all_const xT)$ Abs(x,xT,HOLogic.mk_eq(v$(Bound 0), t)))
      | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t))
    fun tryext x = (x RS ext2 handle THM _ =>  x)
    val cong = (Goal.prove ctxt'' [] (map mk_def env)
-			  (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
-			  (fn x => LocalDefs.unfold_tac (#context x) (map tryext (#prems x))
-							THEN rtac th' 1)) RS sym
+                          (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
+                          (fn x => LocalDefs.unfold_tac (#context x) (map tryext (#prems x))
+                                                        THEN rtac th' 1)) RS sym
 
    val (cong' :: vars') =
        Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs)
@@ -134,15 +134,15 @@
                  val (xn,ta) = variant_abs (xn,xT,ta)
                  val x = Free(xn,xT)
                  val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT)
-		          of NONE => error "tryabsdecomp: Type not found in the Environement"
+                          of NONE => error "tryabsdecomp: Type not found in the Environement"
                            | SOME (bsT,atsT) =>
                              (AList.update Type.could_unify (HOLogic.listT xT, ((x::bsT), atsT)) bds))
                in (([(ta, ctxt')],
                     fn ([th], bds) =>
                       (hd (Variable.export ctxt' ctxt [(forall_intr (cert x) th) COMP allI]),
                        let val (bsT,asT) = the(AList.lookup Type.could_unify bds (HOLogic.listT xT))
-		       in AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) bds
-		       end)),
+                       in AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) bds
+                       end)),
                    bds)
                end)
            | _ => da (s,ctxt) bds)
@@ -157,9 +157,9 @@
                 (Vartab.empty, Vartab.empty)
             val (fnvs,invs) = List.partition (fn ((vn,_),_) => vn mem vns) (Vartab.dest tmenv)
             val (fts,its) =
-	      (map (snd o snd) fnvs,
+              (map (snd o snd) fnvs,
                map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) invs)
-	    val ctyenv = map (fn ((vn,vi),(s,ty)) => (certy (TVar((vn,vi),s)), certy ty)) (Vartab.dest tyenv)
+            val ctyenv = map (fn ((vn,vi),(s,ty)) => (certy (TVar((vn,vi),s)), certy ty)) (Vartab.dest tyenv)
           in ((fts ~~ (replicate (length fts) ctxt),
                Library.apfst (FWD (instantiate (ctyenv, its) cong))), bds)
           end)
@@ -174,9 +174,9 @@
           let
             val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
           in exists_Const
-	    (fn (n,ty) => n = @{const_name "List.nth"}
+            (fn (n,ty) => n = @{const_name "List.nth"}
                           andalso
-			  AList.defined Type.could_unify bds (domain_type ty)) rhs
+                          AList.defined Type.could_unify bds (domain_type ty)) rhs
             andalso Type.could_unify (fastype_of rhs, tT)
           end
 
@@ -266,7 +266,7 @@
 
         val bds = AList.make (fn _ => ([],[])) tys
         val eqs = map (fn eq => eq |> prop_of |> HOLogic.dest_Trueprop
-  	                           |> HOLogic.dest_eq |> fst |> strip_comb |> snd |> tl
+                                   |> HOLogic.dest_eq |> fst |> strip_comb |> snd |> tl
                                    |> (insteq eq)) raw_eqs
         val (ps,congs) = split_list (map (mk_congeq ctxt' fs) eqs)
       in (ps ~~ (Variable.export ctxt' ctxt congs), bds)
@@ -278,14 +278,14 @@
     fun is_listVar (Var (_,t)) = can dest_listT t
          | is_listVar _ = false
     val vars = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
-	          |> strip_comb |> snd |> filter is_listVar
+                  |> strip_comb |> snd |> filter is_listVar
     val cert = cterm_of (ProofContext.theory_of ctxt)
     val cvs = map (fn (v as Var(n,t)) => (cert v,
                   the (AList.lookup Type.could_unify bds t) |> snd |> HOLogic.mk_list (dest_listT t) |> cert)) vars
     val th' = instantiate ([], cvs) th
     val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th'
     val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t')))
-	       (fn _ => simp_tac (simpset_of ctxt) 1)
+               (fn _ => simp_tac (simpset_of ctxt) 1)
   in FWD trans [th'',th']
   end
 
--- a/src/HOL/List.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/List.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -3935,15 +3935,15 @@
       assume L: ?L
       show ?R
       proof clarify
-	fix n assume n: "n : {i..<j}"
-	show "P n"
-	proof cases
-	  assume "n = i" thus "P n" using L by simp
-	next
-	  assume "n ~= i"
-	  hence "i+1 <= n" using n by auto
-	  thus "P n" using L n by simp
-	qed
+        fix n assume n: "n : {i..<j}"
+        show "P n"
+        proof cases
+          assume "n = i" thus "P n" using L by simp
+        next
+          assume "n ~= i"
+          hence "i+1 <= n" using n by auto
+          thus "P n" using L n by simp
+        qed
       qed
     next
       assume R: ?R thus ?L using `?yes` 1 by auto
@@ -3990,15 +3990,15 @@
       assume L: ?L
       show ?R
       proof clarify
-	fix n assume n: "n : {i..j}"
-	show "P n"
-	proof cases
-	  assume "n = i" thus "P n" using L by simp
-	next
-	  assume "n ~= i"
-	  hence "i+1 <= n" using n by auto
-	  thus "P n" using L n by simp
-	qed
+        fix n assume n: "n : {i..j}"
+        show "P n"
+        proof cases
+          assume "n = i" thus "P n" using L by simp
+        next
+          assume "n ~= i"
+          hence "i+1 <= n" using n by auto
+          thus "P n" using L n by simp
+        qed
       qed
     next
       assume R: ?R thus ?L using `?yes` 1 by auto
--- a/src/HOL/Matrix/Matrix.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Matrix/Matrix.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -88,7 +88,7 @@
       fix y
       assume hyp: "Rep_matrix x (fst y) (snd y) \<noteq> 0"
       show "\<exists> a b. (Rep_matrix x b a \<noteq> 0 & y = (b,a))"
-	by (rule exI[of _ "snd y"], rule exI[of _ "fst y"]) (simp add: hyp)
+        by (rule exI[of _ "snd y"], rule exI[of _ "fst y"]) (simp add: hyp)
     qed
   then have "finite (?swap`?A)"
     proof -
--- a/src/HOL/Matrix/cplex/fspmlp.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Matrix/cplex/fspmlp.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -263,13 +263,13 @@
 
         val g = propagate_sure_bounds safe_propagation names g
 
-	val (r1, r2) = split_graph g
+        val (r1, r2) = split_graph g
 
         fun add_row_entry m index f vname value =
             let
-		val v = (case value of 
-			     SOME value => FloatSparseMatrixBuilder.mk_spvec_entry 0 value
-			   | NONE => FloatSparseMatrixBuilder.mk_spvec_entry' 0 (f $ (Var ((vname,0), HOLogic.realT))))
+                val v = (case value of 
+                             SOME value => FloatSparseMatrixBuilder.mk_spvec_entry 0 value
+                           | NONE => FloatSparseMatrixBuilder.mk_spvec_entry' 0 (f $ (Var ((vname,0), HOLogic.realT))))
                 val vec = cons_spvec v empty_spvec
             in
                 cons_spmat (FloatSparseMatrixBuilder.mk_spmat_entry index vec) m
@@ -286,7 +286,7 @@
                     val b2 = Inttab.lookup r2 index
                 in
                     (add_row_entry r12_1 index @{term "lbound :: real => real"} ((names index)^"l") b1, 
-		     add_row_entry r12_2 index @{term "ubound :: real => real"} ((names index)^"u") b2)
+                     add_row_entry r12_2 index @{term "ubound :: real => real"} ((names index)^"u") b2)
                 end
 
         val (r1, r2) = abs_estimate xlen r1 r2
@@ -300,7 +300,7 @@
         val prog = Cplex.load_cplexFile filename
         val prog = Cplex.elim_nonfree_bounds prog
         val prog = Cplex.relax_strict_ineqs prog
-        val (maximize, c, A, b, (xlen, names, _)) = CplexFloatSparseMatrixConverter.convert_prog prog			    
+        val (maximize, c, A, b, (xlen, names, _)) = CplexFloatSparseMatrixConverter.convert_prog prog                       
         val (r1, r2) = calcr safe_propagation xlen names prec A b
         val _ = if maximize then () else raise Load "sorry, cannot handle minimization problems"
         val (dualprog, indexof) = FloatSparseMatrixBuilder.dual_cplexProg c A b
--- a/src/HOL/Matrix/cplex/matrixlp.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Matrix/cplex/matrixlp.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -18,7 +18,7 @@
 
 fun inst_real thm =
   let val certT = ctyp_of (Thm.theory_of_thm thm) in
-    standard (Thm.instantiate
+    Drule.standard (Thm.instantiate
       ([(certT (TVar (hd (OldTerm.term_tvars (prop_of thm)))), certT HOLogic.realT)], []) thm)
   end
 
@@ -59,7 +59,7 @@
         val ord = prod_ord (prod_ord string_ord int_ord) (list_ord string_ord)
         val v = TVar (hd (sort ord (OldTerm.term_tvars (prop_of thm))))
     in
-        standard (Thm.instantiate ([(certT v, certT ty)], []) thm)
+        Drule.standard (Thm.instantiate ([(certT v, certT ty)], []) thm)
     end
 
 fun inst_tvars [] thms = thms
--- a/src/HOL/MetisExamples/Message.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MetisExamples/Message.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,7 +1,7 @@
 (*  Title:      HOL/MetisTest/Message.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 
-Testing the metis method
+Testing the metis method.
 *)
 
 theory Message imports Main begin
@@ -34,13 +34,13 @@
   agent = Server | Friend nat | Spy
 
 datatype
-     msg = Agent  agent	    --{*Agent names*}
+     msg = Agent  agent     --{*Agent names*}
          | Number nat       --{*Ordinary integers, timestamps, ...*}
          | Nonce  nat       --{*Unguessable nonces*}
          | Key    key       --{*Crypto keys*}
-	 | Hash   msg       --{*Hashing*}
-	 | MPair  msg msg   --{*Compound messages*}
-	 | Crypt  key msg   --{*Encryption, public- or shared-key*}
+         | Hash   msg       --{*Hashing*}
+         | MPair  msg msg   --{*Compound messages*}
+         | Crypt  key msg   --{*Encryption, public- or shared-key*}
 
 
 text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*}
--- a/src/HOL/MetisExamples/Tarski.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MetisExamples/Tarski.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,7 @@
 (*  Title:      HOL/MetisTest/Tarski.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 
-Testing the metis method
+Testing the metis method.
 *)
 
 header {* The Full Theorem of Tarski *}
@@ -556,17 +555,17 @@
   Collect
    (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r)) (COMBC op \<in> A))"
   assume 1: "f (lub (Collect
-	   (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
-	     (COMBC op \<in> A)))
+           (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
+             (COMBC op \<in> A)))
       cl) \<noteq>
   lub (Collect
-	(COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
-	  (COMBC op \<in> A)))
+        (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
+          (COMBC op \<in> A)))
    cl"
   have 2: "f (lub H cl) \<noteq>
   lub (Collect
-	(COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
-	  (COMBC op \<in> A)))
+        (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
+          (COMBC op \<in> A)))
    cl"
     by (metis 1 0)
   have 3: "f (lub H cl) \<noteq> lub H cl"
--- a/src/HOL/MicroJava/BV/BVNoTypeError.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/BV/BVNoTypeError.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/BV/BVNoTypeErrors.thy
-    ID:         $Id$
     Author:     Gerwin Klein
 *)
 
@@ -294,7 +293,7 @@
       ultimately show ?thesis using Getfield field "class" stk hconf wf
         apply clarsimp
         apply (fastsimp intro: wf_prog_ws_prog
-	  dest!: hconfD widen_cfs_fields oconf_objD)
+          dest!: hconfD widen_cfs_fields oconf_objD)
         done
     next
       case (Putfield F C)
--- a/src/HOL/MicroJava/BV/BVSpecTypeSafe.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/BV/BVSpecTypeSafe.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/BV/BVSpecTypeSafe.thy
-    ID:         $Id$
     Author:     Cornelia Pusch, Gerwin Klein
     Copyright   1999 Technische Universitaet Muenchen
 *)
@@ -1332,7 +1331,7 @@
   apply (auto simp add: fun_upd_apply blank_def oconf_def split: split_if_asm)
   apply (simp add: fields_is_type 
           [OF _ wf [THEN wf_prog_ws_prog] 
-	        is_class_xcpt [OF wf [THEN wf_prog_ws_prog]]])+
+                is_class_xcpt [OF wf [THEN wf_prog_ws_prog]]])+
   done
     
 lemma 
--- a/src/HOL/MicroJava/BV/JVM.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/BV/JVM.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/BV/JVM.thy
-    ID:         $Id$
     Author:     Tobias Nipkow, Gerwin Klein
     Copyright   2000 TUM
 *)
@@ -37,7 +36,7 @@
   apply (rule is_bcv_kildall)
        apply (simp (no_asm) add: sl_triple_conv [symmetric]) 
        apply (force intro!: semilat_JVM_slI dest: wf_acyclic 
-	 simp add: symmetric sl_triple_conv)
+         simp add: symmetric sl_triple_conv)
       apply (simp (no_asm) add: JVM_le_unfold)
       apply (blast intro!: order_widen wf_converse_subcls1_impl_acc_subtype
                    dest: wf_subcls1 wfP_acyclicP wf_prog_ws_prog)
--- a/src/HOL/MicroJava/Comp/CorrComp.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/Comp/CorrComp.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/Comp/CorrComp.thy
-    ID:         $Id$
     Author:     Martin Strecker
 *)
 
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -622,7 +622,7 @@
 prefer 2 apply (rule_tac x="pc - length mt_pre" in exI, arith)
 
 apply (drule_tac x=pc'' in spec)
-apply (drule mp) apply arith	(* show pc'' < length mt *)
+apply (drule mp) apply arith    (* show pc'' < length mt *)
 apply clarify
 
 apply (rule conjI)
@@ -1398,7 +1398,7 @@
     apply (rule max_of_list_sublist)
     apply (simp (no_asm_simp) only: set_append set.simps map.simps) apply blast
   apply (simp (no_asm_simp))
-  apply simp			(* subgoal bc3 = [] *)
+  apply simp                    (* subgoal bc3 = [] *)
   apply (simp add: comb_nil_def) (* subgoal mt3 = [] \<and> sttp2 = sttp3 *)
 
   (* case start_sttp_resp_cons f3 *)
@@ -1440,7 +1440,7 @@
 apply (rule check_type_mono, assumption)
 apply (simp only: max_ssize_def) apply (rule max_of_list_sublist) apply (simp (no_asm_simp))
 apply blast
-  apply simp			(* subgoal bc3 = [] *)
+  apply simp                    (* subgoal bc3 = [] *)
   apply (simp add: comb_nil_def) (* subgoal mt3 = [] \<and> sttp2 = sttp3 *)
 
 
@@ -1735,7 +1735,7 @@
   apply (simp only: compTpExpr_LT_ST)+
   apply (simp add: eff_def norm_eff_def popST_def pushST_def mt_sttp_flatten_def)
   apply (case_tac Ta) apply (simp (no_asm_simp)) apply (simp (no_asm_simp))
-  apply (rule contracting_popST)		(* contracting (popST 2)  *)
+  apply (rule contracting_popST)                (* contracting (popST 2)  *)
 
 apply (drule_tac ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2 @ [Ifcmpeq 3]"
   and ?bc2.0 = "[LitPush (Bool False)]" 
@@ -1766,7 +1766,7 @@
   apply (simp (no_asm_simp) add: length_compTpExpr)
   apply (simp only: compTpExpr_LT_ST)+
   apply (simp add: eff_def norm_eff_def popST_def pushST_def)
-  apply (rule contracting_popST)		(* contracting (popST 1) *)
+  apply (rule contracting_popST)                (* contracting (popST 1) *)
 
 apply (drule_tac 
   ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2 @ [Ifcmpeq 3, LitPush (Bool False), Goto 2]" 
@@ -2282,7 +2282,7 @@
 apply (rule_tac vn=vn and ty=ty in wt_method_compTpInit_corresp)
 apply assumption+
 apply (simp (no_asm_simp))+ 
-apply (simp add: wf_java_mdecl_def)	(* is_type G ty *)
+apply (simp add: wf_java_mdecl_def)     (* is_type G ty *)
 apply (simp add: compTpInit_def storeST_def pushST_def)
 apply simp
 done
--- a/src/HOL/MicroJava/Comp/Index.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/Comp/Index.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/Comp/Index.thy
-    ID:         $Id$
     Author:     Martin Strecker
 *)
 
@@ -105,7 +104,7 @@
 (* This corresponds to the original def in wf_java_mdecl:
   "disjoint_varnames pns lvars \<equiv> 
   nodups pns \<and> unique lvars \<and> This \<notin> set pns \<and> This \<notin> set (map fst lvars) \<and> 
-	(\<forall>pn\<in>set pns. map_of lvars pn = None)"
+        (\<forall>pn\<in>set pns. map_of lvars pn = None)"
 *)
 
   "disjoint_varnames pns lvars \<equiv> 
--- a/src/HOL/MicroJava/Comp/LemmasComp.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/Comp/LemmasComp.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/Comp/LemmasComp.thy
-    ID:         $Id$
     Author:     Martin Strecker
 *)
 
@@ -56,13 +55,13 @@
     proof 
       assume as: "fst a \<in> fst ` set list" 
       then obtain x where fst_a_x: "x\<in>set list \<and> fst a = fst x" 
-	by (auto simp add: image_iff)
+        by (auto simp add: image_iff)
       then have "fst (f a) = fst (f x)" by (simp add: fst_eq)
       with as show "(fst (f a) \<in> fst ` f ` set list)" by (simp add: fst_a_x)
     next
       assume as: "fst (f a) \<in> fst ` f ` set list"
       then obtain x where fst_a_x: "x\<in>set list \<and> fst (f a) = fst (f x)"
-	by (auto simp add: image_iff)
+        by (auto simp add: image_iff)
       then have "fst a = fst x" by (simp add: fst_eq)
       with as show "fst a \<in> fst ` set list" by (simp add: fst_a_x)
     qed
@@ -307,7 +306,7 @@
 apply (rule_tac f="(\<lambda>(s, m). (s, Object, m))" and
   g="(Fun.comp (\<lambda>(s, m). (s, Object, m)) (compMethod G Object))" 
   in map_of_map_fst)
-defer				(* inj \<dots> *)
+defer                           (* inj \<dots> *)
 apply (simp add: inj_on_def split_beta) 
 apply (simp add: split_beta compMethod_def)
 apply (simp add: map_of_map [THEN sym])
--- a/src/HOL/MicroJava/Comp/TranslComp.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/Comp/TranslComp.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/Comp/TranslComp.thy
-    ID:         $Id$
     Author:     Martin Strecker
 *)
 
@@ -69,30 +68,30 @@
 primrec
 (** code generation for statements **)
 
-"compStmt jmb Skip = []"			
+"compStmt jmb Skip = []"                        
 
 "compStmt jmb (Expr e) = ((compExpr jmb e) @ [Pop])"
 
 "compStmt jmb (c1;; c2) = ((compStmt jmb c1) @ (compStmt jmb c2))"
 
 "compStmt jmb (If(e) c1 Else c2) =
-	(let cnstf = LitPush (Bool False);
+        (let cnstf = LitPush (Bool False);
              cnd   = compExpr jmb e;
-	     thn   = compStmt jmb c1;
-	     els   = compStmt jmb c2;
-	     test  = Ifcmpeq (int(length thn +2)); 
-	     thnex = Goto (int(length els +1))
-	 in
-	 [cnstf] @ cnd @ [test] @ thn @ [thnex] @ els)"
+             thn   = compStmt jmb c1;
+             els   = compStmt jmb c2;
+             test  = Ifcmpeq (int(length thn +2)); 
+             thnex = Goto (int(length els +1))
+         in
+         [cnstf] @ cnd @ [test] @ thn @ [thnex] @ els)"
 
 "compStmt jmb (While(e) c) =
-	(let cnstf = LitPush (Bool False);
+        (let cnstf = LitPush (Bool False);
              cnd   = compExpr jmb e;
-	     bdy   = compStmt jmb c;
-	     test  = Ifcmpeq (int(length bdy +2)); 
-	     loop  = Goto (-(int((length bdy) + (length cnd) +2)))
-	 in
-	 [cnstf] @ cnd @ [test] @ bdy @ [loop])"
+             bdy   = compStmt jmb c;
+             test  = Ifcmpeq (int(length bdy +2)); 
+             loop  = Goto (-(int((length bdy) + (length cnd) +2)))
+         in
+         [cnstf] @ cnd @ [test] @ bdy @ [loop])"
 
 (**********************************************************************)
 
--- a/src/HOL/MicroJava/Comp/TranslCompTp.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/Comp/TranslCompTp.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/Comp/TranslCompTp.thy
-    ID:         $Id$
     Author:     Martin Strecker
 *)
 
@@ -43,8 +42,8 @@
 (**********************************************************************)
 
 syntax
-  mt_of		:: "method_type \<times> state_type \<Rightarrow> method_type"
-  sttp_of	:: "method_type \<times> state_type \<Rightarrow> state_type"
+  mt_of         :: "method_type \<times> state_type \<Rightarrow> method_type"
+  sttp_of       :: "method_type \<times> state_type \<Rightarrow> state_type"
 
 translations
   "mt_of"     => "fst"
--- a/src/HOL/MicroJava/J/WellForm.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/MicroJava/J/WellForm.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/J/WellForm.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   1999 Technische Universitaet Muenchen
 *)
@@ -585,7 +584,7 @@
 apply (simp (no_asm_simp) add: wf_prog_ws_prog)
 
 apply (case_tac "Da=C")
-apply blast			(* case Da=C *)
+apply blast                     (* case Da=C *)
 
 apply (subgoal_tac "((vn, Da), T) \<in> set (fields (G, D))") apply blast
 apply (erule thin_rl)
--- a/src/HOL/Modelcheck/MuckeSyn.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Modelcheck/MuckeSyn.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -76,17 +76,17 @@
 
 (* extended for treatment of nu (TH) *)
 fun search_mu ((Const("MuCalculus.mu",tp)) $ t2) = 
-	(case (search_mu t2) of
-	      SOME t => SOME t 
-	    | NONE => SOME ((Const("MuCalculus.mu",tp)) $ t2))
+        (case (search_mu t2) of
+              SOME t => SOME t 
+            | NONE => SOME ((Const("MuCalculus.mu",tp)) $ t2))
   | search_mu ((Const("MuCalculus.nu",tp)) $ t2) =
         (case (search_mu t2) of
               SOME t => SOME t
             | NONE => SOME ((Const("MuCalculus.nu",tp)) $ t2))
   | search_mu (t1 $ t2) = 
-	(case (search_mu t1) of
-	      SOME t => SOME t 
-	    | NONE     => search_mu t2)
+        (case (search_mu t1) of
+              SOME t => SOME t 
+            | NONE     => search_mu t2)
   | search_mu (Abs(_,_,t)) = search_mu t
   | search_mu _ = NONE;
 
@@ -99,12 +99,12 @@
 fun search_var s t =
 case t of
      t1 $ t2 => (case (search_var s t1) of
-		             SOME tt => SOME tt |
-			     NONE => search_var s t2) |
+                             SOME tt => SOME tt |
+                             NONE => search_var s t2) |
      Abs(_,_,t) => search_var s t |
      Var((s1,_),_) => if s = s1 then SOME t else NONE |
      _ => NONE;
-	
+        
 
 (* function move_mus:
    Mucke can't deal with nested Mu terms. move_mus i searches for 
@@ -119,8 +119,8 @@
 local
 
   val move_thm = OldGoals.prove_goal @{theory} "[| a = b ==> P a; a = b |] ==> P b"
-	(fn prems => [cut_facts_tac prems 1, dtac sym 1, hyp_subst_tac 1,
-		     REPEAT (resolve_tac prems 1)]);
+        (fn prems => [cut_facts_tac prems 1, dtac sym 1, hyp_subst_tac 1,
+                     REPEAT (resolve_tac prems 1)]);
 
   val sig_move_thm = Thm.theory_of_thm move_thm;
   val bCterm = cterm_of sig_move_thm (valOf (search_var "b" (concl_of move_thm)));
@@ -134,20 +134,20 @@
     val concl = Logic.strip_imp_concl subgoal; (* recursive mu's in prems? *)
     val redex = search_mu concl;
     val idx = let val t = #maxidx (rep_thm state) in 
-	      if t < 0 then 1 else t+1 end;
+              if t < 0 then 1 else t+1 end;
 in
 case redex of
      NONE => all_tac state |
      SOME redexterm => 
-	let val Credex = cterm_of sign redexterm;
-	    val aiCterm = 
-		cterm_of sig_move_thm (Logic.incr_indexes ([],idx) (term_of aCterm));
-	    val inst_move_thm = cterm_instantiate 
-				[(bCterm,Credex),(aCterm,aiCterm)] move_thm;
-	in
+        let val Credex = cterm_of sign redexterm;
+            val aiCterm = 
+                cterm_of sig_move_thm (Logic.incr_indexes ([],idx) (term_of aCterm));
+            val inst_move_thm = cterm_instantiate 
+                                [(bCterm,Credex),(aCterm,aiCterm)] move_thm;
+        in
             ((rtac inst_move_thm i) THEN (dtac eq_reflection i) 
-		THEN (move_mus i)) state
-	end
+                THEN (move_mus i)) state
+        end
 end;
 end;
 
@@ -192,12 +192,12 @@
   | mk_lam_def [] ((Const("==",_) $ (Const _)) $ RHS) t = SOME t
   | mk_lam_def [] ((Const("==",_) $ LHS) $ RHS) t = 
     let val thy = theory_of_thm t;
-	val fnam = Syntax.string_of_term_global thy (getfun LHS);
-	val rhs = Syntax.string_of_term_global thy (freeze_thaw RHS)
-	val gl = delete_bold_string (fnam ^" == % " ^ (getargs LHS) ^" . " ^ rhs);
+        val fnam = Syntax.string_of_term_global thy (getfun LHS);
+        val rhs = Syntax.string_of_term_global thy (freeze_thaw RHS)
+        val gl = delete_bold_string (fnam ^" == % " ^ (getargs LHS) ^" . " ^ rhs);
     in
-	SOME (OldGoals.prove_goal thy gl (fn prems =>
-  		[(REPEAT (rtac ext_rl 1)), (rtac t 1) ]))
+        SOME (OldGoals.prove_goal thy gl (fn prems =>
+                [(REPEAT (rtac ext_rl 1)), (rtac t 1) ]))
     end
 | mk_lam_def [] _ t= NONE; 
 
--- a/src/HOL/Modelcheck/mucke_oracle.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Modelcheck/mucke_oracle.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -46,7 +46,7 @@
   if (s="bool") then [] else [b,a1,a2]
   end |
   contains_if (a $ b) = if ((contains_if b)=[]) then (contains_if a)
-		        else (contains_if b) |
+                        else (contains_if b) |
   contains_if _ = [];
 
   fun find_replace_term (Abs(a,T,t)) = find_replace_term (snd(Syntax.variant_abs(a,T,t))) |
@@ -58,7 +58,7 @@
   fun if_substi (Abs(a,T,t)) trm = Abs(a,T,t) |
   if_substi (Const("HOL.If",T) $ b $ a1 $ a2) t = t |
   if_substi (a $ b) t = if ((contains_if b)=[]) then ((if_substi a t)$b)
-		        else (a$(if_substi b t)) |
+                        else (a$(if_substi b t)) |
   if_substi t v = t;
 
   fun deliver_term (t,[]) = [] |
@@ -97,8 +97,8 @@
 let val sign = Thm.theory_of_thm state;
         val (subgoal::_) = Library.drop(i-1,prems_of state);
         val prems = Logic.strip_imp_prems subgoal;
-	val concl = Logic.strip_imp_concl subgoal;
-	val prems = prems @ [concl];
+        val concl = Logic.strip_imp_concl subgoal;
+        val prems = prems @ [concl];
         val itrm = search_ifs prems;
 in
 if (itrm = []) then no_tac state else
@@ -209,7 +209,7 @@
         val (x,y) = snd(ispair a)
        in
         con_term_list_of (Const("pair",Type("fun",[x,Type("fun",[y,a])])))
-			 (arglist_of sg tl [x,y])
+                         (arglist_of sg tl [x,y])
        end)
  else
  (let
@@ -219,7 +219,7 @@
                         val trm = OldGoals.simple_read_term sg ft c;
                         in
                         (con_term_list_of trm (arglist_of sg tl tyl))
-			@(deliver_erg sg tl typ r)
+                        @(deliver_erg sg tl typ r)
                         end;
   val cl = search_in_keylist tl a;
   in
@@ -239,7 +239,7 @@
 (*
 fun force_Abs (Abs s_T_t) = Abs s_T_t
   | force_Abs t = Abs("x", hd(fst(strip_type (type_of t))),
-		      (incr_boundvars 1 t) $ (Bound 0));
+                      (incr_boundvars 1 t) $ (Bound 0));
 
 fun etaexp_dest_Abs t = dest_Abs (force_Abs t);
 *)
@@ -269,7 +269,7 @@
 
 fun make_MuTerm muDeclTerm ParamDeclTerm muTerm isaType =
   let val constMu = Const("_mu",
-			  make_fun_type [isaType,isaType,isaType,isaType]);
+                          make_fun_type [isaType,isaType,isaType,isaType]);
       val t1 = constMu $ muDeclTerm;
       val t2 = t1 $ ParamDeclTerm;
       val t3 = t2 $  muTerm
@@ -327,15 +327,15 @@
   | get_decls sign Clist trm = (Clist,trm);
 
 fun make_mu_def sign ((tt $ LHS) $ (ttt $ RHS)) =
-  let 	val LHSStr = fst (dest_atom LHS);
-	val MuType = Type("bool",[]); (* always ResType of mu, also serves
-					 as dummy type *)
-	val (_,_,PAbs) = dest_Abs (RHS); (* RHS is %Q. ... *)
-  	val (PCon_LHS,MMuTerm) = get_decls sign [] (subst_bound (LHS,PAbs));
-	val PConsts = rev PCon_LHS;
-	val muDeclTerm = make_decl "bool" LHSStr MuType;
-	val PDeclsTerm = make_decls sign MuType PConsts;
-	val MuDefTerm = make_MuTerm muDeclTerm PDeclsTerm MMuTerm MuType;		
+  let   val LHSStr = fst (dest_atom LHS);
+        val MuType = Type("bool",[]); (* always ResType of mu, also serves
+                                         as dummy type *)
+        val (_,_,PAbs) = dest_Abs (RHS); (* RHS is %Q. ... *)
+        val (PCon_LHS,MMuTerm) = get_decls sign [] (subst_bound (LHS,PAbs));
+        val PConsts = rev PCon_LHS;
+        val muDeclTerm = make_decl "bool" LHSStr MuType;
+        val PDeclsTerm = make_decls sign MuType PConsts;
+        val MuDefTerm = make_MuTerm muDeclTerm PDeclsTerm MMuTerm MuType;               
   in MuDefTerm end;
 
 fun make_mu_decl sign ((tt $ LHS) $ (ttt $ RHS)) =
@@ -375,11 +375,11 @@
   in NuDeclTerm end;
 
 fun make_FunMuckeTerm FunDeclTerm ParamDeclTerm Term isaType =
-  let 	val constFun = Const("_fun",
-			    make_fun_type [isaType,isaType,isaType,isaType]);
-      	val t1 = constFun $ FunDeclTerm;
-      	val t2 = t1 $ ParamDeclTerm;
-      	val t3 = t2 $  Term
+  let   val constFun = Const("_fun",
+                            make_fun_type [isaType,isaType,isaType,isaType]);
+        val t1 = constFun $ FunDeclTerm;
+        val t2 = t1 $ ParamDeclTerm;
+        val t3 = t2 $  Term
   in t3 end;
 
 fun make_FunMuckeDecl FunDeclTerm ParamDeclTerm isaType =
@@ -394,26 +394,26 @@
 | is_fundef _ = false; 
 
 fun make_mucke_fun_def sign ((_ $ LHS) $ RHS) =
-  let 	(* fun dest_atom (Const t) = dest_Const (Const t)
+  let   (* fun dest_atom (Const t) = dest_Const (Const t)
           | dest_atom (Free t)  = dest_Free (Free t); *)
-	val LHSStr = fst (dest_atom LHS);
-	val LHSResType = body_type(snd(dest_atom LHS));
-	val LHSResTypeStr = type_to_string_OUTPUT LHSResType;
-(*	val (_,AbsType,RawTerm) = dest_Abs(RHS);
-*)	val (Consts_LHS_rev,Free_RHS) = get_decls sign [] RHS;
-	val Consts_LHS = rev Consts_LHS_rev;
-	val PDeclsTerm = make_decls sign LHSResType Consts_LHS; 
-		(* Boolean functions only, list necessary in general *)
-	val DeclTerm = make_decl LHSResTypeStr LHSStr LHSResType;
-	val MuckeDefTerm = make_FunMuckeTerm DeclTerm PDeclsTerm Free_RHS
-					 LHSResType;	
+        val LHSStr = fst (dest_atom LHS);
+        val LHSResType = body_type(snd(dest_atom LHS));
+        val LHSResTypeStr = type_to_string_OUTPUT LHSResType;
+(*      val (_,AbsType,RawTerm) = dest_Abs(RHS);
+*)      val (Consts_LHS_rev,Free_RHS) = get_decls sign [] RHS;
+        val Consts_LHS = rev Consts_LHS_rev;
+        val PDeclsTerm = make_decls sign LHSResType Consts_LHS; 
+                (* Boolean functions only, list necessary in general *)
+        val DeclTerm = make_decl LHSResTypeStr LHSStr LHSResType;
+        val MuckeDefTerm = make_FunMuckeTerm DeclTerm PDeclsTerm Free_RHS
+                                         LHSResType;    
   in MuckeDefTerm end;
 
 fun make_mucke_fun_decl sign ((_ $ LHS) $ RHS) =
   let   (* fun dest_atom (Const t) = dest_Const (Const t)
           | dest_atom (Free t)  = dest_Free (Free t); *)
         val LHSStr = fst (dest_atom LHS);
-	val LHSResType = body_type(snd(dest_atom LHS));
+        val LHSResType = body_type(snd(dest_atom LHS));
         val LHSResTypeStr = type_to_string_OUTPUT LHSResType;
 (*      val (_,AbsType,RawTerm) = dest_Abs(RHS);
 *)      val (Consts_LHS_rev,Free_RHS) = get_decls sign [] RHS;
@@ -421,34 +421,34 @@
         val PDeclsTerm = make_decls sign LHSResType Consts_LHS;
                 (* Boolean functions only, list necessary in general *)
         val DeclTerm = make_decl LHSResTypeStr LHSStr LHSResType;
-	val MuckeDeclTerm = make_FunMuckeDecl DeclTerm PDeclsTerm LHSResType;
+        val MuckeDeclTerm = make_FunMuckeDecl DeclTerm PDeclsTerm LHSResType;
 in MuckeDeclTerm end;
 
 fun elim_quantifications sign ((Const("Ex",_)) $ Abs (str,tp,t)) =
     (let val ExConst = Const("_Ex",make_fun_type [tp,tp,tp,tp]);
-     	 val TypeConst = Const (type_to_string_OUTPUT tp,tp);
-	 val VarAbs = Syntax.variant_abs(str,tp,t);
-	 val BoundConst = Const(fst VarAbs,tp);
-	 val t1 = ExConst $ TypeConst;
-	 val t2 = t1 $ BoundConst;
- 	 val t3 = elim_quantifications sign (snd VarAbs)
+         val TypeConst = Const (type_to_string_OUTPUT tp,tp);
+         val VarAbs = Syntax.variant_abs(str,tp,t);
+         val BoundConst = Const(fst VarAbs,tp);
+         val t1 = ExConst $ TypeConst;
+         val t2 = t1 $ BoundConst;
+         val t3 = elim_quantifications sign (snd VarAbs)
      in t2 $ t3 end)
   |  elim_quantifications sign ((Const("All",_)) $ Abs (str,tp,t)) =
     (let val AllConst = Const("_All",make_fun_type [tp,tp,tp,tp]);
-    	 val TypeConst = Const (type_to_string_OUTPUT tp,tp);
-	 val VarAbs = Syntax.variant_abs(str,tp,t);
-	 val BoundConst = Const(fst VarAbs,tp);
-	 val t1 = AllConst $ TypeConst;
-	 val t2 = t1 $ BoundConst;
-	 val t3 = elim_quantifications sign (snd VarAbs)
+         val TypeConst = Const (type_to_string_OUTPUT tp,tp);
+         val VarAbs = Syntax.variant_abs(str,tp,t);
+         val BoundConst = Const(fst VarAbs,tp);
+         val t1 = AllConst $ TypeConst;
+         val t2 = t1 $ BoundConst;
+         val t3 = elim_quantifications sign (snd VarAbs)
      in t2 $ t3 end)
   | elim_quantifications sign (t1 $ t2) = 
-   	(elim_quantifications sign t1) $ (elim_quantifications sign t2)
+        (elim_quantifications sign t1) $ (elim_quantifications sign t2)
   | elim_quantifications sign (Abs(_,_,t)) = elim_quantifications sign t
   | elim_quantifications sign t = t;
 fun elim_quant_in_list sign [] = []
   | elim_quant_in_list sign (trm::list) = 
-			(elim_quantifications sign trm)::(elim_quant_in_list sign list);
+                        (elim_quantifications sign trm)::(elim_quant_in_list sign list);
 
 fun dummy true = writeln "True\n" |
     dummy false = writeln "Fals\n";
@@ -458,12 +458,12 @@
       if is_mudef trm 
       then (make_mu_def sign trm)::(transform_definitions sign list)
       else 
-	if is_nudef trm
-	 then (make_nu_def sign trm)::(transform_definitions sign list)
-     	 else 
-	   if is_fundef trm
- 	   then (make_mucke_fun_def sign trm)::(transform_definitions sign list)
-		     else trm::(transform_definitions sign list);
+        if is_nudef trm
+         then (make_nu_def sign trm)::(transform_definitions sign list)
+         else 
+           if is_fundef trm
+           then (make_mucke_fun_def sign trm)::(transform_definitions sign list)
+                     else trm::(transform_definitions sign list);
 
 fun terms_to_decls sign [] = []
  | terms_to_decls sign (trm::list) =
@@ -484,7 +484,7 @@
 
 fun string_of_terms sign terms =
 let fun make_string sign [] = "" |
- 	make_string sign (trm::list) =
+        make_string sign (trm::list) =
            Syntax.string_of_term_global sign trm ^ "\n" ^ make_string sign list
 in
   PrintMode.setmp ["Mucke"] (make_string sign) terms
@@ -531,7 +531,7 @@
     | delete_blanks (":"::xs) = delete_blanks xs
     | delete_blanks (x::xs) = 
         if (is_blank x) then (delete_blanks xs)
-	       	        else x::(delete_blanks xs);
+                        else x::(delete_blanks xs);
   
   fun delete_blanks_string s = implode(delete_blanks (explode s));
 
@@ -580,12 +580,12 @@
   check_complex ((a,[])::r) (t,i) = check_complex r (t,i+1) |
   check_complex ((a,_)::r) (t,i) = check_complex r (true,i+1);
  in
-	if (a=t) then (check_complex l (false,0)) else (tl_contains_complex r t)
+        if (a=t) then (check_complex l (false,0)) else (tl_contains_complex r t)
  end;
  fun check_head_for_case sg (Const(s,_)) st 0 = 
-	if (post_last_dot(s) = (st ^ "_case")) then true else false |
+        if (post_last_dot(s) = (st ^ "_case")) then true else false |
  check_head_for_case sg (a $ (Const(s,_))) st 0 = 
-	if (post_last_dot(s) = (st ^ "_case")) then true else false |
+        if (post_last_dot(s) = (st ^ "_case")) then true else false |
  check_head_for_case _ _ _ 0 = false |
  check_head_for_case sg (a $ b) st n = check_head_for_case sg a st (n-1) |
  check_head_for_case _ _ _ _ = false;
@@ -636,22 +636,22 @@
  heart_of_trm t = t;
  fun replace_termlist_with_args _ _ trm _ [] _ ([],[]) = trm (* should never occur *) |
  replace_termlist_with_args sg _ trm _ [a] _ ([],[]) =
-	if ((heart_of_trm trm)= Const("False",Type("bool",[]))) then trm else 
-	(enrich_case_with_terms sg a trm) |
+        if ((heart_of_trm trm)= Const("False",Type("bool",[]))) then trm else 
+        (enrich_case_with_terms sg a trm) |
  replace_termlist_with_args sg tl trm con [] t (l1,l2) =
-	(replace_termlist_with_constrs sg tl l1 l2 t) | 
+        (replace_termlist_with_constrs sg tl l1 l2 t) | 
  replace_termlist_with_args sg tl trm con (a::r) t (l1,l2) =
  let
   val tty = type_of_term t;
   val con_term = con_term_of con a;
  in
-	(Const("HOL.If",Type("fun",[Type("bool",[]),
+        (Const("HOL.If",Type("fun",[Type("bool",[]),
         Type("fun",[Type("bool",[]),Type("fun",[Type("bool",[]),Type("bool",[])])])])) $
-	(Const("op =",Type("fun",[tty,Type("fun",[tty,Type("bool",[])])])) $
+        (Const("op =",Type("fun",[tty,Type("fun",[tty,Type("bool",[])])])) $
         t $ con_term) $
-	(if ((heart_of_trm trm)= Const("False",Type("bool",[]))) then trm else 
-	(enrich_case_with_terms sg a trm)) $
-	(replace_termlist_with_args sg tl trm con r t (l1,l2)))
+        (if ((heart_of_trm trm)= Const("False",Type("bool",[]))) then trm else 
+        (enrich_case_with_terms sg a trm)) $
+        (replace_termlist_with_args sg tl trm con r t (l1,l2)))
  end;
  val arglist = arglist_of sg tl (snd c);
  val tty = type_of_term t;
@@ -670,7 +670,7 @@
  val ty = type_of_term trm;
  val constr_list = search_in_keylist tl ty;
 in
-	replace_termlist_with_constrs sg tl trm_list constr_list trm
+        replace_termlist_with_constrs sg tl trm_list constr_list trm
 end;  
 
 in
@@ -683,7 +683,7 @@
  (* rc_in_term changes the case statement over term x to a cascading if; the last *)
  (* indicates the distance to the "case"-constant                                 *)
  fun rc_in_term sg tl (a $ b) l x 0 =
-	 (replace_case sg tl a 0) $ (rc_in_termlist sg tl l x) |  
+         (replace_case sg tl a 0) $ (rc_in_termlist sg tl l x) |  
  rc_in_term sg tl  _ l x 0 = rc_in_termlist sg tl l x |
  rc_in_term sg tl (a $ b) l x n = rc_in_term sg tl a (b::l) x (n-1) |
  rc_in_term sg _ trm _ _ n =
@@ -691,11 +691,11 @@
  val (bv,n) = check_case sg tl (a $ b);
 in
  if (bv) then 
-	(let
-	val t = (replace_case sg tl a n) 
-	in 
-	rc_in_term sg tl t [] b n	
-	end)
+        (let
+        val t = (replace_case sg tl a n) 
+        in 
+        rc_in_term sg tl t [] b n       
+        end)
  else ((replace_case sg tl a 0) $ (replace_case sg tl b 0))
 end |
 replace_case sg tl (a $ b) 1 = a $ (replace_case sg tl b 0) |
@@ -786,7 +786,7 @@
 (* fst vt muss frei vorkommen, als ECHTER TeilKonstruktorterm *)
 then (Abs(x,ty,
         abstract_over(Free(fst vt,ty),
-	if_term sg ((Type("bool",[]),[("True",[]),("False",[])])::tl) (fst vt) ty tt)))
+        if_term sg ((Type("bool",[]),[("True",[]),("False",[])])::tl) (fst vt) ty tt)))
 else Abs(x,ty,abstract_over(Free(fst vt,ty),tt))
 end |
 remove_vars sg tl (a $ b) = (remove_vars sg tl a) $ (remove_vars sg tl b) |
@@ -795,7 +795,7 @@
 (* dissolves equalities "=" of boolean terms, where one of them is a complex term *)
 fun remove_equal sg tl (Abs(x,ty,trm)) = Abs(x,ty,remove_equal sg tl trm) |
 remove_equal sg tl (Const("op =",Type("fun",[Type("bool",[]),
-	Type("fun",[Type("bool",[]),Type("bool",[])])])) $ lhs $ rhs) =
+        Type("fun",[Type("bool",[]),Type("bool",[])])])) $ lhs $ rhs) =
 let
 fun complex_trm (Abs(_,_,_)) = true |
 complex_trm (_ $ _) = true |
@@ -811,16 +811,16 @@
 Type("fun",[Type("bool",[]),Type("fun",[Type("bool",[]),Type("bool",[])])])) $
 (Const("op -->",
  Type("fun",[Type("bool",[]),Type("fun",[Type("bool",[]),Type("bool",[])])])) $
-	lhs $ rhs) $
+        lhs $ rhs) $
 (Const("op -->",
  Type("fun",[Type("bool",[]),Type("fun",[Type("bool",[]),Type("bool",[])])])) $
-	rhs $ lhs)
+        rhs $ lhs)
 )
 end)
 else
 (Const("op =",
  Type("fun",[Type("bool",[]),Type("fun",[Type("bool",[]),Type("bool",[])])])) $
-	lhs $ rhs))
+        lhs $ rhs))
 end |
 remove_equal sg tl (a $ b) = (remove_equal sg tl a) $ (remove_equal sg tl b) |
 remove_equal sg tl trm = trm;
@@ -833,7 +833,7 @@
 (* OUTPUT - relevant *)
 (* transforms constructor term to a mucke-suitable output *)
 fun term_OUTPUT sg (Const("Pair",_) $ a $ b) =
-		(term_OUTPUT sg a) ^ "_I_" ^ (term_OUTPUT sg b) |
+                (term_OUTPUT sg a) ^ "_I_" ^ (term_OUTPUT sg b) |
 term_OUTPUT sg (a $ b) = (term_OUTPUT sg a) ^ "_I_" ^ (term_OUTPUT sg b) |
 term_OUTPUT sg (Const(s,t)) = post_last_dot s |
 term_OUTPUT _ _ = 
@@ -846,7 +846,7 @@
 
 in
         if (contains_bound 0 (a $ b)) 
-	then ((rewrite_dt_term sg tl a) $ (rewrite_dt_term sg tl b))
+        then ((rewrite_dt_term sg tl a) $ (rewrite_dt_term sg tl b))
         else
         (
         let
@@ -893,10 +893,10 @@
          (generate_constnames_OUTPUT (c ^ "_I_") l tl) @ (concat_constrs r tl);
  fun concat_types _ [] _ _ = [] |
  concat_types prestring (a::q) [] tl = [prestring ^ a] 
-				       @ (concat_types prestring q [] tl) |
+                                       @ (concat_types prestring q [] tl) |
  concat_types prestring (a::q) r tl = 
-		(generate_constnames_OUTPUT (prestring ^ a ^ "_I_") r tl) 
-		@ (concat_types prestring q r tl);
+                (generate_constnames_OUTPUT (prestring ^ a ^ "_I_") r tl) 
+                @ (concat_types prestring q r tl);
  val g = concat_constrs (search_in_keylist tl a) tl;
 in
  concat_types prestring g r tl
@@ -970,15 +970,15 @@
         then (preprocess_td sg b done)
         else
         (let
-	 fun qtn (Type(a,tl)) = (a,tl) |
-	 qtn _ = error "unexpected type variable in preprocess_td";
-	 val s =  post_last_dot(fst(qtn a));
-	 fun param_types ((Const("Trueprop",_)) $ (_ $ (Var(_,Type(_,t)))))  = t |
-	 param_types _ = error "malformed induct-theorem in preprocess_td";	
-	 val induct_rule = PureThy.get_thm sg (s ^ ".induct");
-	 val pl = param_types (concl_of induct_rule);
+         fun qtn (Type(a,tl)) = (a,tl) |
+         qtn _ = error "unexpected type variable in preprocess_td";
+         val s =  post_last_dot(fst(qtn a));
+         fun param_types ((Const("Trueprop",_)) $ (_ $ (Var(_,Type(_,t)))))  = t |
+         param_types _ = error "malformed induct-theorem in preprocess_td";     
+         val induct_rule = PureThy.get_thm sg (s ^ ".induct");
+         val pl = param_types (concl_of induct_rule);
          val l = split_constrs sg (snd(qtn a)) pl (prems_of induct_rule);
-	 val ntl = new_types l;
+         val ntl = new_types l;
         in 
         ((a,l) :: (preprocess_td sg (ntl @ b) (a :: done)))
         end)
@@ -1015,40 +1015,40 @@
   val sign = Thm.theory_of_cterm csubgoal;
   val subgoal = Thm.term_of csubgoal;
 
-	val Freesubgoal = freeze_thaw subgoal;
+        val Freesubgoal = freeze_thaw subgoal;
 
-	val prems = Logic.strip_imp_prems Freesubgoal; 
-	val concl = Logic.strip_imp_concl Freesubgoal; 
-	
-	val Muckedecls = terms_to_decls sign prems;
-	val decls_str = string_of_terms sign Muckedecls;
-	
-	val type_list = (extract_type_names_prems sign (prems@[concl]));
-	val ctl =  preprocess_td sign type_list [];
-	val b = if (ctl=[]) then true else (check_finity [Type("bool",[])] [] ctl false);
-	val type_str = make_type_decls ctl 
-				((Type("bool",[]),("True",[])::("False",[])::[])::ctl);
-	
-	val mprems = rewrite_dt_terms sign ctl prems;
-	val mconcl = rewrite_dt_terms sign ctl [concl];
+        val prems = Logic.strip_imp_prems Freesubgoal; 
+        val concl = Logic.strip_imp_concl Freesubgoal; 
+        
+        val Muckedecls = terms_to_decls sign prems;
+        val decls_str = string_of_terms sign Muckedecls;
+        
+        val type_list = (extract_type_names_prems sign (prems@[concl]));
+        val ctl =  preprocess_td sign type_list [];
+        val b = if (ctl=[]) then true else (check_finity [Type("bool",[])] [] ctl false);
+        val type_str = make_type_decls ctl 
+                                ((Type("bool",[]),("True",[])::("False",[])::[])::ctl);
+        
+        val mprems = rewrite_dt_terms sign ctl prems;
+        val mconcl = rewrite_dt_terms sign ctl [concl];
 
-	val Muckeprems = transform_terms sign mprems;
+        val Muckeprems = transform_terms sign mprems;
         val prems_str = transform_case(string_of_terms sign Muckeprems);
 
         val Muckeconcl = elim_quant_in_list sign mconcl;
-	val concl_str = transform_case(string_of_terms sign Muckeconcl);
+        val concl_str = transform_case(string_of_terms sign Muckeconcl);
 
-	val MCstr = (
-		"/**** type declarations: ****/\n" ^ type_str ^
-		"/**** declarations: ****/\n" ^ decls_str ^
-		"/**** definitions: ****/\n" ^ prems_str ^ 
-		"/**** formula: ****/\n" ^ concl_str ^";");
-	val result = callmc MCstr;
+        val MCstr = (
+                "/**** type declarations: ****/\n" ^ type_str ^
+                "/**** declarations: ****/\n" ^ decls_str ^
+                "/**** definitions: ****/\n" ^ prems_str ^ 
+                "/**** formula: ****/\n" ^ concl_str ^";");
+        val result = callmc MCstr;
   in
 (if !trace_mc then 
-	(writeln ("\nmodelchecker input:\n" ^ MCstr ^ "\n/**** end ****/"))
+        (writeln ("\nmodelchecker input:\n" ^ MCstr ^ "\n/**** end ****/"))
           else ();
 (case (extract_result concl_str result) of 
-	true  =>  cterm_of sign (Logic.strip_imp_concl subgoal) | 
-	false => (error ("Mucke couldn't solve subgoal: \n" ^result)))) 
+        true  =>  cterm_of sign (Logic.strip_imp_concl subgoal) | 
+        false => (error ("Mucke couldn't solve subgoal: \n" ^result)))) 
   end;
--- a/src/HOL/NSA/NSA.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/NSA/NSA.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
-(*  Title       : NSA.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-
-Converted to Isar and polished by lcp
+(*  Title:      HOL/NSA/NSA.thy
+    Author:     Jacques D. Fleuriot, University of Cambridge
+    Author:     Lawrence C Paulson, University of Cambridge
 *)
 
 header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*}
@@ -2289,10 +2287,10 @@
      "\<forall>n. norm(X n - Y n) < inverse(real(Suc n))
       ==> star_n X - star_n Y \<in> Infinitesimal"
 by (auto intro!: bexI
-	 dest: FreeUltrafilterNat_inverse_real_of_posnat 
-	       FreeUltrafilterNat.Int
-	 intro: order_less_trans FreeUltrafilterNat.subset 
-	 simp add: Infinitesimal_FreeUltrafilterNat_iff star_n_diff 
+         dest: FreeUltrafilterNat_inverse_real_of_posnat 
+               FreeUltrafilterNat.Int
+         intro: order_less_trans FreeUltrafilterNat.subset 
+         simp add: Infinitesimal_FreeUltrafilterNat_iff star_n_diff 
                    star_n_inverse)
 
 end
--- a/src/HOL/NanoJava/Decl.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/NanoJava/Decl.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/NanoJava/Decl.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   2001 Technische Universitaet Muenchen
 *)
@@ -13,27 +12,27 @@
   | Class cname  --{* class type *}
 
 text{* Field declaration *}
-types	fdecl		
-	= "fname \<times> ty"
+types   fdecl           
+        = "fname \<times> ty"
 
-record  methd		
-	= par :: ty 
+record  methd           
+        = par :: ty 
           res :: ty 
           lcl ::"(vname \<times> ty) list"
           bdy :: stmt
 
 text{* Method declaration *}
-types	mdecl
+types   mdecl
         = "mname \<times> methd"
 
-record	"class"
-	= super   :: cname
+record  "class"
+        = super   :: cname
           flds    ::"fdecl list"
           methods ::"mdecl list"
 
 text{* Class declaration *}
-types	cdecl
-	= "cname \<times> class"
+types   cdecl
+        = "cname \<times> class"
 
 types   prog
         = "cdecl list"
@@ -47,12 +46,12 @@
 
 consts
 
-  Prog    :: prog	--{* program as a global value *}
-  Object  :: cname	--{* name of root class *}
+  Prog    :: prog       --{* program as a global value *}
+  Object  :: cname      --{* name of root class *}
 
 
 constdefs
-  "class"	     :: "cname \<rightharpoonup> class"
+ "class"     :: "cname \<rightharpoonup> class"
  "class      \<equiv> map_of Prog"
 
   is_class   :: "cname => bool"
--- a/src/HOL/NanoJava/Example.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/NanoJava/Example.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/NanoJava/Example.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   2001 Technische Universitaet Muenchen
 *)
@@ -30,9 +29,9 @@
 
   public static void main(String[] args) // test x+1=1+x
     {
-	Nat one = new Nat().suc();
-	Nat x   = new Nat().suc().suc().suc().suc();
-	Nat ok = x.suc().eq(x.add(one));
+        Nat one = new Nat().suc();
+        Nat x   = new Nat().suc().suc().suc().suc();
+        Nat ok = x.suc().eq(x.add(one));
         System.out.println(ok != null);
     }
 }
--- a/src/HOL/NanoJava/OpSem.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/NanoJava/OpSem.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/NanoJava/OpSem.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   2001 Technische Universitaet Muenchen
 *)
@@ -56,21 +55,21 @@
 
 
 inductive_cases exec_elim_cases':
-				  "s -Skip            -n\<rightarrow> t"
-				  "s -c1;; c2         -n\<rightarrow> t"
-				  "s -If(e) c1 Else c2-n\<rightarrow> t"
-				  "s -While(x) c      -n\<rightarrow> t"
-				  "s -x:==e           -n\<rightarrow> t"
-				  "s -e1..f:==e2      -n\<rightarrow> t"
+                                  "s -Skip            -n\<rightarrow> t"
+                                  "s -c1;; c2         -n\<rightarrow> t"
+                                  "s -If(e) c1 Else c2-n\<rightarrow> t"
+                                  "s -While(x) c      -n\<rightarrow> t"
+                                  "s -x:==e           -n\<rightarrow> t"
+                                  "s -e1..f:==e2      -n\<rightarrow> t"
 inductive_cases Meth_elim_cases:  "s -Meth Cm         -n\<rightarrow> t"
 inductive_cases Impl_elim_cases:  "s -Impl Cm         -n\<rightarrow> t"
 lemmas exec_elim_cases = exec_elim_cases' Meth_elim_cases Impl_elim_cases
 inductive_cases eval_elim_cases:
-				  "s -new C         \<succ>v-n\<rightarrow> t"
-				  "s -Cast C e      \<succ>v-n\<rightarrow> t"
-				  "s -LAcc x        \<succ>v-n\<rightarrow> t"
-				  "s -e..f          \<succ>v-n\<rightarrow> t"
-				  "s -{C}e1..m(e2)  \<succ>v-n\<rightarrow> t"
+                                  "s -new C         \<succ>v-n\<rightarrow> t"
+                                  "s -Cast C e      \<succ>v-n\<rightarrow> t"
+                                  "s -LAcc x        \<succ>v-n\<rightarrow> t"
+                                  "s -e..f          \<succ>v-n\<rightarrow> t"
+                                  "s -{C}e1..m(e2)  \<succ>v-n\<rightarrow> t"
 
 lemma exec_eval_mono [rule_format]: 
   "(s -c  -n\<rightarrow> t \<longrightarrow> (\<forall>m. n \<le> m \<longrightarrow> s -c  -m\<rightarrow> t)) \<and>
--- a/src/HOL/NanoJava/State.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/NanoJava/State.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/NanoJava/State.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   2001 Technische Universitaet Muenchen
 *)
@@ -20,8 +19,8 @@
   = Null        --{* null reference *}
   | Addr loc    --{* address, i.e. location of object *}
 
-types	fields
-	= "(fname \<rightharpoonup> val)"
+types   fields
+        = "(fname \<rightharpoonup> val)"
 
         obj = "cname \<times> fields"
 
@@ -36,12 +35,12 @@
  "init_vars m == Option.map (\<lambda>T. Null) o m"
   
 text {* private: *}
-types	heap   = "loc   \<rightharpoonup> obj"
-        locals = "vname \<rightharpoonup> val"	
+types   heap   = "loc   \<rightharpoonup> obj"
+        locals = "vname \<rightharpoonup> val"  
 
 text {* private: *}
 record  state
-	= heap   :: heap
+        = heap   :: heap
           locals :: locals
 
 translations
@@ -97,7 +96,7 @@
   upd_obj    :: "loc => fname => val => state => state"
  "upd_obj a f v s \<equiv> let (C,fs) = the (heap s a) in hupd(a\<mapsto>(C,fs(f\<mapsto>v))) s"
 
-  new_Addr	:: "state => val"
+  new_Addr      :: "state => val"
  "new_Addr s == SOME v. (\<exists>a. v = Addr a \<and> (heap s) a = None) | v = Null"
 
 
--- a/src/HOL/Nominal/Examples/Class.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/Examples/Class.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -8497,63 +8497,63 @@
     { assume asm: "N\<noteq>Ax y b"
       have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} = 
         (Cut <a>.NotR (u).(M{y:=<c>.P}) a (v).NotL <b>.(N{y:=<c>.P}) v)" using prems
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>l (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using prems
-	by (auto intro: l_redu.intros simp add: subst_fresh)
+        by (auto intro: l_redu.intros simp add: subst_fresh)
       also have "\<dots> = (Cut <b>.N (u).M){y:=<c>.P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally have ?thesis by auto
     }
     moreover
     { assume asm: "N=Ax y b"
       have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} = 
         (Cut <a>.NotR (u).(M{y:=<c>.P}) a (v).NotL <b>.(N{y:=<c>.P}) v)" using prems
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <b>.(Cut <c>.P (y).Ax y b) (u).(M{y:=<c>.P}))" using prems
-	by simp
+        by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(P[c\<turnstile>c>b]) (u).(M{y:=<c>.P}))" 
       proof (cases "fic P c")
-	case True 
-	assume "fic P c"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutL_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxR_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fic P c"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutL_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxR_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fic P c" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_left)
-	  apply(simp)
-	  apply(simp add: subst_with_ax2)
-	  done
+        case False 
+        assume "\<not>fic P c" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_left)
+          apply(simp)
+          apply(simp add: subst_with_ax2)
+          done
       qed
       also have "\<dots> = (Cut <b>.N (u).M){y:=<c>.P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule crename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule crename_swap)
+        apply(simp)
+        done
       finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <b>.N (u).M){y:=<c>.P}" 
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -8564,72 +8564,72 @@
     { assume asm: "M1\<noteq>Ax y a1"
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} = 
         Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL1 (u).(N{y:=<c>.P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a1>.M1 (u).N){y:=<c>.P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "M1=Ax y a1"
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} = 
         Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL1 (u).(N{y:=<c>.P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a1>.(Cut <c>.P (y). Ax y a1) (u).(N{y:=<c>.P})" 
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.P[c\<turnstile>c>a1] (u).(N{y:=<c>.P})"
       proof (cases "fic P c")
-	case True 
-	assume "fic P c"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutL_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxR_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fic P c"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutL_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxR_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fic P c" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_left)
-	  apply(simp)
-	  apply(simp add: subst_with_ax2)
-	  done
+        case False 
+        assume "\<not>fic P c" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_left)
+          apply(simp)
+          apply(simp add: subst_with_ax2)
+          done
       qed
       also have "\<dots> = (Cut <a1>.M1 (u).N){y:=<c>.P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule crename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule crename_swap)
+        apply(simp)
+        done
       finally 
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -8640,72 +8640,72 @@
     { assume asm: "M2\<noteq>Ax y a2"
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} = 
         Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL2 (u).(N{y:=<c>.P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a2>.M2 (u).N){y:=<c>.P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "M2=Ax y a2"
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} = 
         Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL2 (u).(N{y:=<c>.P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a2>.(Cut <c>.P (y). Ax y a2) (u).(N{y:=<c>.P})" 
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.P[c\<turnstile>c>a2] (u).(N{y:=<c>.P})"
       proof (cases "fic P c")
-	case True 
-	assume "fic P c"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutL_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxR_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fic P c"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutL_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxR_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fic P c" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_left)
-	  apply(simp)
-	  apply(simp add: subst_with_ax2)
-	  done
+        case False 
+        assume "\<not>fic P c" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_left)
+          apply(simp)
+          apply(simp add: subst_with_ax2)
+          done
       qed
       also have "\<dots> = (Cut <a2>.M2 (u).N){y:=<c>.P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule crename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule crename_swap)
+        apply(simp)
+        done
       finally 
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -8716,72 +8716,72 @@
     { assume asm: "M\<noteq>Ax y a"
       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
         Cut <b>.OrR1 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a>.M (x1).N1){y:=<c>.P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "M=Ax y a"
       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
         Cut <b>.OrR1 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a>.(Cut <c>.P (y). Ax y a) (x1).(N1{y:=<c>.P})" 
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.P[c\<turnstile>c>a] (x1).(N1{y:=<c>.P})"
       proof (cases "fic P c")
-	case True 
-	assume "fic P c"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutL_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxR_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fic P c"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutL_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxR_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fic P c" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_left)
-	  apply(simp)
-	  apply(simp add: subst_with_ax2)
-	  done
+        case False 
+        assume "\<not>fic P c" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_left)
+          apply(simp)
+          apply(simp add: subst_with_ax2)
+          done
       qed
       also have "\<dots> = (Cut <a>.M (x1).N1){y:=<c>.P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule crename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule crename_swap)
+        apply(simp)
+        done
       finally 
       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -8792,72 +8792,72 @@
     { assume asm: "M\<noteq>Ax y a"
       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
         Cut <b>.OrR2 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a>.M (x2).N2){y:=<c>.P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "M=Ax y a"
       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
         Cut <b>.OrR2 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a>.(Cut <c>.P (y). Ax y a) (x2).(N2{y:=<c>.P})" 
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.P[c\<turnstile>c>a] (x2).(N2{y:=<c>.P})"
       proof (cases "fic P c")
-	case True 
-	assume "fic P c"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutL_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxR_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fic P c"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutL_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxR_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fic P c" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_left)
-	  apply(simp)
-	  apply(simp add: subst_with_ax2)
-	  done
+        case False 
+        assume "\<not>fic P c" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_left)
+          apply(simp)
+          apply(simp add: subst_with_ax2)
+          done
       qed
       also have "\<dots> = (Cut <a>.M (x2).N2){y:=<c>.P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule crename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule crename_swap)
+        apply(simp)
+        done
       finally 
       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -8868,80 +8868,80 @@
     { assume asm: "N\<noteq>Ax y d"
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} = 
         Cut <b>.ImpR (x).<a>.(M{y:=<c>.P}) b (z).ImpL <d>.(N{y:=<c>.P}) (u).(Q{y:=<c>.P}) z" 
-	using prems by (simp add: fresh_prod abs_fresh fresh_atm)
+        using prems by (simp add: fresh_prod abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P})  (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a>.(Cut <d>.N  (x).M) (u).Q){y:=<c>.P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^isub>a* 
                      (Cut <a>.(Cut <d>.N  (x).M) (u).Q){y:=<c>.P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "N=Ax y d"
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} = 
         Cut <b>.ImpR (x).<a>.(M{y:=<c>.P}) b (z).ImpL <d>.(N{y:=<c>.P}) (u).(Q{y:=<c>.P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P})  (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a>.(Cut <d>.(Cut <c>.P (y).Ax y d)  (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(P[c\<turnstile>c>d]) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
       proof (cases "fic P c")
-	case True 
-	assume "fic P c"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutL_intro)
-	  apply(rule a_Cut_l)
-	  apply(simp add: subst_fresh abs_fresh)
-	  apply(simp add: abs_fresh fresh_atm)
-	  apply(rule al_redu)
-	  apply(rule better_LAxR_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fic P c"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutL_intro)
+          apply(rule a_Cut_l)
+          apply(simp add: subst_fresh abs_fresh)
+          apply(simp add: abs_fresh fresh_atm)
+          apply(rule al_redu)
+          apply(rule better_LAxR_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fic P c" 
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_left)
-	  apply(simp)
-	  apply(simp add: subst_with_ax2)
-	  done
+        case False 
+        assume "\<not>fic P c" 
+        then show ?thesis using prems
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_CutL)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_left)
+          apply(simp)
+          apply(simp add: subst_with_ax2)
+          done
       qed
       also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha)
-	apply(rule sym)
-	apply(rule crename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha)
+        apply(rule sym)
+        apply(rule crename_swap)
+        apply(simp)
+        done
       finally 
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^isub>a* 
-	       (Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}"
-	by simp
+               (Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}"
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -9007,63 +9007,63 @@
     { assume asm: "M\<noteq>Ax u c"
       have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} = 
         (Cut <a>.NotR (u).(M{c:=(y).P}) a (v).NotL <b>.(N{c:=(y).P}) v)" using prems
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>l (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using prems
-	by (auto intro: l_redu.intros simp add: subst_fresh)
+        by (auto intro: l_redu.intros simp add: subst_fresh)
       also have "\<dots> = (Cut <b>.N (u).M){c:=(y).P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally have ?thesis by auto
     }
     moreover
     { assume asm: "M=Ax u c"
       have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} = 
         (Cut <a>.NotR (u).(M{c:=(y).P}) a (v).NotL <b>.(N{c:=(y).P}) v)" using prems
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <b>.(N{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P))" using prems
-	by simp
+        by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{c:=(y).P})  (u).(P[y\<turnstile>n>u]))" 
       proof (cases "fin P y")
-	case True 
-	assume "fin P y"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutR_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxL_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fin P y"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutR_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxL_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fin P y" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_right)
-	  apply(simp)
-	  apply(simp add: subst_with_ax1)
-	  done
+        case False 
+        assume "\<not>fin P y" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_right)
+          apply(simp)
+          apply(simp add: subst_with_ax1)
+          done
       qed
       also have "\<dots> = (Cut <b>.N (u).M){c:=(y).P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule nrename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule nrename_swap)
+        apply(simp)
+        done
       finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <b>.N (u).M){c:=(y).P}" 
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -9074,72 +9074,72 @@
     { assume asm: "N\<noteq>Ax u c"
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} = 
         Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a1>.M1 (u).N){c:=(y).P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "N=Ax u c"
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} = 
         Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a1>.(M1{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P)" 
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
       proof (cases "fin P y")
-	case True 
-	assume "fin P y"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutR_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxL_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fin P y"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutR_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxL_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fin P y" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_right)
-	  apply(simp)
-	  apply(simp add: subst_with_ax1)
-	  done
+        case False 
+        assume "\<not>fin P y" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_right)
+          apply(simp)
+          apply(simp add: subst_with_ax1)
+          done
       qed
       also have "\<dots> = (Cut <a1>.M1 (u).N){c:=(y).P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule nrename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule nrename_swap)
+        apply(simp)
+        done
       finally 
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -9150,72 +9150,72 @@
     { assume asm: "N\<noteq>Ax u c"
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} = 
         Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a2>.M2 (u).N){c:=(y).P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "N=Ax u c"
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} = 
         Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a2>.(M2{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P)" 
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
       proof (cases "fin P y")
-	case True 
-	assume "fin P y"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutR_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxL_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fin P y"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutR_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxL_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fin P y" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_right)
-	  apply(simp)
-	  apply(simp add: subst_with_ax1)
-	  done
+        case False 
+        assume "\<not>fin P y" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_right)
+          apply(simp)
+          apply(simp add: subst_with_ax1)
+          done
       qed
       also have "\<dots> = (Cut <a2>.M2 (u).N){c:=(y).P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule nrename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule nrename_swap)
+        apply(simp)
+        done
       finally 
       have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -9226,72 +9226,72 @@
     { assume asm: "N1\<noteq>Ax x1 c"
       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
         Cut <b>.OrR1 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a>.M (x1).N1){c:=(y).P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "N1=Ax x1 c"
       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
         Cut <b>.OrR1 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a>.(M{c:=(y).P}) (x1).(Cut <c>.(Ax x1 c) (y).P)" 
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P})   (x1).(P[y\<turnstile>n>x1])"
       proof (cases "fin P y")
-	case True 
-	assume "fin P y"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutR_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxL_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fin P y"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutR_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxL_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fin P y" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_right)
-	  apply(simp)
-	  apply(simp add: subst_with_ax1)
-	  done
+        case False 
+        assume "\<not>fin P y" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_right)
+          apply(simp)
+          apply(simp add: subst_with_ax1)
+          done
       qed
       also have "\<dots> = (Cut <a>.M (x1).N1){c:=(y).P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule nrename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule nrename_swap)
+        apply(simp)
+        done
       finally 
       have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -9302,72 +9302,72 @@
     { assume asm: "N2\<noteq>Ax x2 c"
       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
         Cut <b>.OrR2 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a>.M (x2).N2){c:=(y).P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "N2=Ax x2 c"
       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
         Cut <b>.OrR2 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a>.(M{c:=(y).P}) (x2).(Cut <c>.(Ax x2 c) (y).P)" 
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(P[y\<turnstile>n>x2])"
       proof (cases "fin P y")
-	case True 
-	assume "fin P y"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_starI)
-	  apply(rule better_CutR_intro)
-	  apply(rule al_redu)
-	  apply(rule better_LAxL_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fin P y"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_starI)
+          apply(rule better_CutR_intro)
+          apply(rule al_redu)
+          apply(rule better_LAxL_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fin P y" 
-	then show ?thesis
-	  apply -
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_right)
-	  apply(simp)
-	  apply(simp add: subst_with_ax1)
-	  done
+        case False 
+        assume "\<not>fin P y" 
+        then show ?thesis
+          apply -
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_right)
+          apply(simp)
+          apply(simp add: subst_with_ax1)
+          done
       qed
       also have "\<dots> = (Cut <a>.M (x2).N2){c:=(y).P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(rule sym)
-	apply(rule nrename_swap)
-	apply(simp)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(rule sym)
+        apply(rule nrename_swap)
+        apply(simp)
+        done
       finally 
       have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
-	by simp
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -9378,213 +9378,213 @@
     { assume asm: "M\<noteq>Ax x c \<and> Q\<noteq>Ax u c"
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
         Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
-	using prems by (simp add: fresh_prod abs_fresh fresh_atm)
+        using prems by (simp add: fresh_prod abs_fresh fresh_atm)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = (Cut <a>.(Cut <d>.N  (x).M) (u).Q){c:=(y).P}" using prems 
-	by (simp add: subst_fresh abs_fresh fresh_atm)
+        by (simp add: subst_fresh abs_fresh fresh_atm)
       finally 
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
                      (Cut <a>.(Cut <d>.N  (x).M) (u).Q){c:=(y).P}"
-	by simp
+        by simp
     } 
     moreover
     { assume asm: "M=Ax x c \<and> Q\<noteq>Ax u c"
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
         Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(Cut <c>.Ax x c (y).P)) (u).(Q{c:=(y).P})"
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(P[y\<turnstile>n>x])) (u).(Q{c:=(y).P})"
       proof (cases "fin P y")
-	case True 
-	assume "fin P y"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule al_redu)
-	  apply(rule better_LAxL_intro)
-	  apply(simp)
-	  apply(simp)
-	  done
+        case True 
+        assume "fin P y"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule al_redu)
+          apply(rule better_LAxL_intro)
+          apply(simp)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fin P y" 
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_right)
-	  apply(simp)
-	  apply(simp add: subst_with_ax1)
-	  done
+        case False 
+        assume "\<not>fin P y" 
+        then show ?thesis using prems
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_right)
+          apply(simp)
+          apply(simp add: subst_with_ax1)
+          done
       qed
       also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha)
-	apply(simp add: nrename_swap)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha)
+        apply(simp add: nrename_swap)
+        done
       finally 
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
-	       (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
-	by simp
+               (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
+        by simp
     }
      moreover
     { assume asm: "M\<noteq>Ax x c \<and> Q=Ax u c"
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
         Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Cut <c>.Ax u c (y).P)"
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(P[y\<turnstile>n>u])"
       proof (cases "fin P y")
-	case True 
-	assume "fin P y"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_star_CutR)
-	  apply(rule a_starI)
-	  apply(rule al_redu)
-	  apply(rule better_LAxL_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fin P y"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_star_CutR)
+          apply(rule a_starI)
+          apply(rule al_redu)
+          apply(rule better_LAxL_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fin P y" 
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_right)
-	  apply(simp)
-	  apply(simp add: subst_with_ax1)
-	  done
+        case False 
+        assume "\<not>fin P y" 
+        then show ?thesis using prems
+          apply -
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_right)
+          apply(simp)
+          apply(simp add: subst_with_ax1)
+          done
       qed
       also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(simp add: alpha fresh_atm)
-	apply(simp add: nrename_swap)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(simp add: alpha fresh_atm)
+        apply(simp add: nrename_swap)
+        done
       finally 
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
-	       (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
-	by simp
+               (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
+        by simp
     }
      moreover
     { assume asm: "M=Ax x c \<and> Q=Ax u c"
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
         Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
-	using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
+        using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
-	using prems
-	apply -
-	apply(rule a_starI)
-	apply(rule al_redu)
-	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
-	done
+        using prems
+        apply -
+        apply(rule a_starI)
+        apply(rule al_redu)
+        apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
+        done
       also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(Cut <c>.Ax x c (y).P)) (u).(Cut <c>.Ax u c (y).P)"
-	using prems by simp
+        using prems by simp
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(Cut <c>.Ax x c (y).P)) (u).(P[y\<turnstile>n>u])"
       proof (cases "fin P y")
-	case True 
-	assume "fin P y"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_star_CutR)
-	  apply(rule a_starI)
-	  apply(rule al_redu)
-	  apply(rule better_LAxL_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fin P y"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_star_CutR)
+          apply(rule a_starI)
+          apply(rule al_redu)
+          apply(rule better_LAxL_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fin P y" 
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_right)
-	  apply(simp)
-	  apply(simp add: subst_with_ax1)
-	  done
+        case False 
+        assume "\<not>fin P y" 
+        then show ?thesis using prems
+          apply -
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_right)
+          apply(simp)
+          apply(simp add: subst_with_ax1)
+          done
       qed
       also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(P[y\<turnstile>n>x])) (u).(P[y\<turnstile>n>u])"
       proof (cases "fin P y")
-	case True 
-	assume "fin P y"
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_CutR)
-	  apply(rule a_starI)
-	  apply(rule al_redu)
-	  apply(rule better_LAxL_intro)
-	  apply(simp)
-	  done
+        case True 
+        assume "fin P y"
+        then show ?thesis using prems
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_CutR)
+          apply(rule a_starI)
+          apply(rule al_redu)
+          apply(rule better_LAxL_intro)
+          apply(simp)
+          done
       next
-	case False 
-	assume "\<not>fin P y" 
-	then show ?thesis using prems
-	  apply -
-	  apply(rule a_star_CutL)
-	  apply(rule a_star_CutR)
-	  apply(rule a_star_trans)
-	  apply(rule a_starI)
-	  apply(rule ac_redu)
-	  apply(rule better_right)
-	  apply(simp)
-	  apply(simp add: subst_with_ax1)
-	  done
+        case False 
+        assume "\<not>fin P y" 
+        then show ?thesis using prems
+          apply -
+          apply(rule a_star_CutL)
+          apply(rule a_star_CutR)
+          apply(rule a_star_trans)
+          apply(rule a_starI)
+          apply(rule ac_redu)
+          apply(rule better_right)
+          apply(simp)
+          apply(simp add: subst_with_ax1)
+          done
       qed
       also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using prems
-	apply -
-	apply(auto simp add: subst_fresh abs_fresh)
-	apply(simp add: trm.inject)
-	apply(rule conjI)
-	apply(simp add: alpha fresh_atm trm.inject)
-	apply(simp add: nrename_swap)
-	apply(simp add: alpha fresh_atm trm.inject)
-	apply(simp add: nrename_swap)
-	done
+        apply -
+        apply(auto simp add: subst_fresh abs_fresh)
+        apply(simp add: trm.inject)
+        apply(rule conjI)
+        apply(simp add: alpha fresh_atm trm.inject)
+        apply(simp add: nrename_swap)
+        apply(simp add: alpha fresh_atm trm.inject)
+        apply(simp add: nrename_swap)
+        done
       finally 
       have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
-	       (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
-	by simp
+               (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
+        by simp
     }
     ultimately show ?thesis by blast
   qed
@@ -10057,12 +10057,12 @@
     proof (induct b rule: SNa.induct)
       case (SNaI y)
       show ?case
-	apply (rule hyp)
-	apply (erule SNa')
-	apply (erule SNa')
-	apply (rule SNa.SNaI)
-	apply (erule SNaI)+
-	done
+        apply (rule hyp)
+        apply (erule SNa')
+        apply (erule SNa')
+        apply (rule SNa.SNaI)
+        apply (erule SNaI)+
+        done
     qed
   qed
   from b_SNa show ?thesis by (rule r)
@@ -13484,7 +13484,7 @@
      moreover
     { assume "<a>:M \<in> NOTRIGHT (NOT B) (\<parallel>(B)\<parallel>)"
       then obtain x' M' where eq: "M = NotR (x').M' a" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
-	using NOTRIGHT_elim by blast
+        using NOTRIGHT_elim by blast
       then have "SNa M'" using ih2 by blast
       then have "SNa M" using eq by (simp add: NotR_in_SNa)
     }
@@ -13506,7 +13506,7 @@
      moreover
     { assume "(x):M \<in> NOTLEFT (NOT B) (\<parallel><B>\<parallel>)"
       then obtain a' M' where eq: "M = NotL <a'>.M' x" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)"
-	using NOTLEFT_elim by blast
+        using NOTLEFT_elim by blast
       then have "SNa M'" using ih1 by blast
       then have "SNa M" using eq by (simp add: NotL_in_SNa)
     }
@@ -13535,7 +13535,7 @@
     { assume "<a>:M \<in> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)"
       then obtain a' M' b' N' where eq: "M = AndR <a'>.M' <b'>.N' a" 
                                 and "<a'>:M' \<in> (\<parallel><A>\<parallel>)" and "<b'>:N' \<in> (\<parallel><B>\<parallel>)"
-	by (erule_tac ANDRIGHT_elim, blast)
+        by (erule_tac ANDRIGHT_elim, blast)
       then have "SNa M'" and "SNa N'" using ih1 ih3 by blast+
       then have "SNa M" using eq by (simp add: AndR_in_SNa)
     }
@@ -13558,14 +13558,14 @@
      moreover
     { assume "(x):M \<in> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>)"
       then obtain x' M' where eq: "M = AndL1 (x').M' x" and "(x'):M' \<in> (\<parallel>(A)\<parallel>)"
-	using ANDLEFT1_elim by blast
+        using ANDLEFT1_elim by blast
       then have "SNa M'" using ih2 by blast
       then have "SNa M" using eq by (simp add: AndL1_in_SNa)
     }
     moreover
     { assume "(x):M \<in> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)"
       then obtain x' M' where eq: "M = AndL2 (x').M' x" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
-	using ANDLEFT2_elim by blast
+        using ANDLEFT2_elim by blast
       then have "SNa M'" using ih4 by blast
       then have "SNa M" using eq by (simp add: AndL2_in_SNa)
     }
@@ -13594,14 +13594,14 @@
     { assume "<a>:M \<in> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>)"
       then obtain a' M' where eq: "M = OrR1 <a'>.M' a" 
                                 and "<a'>:M' \<in> (\<parallel><A>\<parallel>)" 
-	by (erule_tac ORRIGHT1_elim, blast)
+        by (erule_tac ORRIGHT1_elim, blast)
       then have "SNa M'" using ih1 by blast
       then have "SNa M" using eq by (simp add: OrR1_in_SNa)
     }
      moreover
     { assume "<a>:M \<in> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)"
       then obtain a' M' where eq: "M = OrR2 <a'>.M' a" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)" 
-	using ORRIGHT2_elim by blast
+        using ORRIGHT2_elim by blast
       then have "SNa M'" using ih3 by blast
       then have "SNa M" using eq by (simp add: OrR2_in_SNa)
     }
@@ -13625,7 +13625,7 @@
     { assume "(x):M \<in> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)"
       then obtain x' M' y' N' where eq: "M = OrL (x').M' (y').N' x" 
                                 and "(x'):M' \<in> (\<parallel>(A)\<parallel>)" and  "(y'):N' \<in> (\<parallel>(B)\<parallel>)"
-	by (erule_tac ORLEFT_elim, blast)
+        by (erule_tac ORLEFT_elim, blast)
       then have "SNa M'" and "SNa N'" using ih2 ih4 by blast+
       then have "SNa M" using eq by (simp add: OrL_in_SNa)
     }
@@ -13654,7 +13654,7 @@
     { assume "<a>:M \<in> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)"
       then obtain x' a' M' where eq: "M = ImpR (x').<a'>.M' a" 
                            and imp: "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(M'{a':=(z).P}) \<in> \<parallel>(A)\<parallel>"    
-	by (erule_tac IMPRIGHT_elim, blast)
+        by (erule_tac IMPRIGHT_elim, blast)
       obtain z::"name" where fs: "z\<sharp>x'" by (rule_tac exists_fresh, rule fin_supp, blast)
       have "(z):Ax z a'\<in> \<parallel>(B)\<parallel>" by (simp add: Ax_in_CANDs)
       with imp fs have "(x'):(M'{a':=(z).Ax z a'}) \<in> \<parallel>(A)\<parallel>" by (simp add: fresh_prod fresh_atm)
@@ -13685,7 +13685,7 @@
     { assume "(x):M \<in> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)"
       then obtain a' M' y' N' where eq: "M = ImpL <a'>.M' (y').N' x" 
                                 and "<a'>:M' \<in> (\<parallel><A>\<parallel>)" and  "(y'):N' \<in> (\<parallel>(B)\<parallel>)"
-	by (erule_tac IMPLEFT_elim, blast)
+        by (erule_tac IMPLEFT_elim, blast)
       then have "SNa M'" and "SNa N'" using ih1 ih4 by blast+
       then have "SNa M" using eq by (simp add: ImpL_in_SNa)
     }
@@ -13708,7 +13708,7 @@
   apply(auto)
   apply(drule ax_do_not_a_star_reduce)
   apply(auto)
-  done	
+  done  
 
 lemma BINDING_preserved:
   shows "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
@@ -13803,26 +13803,26 @@
       then obtain x' a' N' where eq: "M = ImpR (x').<a'>.N' a" and fic: "fic (ImpR (x').<a'>.N' a) a"
                            and imp1: "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(N'{a':=(z).P}) \<in> \<parallel>(A)\<parallel>" 
                            and imp2: "\<forall>c Q. a'\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><A>\<parallel> \<longrightarrow> <a'>:(N'{x':=<c>.Q}) \<in> \<parallel><B>\<parallel>"
-	using IMPRIGHT_elim by blast
+        using IMPRIGHT_elim by blast
       from eq asm obtain N'' where eq': "M' = ImpR (x').<a'>.N'' a" and red: "N' \<longrightarrow>\<^isub>a* N''" 
-	using a_star_redu_ImpR_elim by (blast)
+        using a_star_redu_ImpR_elim by (blast)
       from imp1 have "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(N''{a':=(z).P}) \<in> \<parallel>(A)\<parallel>" using red ih2
-	apply(auto)
-	apply(drule_tac x="z" in spec)
-	apply(drule_tac x="P" in spec)
-	apply(simp)
-	apply(drule_tac a_star_subst2)
-	apply(blast)
-	done
+        apply(auto)
+        apply(drule_tac x="z" in spec)
+        apply(drule_tac x="P" in spec)
+        apply(simp)
+        apply(drule_tac a_star_subst2)
+        apply(blast)
+        done
       moreover
       from imp2 have "\<forall>c Q. a'\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><A>\<parallel> \<longrightarrow> <a'>:(N''{x':=<c>.Q}) \<in> \<parallel><B>\<parallel>" using red ih3
-	apply(auto)
-	apply(drule_tac x="c" in spec)
-	apply(drule_tac x="Q" in spec)
-	apply(simp)
-	apply(drule_tac a_star_subst1)
-	apply(blast)
-	done
+        apply(auto)
+        apply(drule_tac x="c" in spec)
+        apply(drule_tac x="Q" in spec)
+        apply(simp)
+        apply(drule_tac a_star_subst1)
+        apply(blast)
+        done
       moreover
       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       ultimately have "<a>:M' \<in> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)" using eq' by auto
@@ -13851,10 +13851,10 @@
       then obtain a' T' y' N' where eq: "M = ImpL <a'>.T' (y').N' x" 
                              and fin: "fin (ImpL <a'>.T' (y').N' x) x"
                              and imp1: "<a'>:T' \<in> \<parallel><A>\<parallel>" and imp2: "(y'):N' \<in> \<parallel>(B)\<parallel>"
-	by (erule_tac IMPLEFT_elim, blast)
+        by (erule_tac IMPLEFT_elim, blast)
       from eq asm obtain T'' N'' where eq': "M' = ImpL <a'>.T'' (y').N'' x" 
                                  and red1: "T' \<longrightarrow>\<^isub>a* T''"  and red2: "N' \<longrightarrow>\<^isub>a* N''"
-	using a_star_redu_ImpL_elim by blast
+        using a_star_redu_ImpL_elim by blast
       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       moreover
       from imp1 red1 have "<a'>:T'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
@@ -13892,10 +13892,10 @@
       then obtain a' T' b' N' where eq: "M = AndR <a'>.T' <b'>.N' a" 
                               and fic: "fic (AndR <a'>.T' <b'>.N' a) a"
                            and imp1: "<a'>:T' \<in> \<parallel><A>\<parallel>" and imp2: "<b'>:N' \<in> \<parallel><B>\<parallel>"
-	using ANDRIGHT_elim by blast
+        using ANDRIGHT_elim by blast
       from eq asm obtain T'' N'' where eq': "M' = AndR <a'>.T'' <b'>.N'' a" 
                           and red1: "T' \<longrightarrow>\<^isub>a* T''" and red2: "N' \<longrightarrow>\<^isub>a* N''" 
-	using a_star_redu_AndR_elim by blast
+        using a_star_redu_AndR_elim by blast
       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       moreover
       from imp1 red1 have "<a'>:T'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
@@ -13926,9 +13926,9 @@
     { assume "(x):M \<in> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>)"
       then obtain y' N' where eq: "M = AndL1 (y').N' x" 
                              and fin: "fin (AndL1 (y').N' x) x" and imp: "(y'):N' \<in> \<parallel>(A)\<parallel>"
-	by (erule_tac ANDLEFT1_elim, blast)
+        by (erule_tac ANDLEFT1_elim, blast)
       from eq asm obtain N'' where eq': "M' = AndL1 (y').N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
-	using a_star_redu_AndL1_elim by blast
+        using a_star_redu_AndL1_elim by blast
       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       moreover
       from imp red1 have "(y'):N'' \<in> \<parallel>(A)\<parallel>" using ih2 by simp
@@ -13938,9 +13938,9 @@
     { assume "(x):M \<in> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)"
       then obtain y' N' where eq: "M = AndL2 (y').N' x" 
                              and fin: "fin (AndL2 (y').N' x) x" and imp: "(y'):N' \<in> \<parallel>(B)\<parallel>"
-	by (erule_tac ANDLEFT2_elim, blast)
+        by (erule_tac ANDLEFT2_elim, blast)
       from eq asm obtain N'' where eq': "M' = AndL2 (y').N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
-	using a_star_redu_AndL2_elim by blast
+        using a_star_redu_AndL2_elim by blast
       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       moreover
       from imp red1 have "(y'):N'' \<in> \<parallel>(B)\<parallel>" using ih4 by simp
@@ -13975,9 +13975,9 @@
     { assume "<a>:M \<in> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>)"
       then obtain a' N' where eq: "M = OrR1 <a'>.N' a" 
                               and fic: "fic (OrR1 <a'>.N' a) a" and imp1: "<a'>:N' \<in> \<parallel><A>\<parallel>"
-	using ORRIGHT1_elim by blast
+        using ORRIGHT1_elim by blast
       from eq asm obtain N'' where eq': "M' = OrR1 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^isub>a* N''" 
-	using a_star_redu_OrR1_elim by blast
+        using a_star_redu_OrR1_elim by blast
       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       moreover
       from imp1 red1 have "<a'>:N'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
@@ -13987,9 +13987,9 @@
     { assume "<a>:M \<in> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)"
       then obtain a' N' where eq: "M = OrR2 <a'>.N' a" 
                               and fic: "fic (OrR2 <a'>.N' a) a" and imp1: "<a'>:N' \<in> \<parallel><B>\<parallel>"
-	using ORRIGHT2_elim by blast
+        using ORRIGHT2_elim by blast
       from eq asm obtain N'' where eq': "M' = OrR2 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^isub>a* N''" 
-	using a_star_redu_OrR2_elim by blast
+        using a_star_redu_OrR2_elim by blast
       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       moreover
       from imp1 red1 have "<a'>:N'' \<in> \<parallel><B>\<parallel>" using ih3 by simp
@@ -14019,10 +14019,10 @@
       then obtain y' T' z' N' where eq: "M = OrL (y').T' (z').N' x" 
                              and fin: "fin (OrL (y').T' (z').N' x) x" 
                              and imp1: "(y'):T' \<in> \<parallel>(A)\<parallel>" and imp2: "(z'):N' \<in> \<parallel>(B)\<parallel>"
-	by (erule_tac ORLEFT_elim, blast)
+        by (erule_tac ORLEFT_elim, blast)
       from eq asm obtain T'' N'' where eq': "M' = OrL (y').T'' (z').N'' x" 
                 and red1: "T' \<longrightarrow>\<^isub>a* T''" and red2: "N' \<longrightarrow>\<^isub>a* N''"
-	using a_star_redu_OrL_elim by blast
+        using a_star_redu_OrL_elim by blast
       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       moreover
       from imp1 red1 have "(y'):T'' \<in> \<parallel>(A)\<parallel>" using ih2 by simp
@@ -14057,9 +14057,9 @@
     { assume "<a>:M \<in> NOTRIGHT (NOT A) (\<parallel>(A)\<parallel>)"
       then obtain y' N' where eq: "M = NotR (y').N' a" 
                               and fic: "fic (NotR (y').N' a) a" and imp: "(y'):N' \<in> \<parallel>(A)\<parallel>"
-	using NOTRIGHT_elim by blast
+        using NOTRIGHT_elim by blast
       from eq asm obtain N'' where eq': "M' = NotR (y').N'' a" and red: "N' \<longrightarrow>\<^isub>a* N''" 
-	using a_star_redu_NotR_elim by blast
+        using a_star_redu_NotR_elim by blast
       from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
       moreover
       from imp red have "(y'):N'' \<in> \<parallel>(A)\<parallel>" using ih2 by simp
@@ -14088,9 +14088,9 @@
     { assume "(x):M \<in> NOTLEFT (NOT A) (\<parallel><A>\<parallel>)"
       then obtain a' N' where eq: "M = NotL <a'>.N' x" 
                              and fin: "fin (NotL <a'>.N' x) x" and imp: "<a'>:N' \<in> \<parallel><A>\<parallel>"
-	by (erule_tac NOTLEFT_elim, blast)
+        by (erule_tac NOTLEFT_elim, blast)
       from eq asm obtain N'' where eq': "M' = NotL <a'>.N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
-	using a_star_redu_NotL_elim by blast
+        using a_star_redu_NotL_elim by blast
       from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
       moreover
       from imp red1 have "<a'>:N'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
--- a/src/HOL/Nominal/Examples/Crary.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/Examples/Crary.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -800,7 +800,7 @@
       then have "\<not> y\<sharp>\<Gamma>" by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_atm fresh_prod)
       then have hf:"x\<sharp> Var y" using fs by (auto simp add: fresh_atm)
       then have "((x,s)#\<theta>)<Var y> = \<theta><Var y>" "((x,t)#\<theta>')<Var y> = \<theta>'<Var y>" 
-	using fresh_psubst_simp by blast+ 
+        using fresh_psubst_simp by blast+ 
       moreover have  "\<Gamma>' \<turnstile> \<theta><Var y> is \<theta>'<Var y> : U" using h1 hl by auto
       ultimately have "\<Gamma>' \<turnstile> ((x,s)#\<theta>)<Var y> is ((x,t)#\<theta>')<Var y> : U" by auto
     }
--- a/src/HOL/Nominal/Examples/Fsub.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/Examples/Fsub.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -350,7 +350,7 @@
     proof (induct \<Gamma>')
       case (Cons Y \<Gamma>')
       thus "\<not> (\<exists>T.(TVarB X' T)\<in>set(Y#\<Gamma>'))"
-	by (simp add:  fresh_ty_dom_cons 
+        by (simp add:  fresh_ty_dom_cons 
                        fresh_fin_union[OF pt_tyvrs_inst  at_tyvrs_inst fs_tyvrs_inst]  
                        finite_vrs finite_doms, 
             auto simp add: fresh_atm fresh_singleton [OF pt_tyvrs_inst at_tyvrs_inst])
@@ -375,7 +375,7 @@
     proof (induct \<Gamma>')
       case (Cons y \<Gamma>')
       thus "\<not> (\<exists>T.(VarB x' T)\<in>set(y#\<Gamma>'))" 
-	by (simp add:  fresh_fin_union[OF pt_vrs_inst  at_vrs_inst fs_vrs_inst]  
+        by (simp add:  fresh_fin_union[OF pt_vrs_inst  at_vrs_inst fs_vrs_inst]  
                        finite_vrs finite_doms, 
             auto simp add: fresh_atm fresh_singleton [OF pt_vrs_inst at_vrs_inst])
     qed (simp)
@@ -909,26 +909,26 @@
       have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact
       have lh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact      
       from rh_drv have "T=Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> \<Gamma>\<turnstile>Q\<^isub>2<:T\<^isub>2)" 
-	by (auto elim: S_ArrowE_left)
+        by (auto elim: S_ArrowE_left)
       moreover
       have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in \<Gamma>" 
-	using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)
+        using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)
       hence "(S\<^isub>1 \<rightarrow> S\<^isub>2) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp)
       moreover
       have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok)
       moreover
       { assume "\<exists>T\<^isub>1 T\<^isub>2. T = T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1 \<and> \<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2"
-	then obtain T\<^isub>1 T\<^isub>2 
-	  where T_inst: "T = T\<^isub>1 \<rightarrow> T\<^isub>2" 
-	  and   rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1"
-	  and   rh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force
-	from IH_trans[of "Q\<^isub>1"] 
-	have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>1 lh_drv_prm\<^isub>1 by simp 
-	moreover
-	from IH_trans[of "Q\<^isub>2"] 
-	have "\<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 lh_drv_prm\<^isub>2 by simp
-	ultimately have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by auto
-	then have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" using T_inst by simp
+        then obtain T\<^isub>1 T\<^isub>2 
+          where T_inst: "T = T\<^isub>1 \<rightarrow> T\<^isub>2" 
+          and   rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1"
+          and   rh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force
+        from IH_trans[of "Q\<^isub>1"] 
+        have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>1 lh_drv_prm\<^isub>1 by simp 
+        moreover
+        from IH_trans[of "Q\<^isub>2"] 
+        have "\<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 lh_drv_prm\<^isub>2 by simp
+        ultimately have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by auto
+        then have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" using T_inst by simp
       }
       ultimately show "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" by blast
     next
@@ -938,36 +938,36 @@
       have lh_drv_prm\<^isub>2: "((TVarB X Q\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact
       then have "X\<sharp>\<Gamma>" by (force dest: subtype_implies_ok simp add: valid_ty_dom_fresh)
       then have fresh_cond: "X\<sharp>\<Gamma>" "X\<sharp>Q\<^isub>1" "X\<sharp>T" using rh_drv lh_drv_prm\<^isub>1 
-	by (simp_all add: subtype_implies_fresh)
+        by (simp_all add: subtype_implies_fresh)
       from rh_drv 
       have "T = Top \<or> 
           (\<exists>T\<^isub>1 T\<^isub>2. T = (\<forall>X<:T\<^isub>1. T\<^isub>2) \<and> \<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1 \<and> ((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2)" 
-	using fresh_cond by (simp add: S_ForallE_left)
+        using fresh_cond by (simp add: S_ForallE_left)
       moreover
       have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in ((TVarB X Q\<^isub>1)#\<Gamma>)" 
-	using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)
+        using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)
       then have "(\<forall>X<:S\<^isub>1. S\<^isub>2) closed_in \<Gamma>" by (force simp add: closed_in_def ty.supp abs_supp)
       moreover
       have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok)
       moreover
       { assume "\<exists>T\<^isub>1 T\<^isub>2. T=(\<forall>X<:T\<^isub>1. T\<^isub>2) \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> ((TVarB X T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2"
-	then obtain T\<^isub>1 T\<^isub>2 
-	  where T_inst: "T = (\<forall>X<:T\<^isub>1. T\<^isub>2)" 
-	  and   rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1" 
-	  and   rh_drv_prm\<^isub>2:"((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force
-	have "(\<forall>X<:Q\<^isub>1. Q\<^isub>2) = Q" by fact 
-	then have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" 
-	  using fresh_cond by auto
-	from IH_trans[of "Q\<^isub>1"] 
-	have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using lh_drv_prm\<^isub>1 rh_drv_prm\<^isub>1 Q\<^isub>1\<^isub>2_less by blast
-	moreover
-	from IH_narrow[of "Q\<^isub>1" "[]"] 
-	have "((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" using Q\<^isub>1\<^isub>2_less lh_drv_prm\<^isub>2 rh_drv_prm\<^isub>1 by simp
-	with IH_trans[of "Q\<^isub>2"] 
-	have "((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 by simp
-	ultimately have "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)"
-	  using fresh_cond by (simp add: subtype_of.SA_all)
-	hence "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T" using T_inst by simp
+        then obtain T\<^isub>1 T\<^isub>2 
+          where T_inst: "T = (\<forall>X<:T\<^isub>1. T\<^isub>2)" 
+          and   rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1" 
+          and   rh_drv_prm\<^isub>2:"((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force
+        have "(\<forall>X<:Q\<^isub>1. Q\<^isub>2) = Q" by fact 
+        then have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" 
+          using fresh_cond by auto
+        from IH_trans[of "Q\<^isub>1"] 
+        have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using lh_drv_prm\<^isub>1 rh_drv_prm\<^isub>1 Q\<^isub>1\<^isub>2_less by blast
+        moreover
+        from IH_narrow[of "Q\<^isub>1" "[]"] 
+        have "((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" using Q\<^isub>1\<^isub>2_less lh_drv_prm\<^isub>2 rh_drv_prm\<^isub>1 by simp
+        with IH_trans[of "Q\<^isub>2"] 
+        have "((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 by simp
+        ultimately have "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)"
+          using fresh_cond by (simp add: subtype_of.SA_all)
+        hence "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T" using T_inst by simp
       }
       ultimately show "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T" by blast
     qed
@@ -985,44 +985,44 @@
     proof (induct \<Gamma>\<equiv>"\<Delta>@[(TVarB X Q)]@\<Gamma>" M N arbitrary: \<Gamma> X \<Delta> rule: subtype_of.induct) 
       case (SA_Top _ S \<Gamma> X \<Delta>)
       then have lh_drv_prm\<^isub>1: "\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok" 
-	and lh_drv_prm\<^isub>2: "S closed_in (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp_all
+        and lh_drv_prm\<^isub>2: "S closed_in (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp_all
       have rh_drv: "\<Gamma> \<turnstile> P <: Q" by fact
       hence "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)
       with lh_drv_prm\<^isub>1 have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" by (simp add: replace_type)
       moreover
       from lh_drv_prm\<^isub>2 have "S closed_in (\<Delta>@[(TVarB X P)]@\<Gamma>)" 
-	by (simp add: closed_in_def doms_append)
+        by (simp add: closed_in_def doms_append)
       ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: Top" by (simp add: subtype_of.SA_Top)
     next
       case (SA_trans_TVar Y S _ N \<Gamma> X \<Delta>) 
       then have IH_inner: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: N"
-	and lh_drv_prm: "(TVarB Y S) \<in> set (\<Delta>@[(TVarB X Q)]@\<Gamma>)"
-	and rh_drv: "\<Gamma> \<turnstile> P<:Q"
-	and ok\<^isub>Q: "\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok" by (simp_all add: subtype_implies_ok)
+        and lh_drv_prm: "(TVarB Y S) \<in> set (\<Delta>@[(TVarB X Q)]@\<Gamma>)"
+        and rh_drv: "\<Gamma> \<turnstile> P<:Q"
+        and ok\<^isub>Q: "\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok" by (simp_all add: subtype_implies_ok)
       then have ok\<^isub>P: "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" by (simp add: subtype_implies_ok) 
       show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N"
       proof (cases "X=Y")
-	case False
-	have "X\<noteq>Y" by fact
-	hence "(TVarB Y S)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" using lh_drv_prm by (simp add:binding.inject)
-	with IH_inner show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" by (simp add: subtype_of.SA_trans_TVar)
+        case False
+        have "X\<noteq>Y" by fact
+        hence "(TVarB Y S)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" using lh_drv_prm by (simp add:binding.inject)
+        with IH_inner show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" by (simp add: subtype_of.SA_trans_TVar)
       next
-	case True
-	have memb\<^isub>X\<^isub>Q: "(TVarB X Q)\<in>set (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp
-	have memb\<^isub>X\<^isub>P: "(TVarB X P)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" by simp
-	have eq: "X=Y" by fact 
-	hence "S=Q" using ok\<^isub>Q lh_drv_prm memb\<^isub>X\<^isub>Q by (simp only: uniqueness_of_ctxt)
-	hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q <: N" using IH_inner by simp
-	moreover
-	have "(\<Delta>@[(TVarB X P)]@\<Gamma>) extends \<Gamma>" by (simp add: extends_def)
-	hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening)
-	ultimately have "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: N" by (simp add: transitivity_lemma) 
-	then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by auto
+        case True
+        have memb\<^isub>X\<^isub>Q: "(TVarB X Q)\<in>set (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp
+        have memb\<^isub>X\<^isub>P: "(TVarB X P)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" by simp
+        have eq: "X=Y" by fact 
+        hence "S=Q" using ok\<^isub>Q lh_drv_prm memb\<^isub>X\<^isub>Q by (simp only: uniqueness_of_ctxt)
+        hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q <: N" using IH_inner by simp
+        moreover
+        have "(\<Delta>@[(TVarB X P)]@\<Gamma>) extends \<Gamma>" by (simp add: extends_def)
+        hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening)
+        ultimately have "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: N" by (simp add: transitivity_lemma) 
+        then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by auto
       qed
     next
       case (SA_refl_TVar _ Y \<Gamma> X \<Delta>)
       then have lh_drv_prm\<^isub>1: "\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok" 
-	and lh_drv_prm\<^isub>2: "Y \<in> ty_dom (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp_all
+        and lh_drv_prm\<^isub>2: "Y \<in> ty_dom (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp_all
       have "\<Gamma> \<turnstile> P <: Q" by fact
       hence "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)
       with lh_drv_prm\<^isub>1 have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" by (simp add: replace_type)
@@ -1036,7 +1036,7 @@
       case (SA_all _ T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \<Gamma> X \<Delta>)
       from SA_all(2,4,5,6)
       have IH_inner\<^isub>1: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> T\<^isub>1 <: S\<^isub>1" 
-	and IH_inner\<^isub>2: "(((TVarB Y T\<^isub>1)#\<Delta>)@[(TVarB X P)]@\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" by force+
+        and IH_inner\<^isub>2: "(((TVarB Y T\<^isub>1)#\<Delta>)@[(TVarB X P)]@\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" by force+
       then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> (\<forall>Y<:S\<^isub>1. S\<^isub>2) <: (\<forall>Y<:T\<^isub>1. T\<^isub>2)" by auto
     qed
   } 
@@ -1298,7 +1298,7 @@
     have "\<turnstile> (bs @ \<Delta>) ok" by simp
     moreover from Cons and valid_cons have "x \<sharp> trm_dom (bs @ \<Delta>)"
       by (simp add: doms_append finite_doms
-	fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst])
+        fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst])
     moreover from Cons and valid_cons have "T closed_in (bs @ \<Delta>)"
       by (simp add: closed_in_def doms_append)
     ultimately have "\<turnstile> (VarB x T # bs @ \<Delta>) ok"
@@ -1536,7 +1536,7 @@
     proof
       assume "TVarB Y S \<in> set D"
       then have "TVarB Y (S[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D[X \<mapsto> P]\<^sub>e)"
-	by (rule ctxt_subst_mem_TVarB)
+        by (rule ctxt_subst_mem_TVarB)
       then have "TVarB Y (S[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" by simp
       with neq and ST show ?thesis by auto
     next
@@ -1544,9 +1544,9 @@
       from X\<Gamma>_ok have "X \<sharp> ty_dom \<Gamma>" and "\<turnstile> \<Gamma> ok" by auto
       then have "X \<sharp> \<Gamma>" by (simp add: valid_ty_dom_fresh)
       with Y have "X \<sharp> S"
-	by (induct \<Gamma>) (auto simp add: fresh_list_nil fresh_list_cons)
+        by (induct \<Gamma>) (auto simp add: fresh_list_nil fresh_list_cons)
       with ST have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>S<:T[X \<mapsto> P]\<^sub>\<tau>"
-	by (simp add: type_subst_identity)
+        by (simp add: type_subst_identity)
       moreover from Y have "TVarB Y S \<in> set (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" by simp
       ultimately show ?thesis using neq by auto
     qed
--- a/src/HOL/Nominal/Examples/SN.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/Examples/SN.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -202,12 +202,12 @@
     proof (induct b rule: SN.SN.induct)
       case (SN_intro y)
       show ?case
-	apply (rule hyp)
-	apply (erule SNI')
-	apply (erule SNI')
-	apply (rule SN.SN_intro)
-	apply (erule SN_intro)+
-	done
+        apply (rule hyp)
+        apply (erule SNI')
+        apply (erule SNI')
+        apply (rule SN.SN_intro)
+        apply (erule SN_intro)+
+        done
     qed
   qed
   from b show ?thesis by (rule r)
@@ -432,62 +432,62 @@
       assume as2: "\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>"
       have "CR3_RED (App (Lam [x].t) u) \<sigma>" unfolding CR3_RED_def
       proof(intro strip)
-	fix r
-	assume red: "App (Lam [x].t) u \<longrightarrow>\<^isub>\<beta> r"
-	moreover
-	{ assume "\<exists>t'. t \<longrightarrow>\<^isub>\<beta> t' \<and> r = App (Lam [x].t') u"
-	  then obtain t' where a1: "t \<longrightarrow>\<^isub>\<beta> t'" and a2: "r = App (Lam [x].t') u" by blast
-	  have "App (Lam [x].t') u\<in>RED \<sigma>" using ih1 a1 as1 as2
-	    apply(auto)
-	    apply(drule_tac x="t'" in meta_spec)
-	    apply(simp)
-	    apply(drule meta_mp)
-	    prefer 2
-	    apply(auto)[1]
-	    apply(rule ballI)
-	    apply(drule_tac x="s" in bspec)
-	    apply(simp)
-	    apply(subgoal_tac "CR2 \<sigma>")(*A*)
-	    apply(unfold CR2_def)[1]
-	    apply(drule_tac x="t[x::=s]" in spec)
-	    apply(drule_tac x="t'[x::=s]" in spec)
-	    apply(simp add: beta_subst)
-	    (*A*)
-	    apply(simp add: RED_props)
-	    done
-	  then have "r\<in>RED \<sigma>" using a2 by simp
-	}
-	moreover
-	{ assume "\<exists>u'. u \<longrightarrow>\<^isub>\<beta> u' \<and> r = App (Lam [x].t) u'"
-	  then obtain u' where b1: "u \<longrightarrow>\<^isub>\<beta> u'" and b2: "r = App (Lam [x].t) u'" by blast
-	  have "App (Lam [x].t) u'\<in>RED \<sigma>" using ih2 b1 as1 as2
-	    apply(auto)
-	    apply(drule_tac x="u'" in meta_spec)
-	    apply(simp)
-	    apply(drule meta_mp)
-	    apply(subgoal_tac "CR2 \<tau>")
-	    apply(unfold CR2_def)[1]
-	    apply(drule_tac x="u" in spec)
-	    apply(drule_tac x="u'" in spec)
-	    apply(simp)
-	    apply(simp add: RED_props)
-	    apply(simp)
-	    done
-	  then have "r\<in>RED \<sigma>" using b2 by simp
-	}
-	moreover
-	{ assume "r = t[x::=u]"
-	  then have "r\<in>RED \<sigma>" using as1 as2 by auto
-	}
-	ultimately show "r \<in> RED \<sigma>" 
-	  (* one wants to use the strong elimination principle; for this one 
-	     has to know that x\<sharp>u *)
-	apply(cases) 
-	apply(auto simp add: lam.inject)
-	apply(drule beta_abs)
-	apply(auto)[1]
-	apply(auto simp add: alpha subst_rename)
-	done
+        fix r
+        assume red: "App (Lam [x].t) u \<longrightarrow>\<^isub>\<beta> r"
+        moreover
+        { assume "\<exists>t'. t \<longrightarrow>\<^isub>\<beta> t' \<and> r = App (Lam [x].t') u"
+          then obtain t' where a1: "t \<longrightarrow>\<^isub>\<beta> t'" and a2: "r = App (Lam [x].t') u" by blast
+          have "App (Lam [x].t') u\<in>RED \<sigma>" using ih1 a1 as1 as2
+            apply(auto)
+            apply(drule_tac x="t'" in meta_spec)
+            apply(simp)
+            apply(drule meta_mp)
+            prefer 2
+            apply(auto)[1]
+            apply(rule ballI)
+            apply(drule_tac x="s" in bspec)
+            apply(simp)
+            apply(subgoal_tac "CR2 \<sigma>")(*A*)
+            apply(unfold CR2_def)[1]
+            apply(drule_tac x="t[x::=s]" in spec)
+            apply(drule_tac x="t'[x::=s]" in spec)
+            apply(simp add: beta_subst)
+            (*A*)
+            apply(simp add: RED_props)
+            done
+          then have "r\<in>RED \<sigma>" using a2 by simp
+        }
+        moreover
+        { assume "\<exists>u'. u \<longrightarrow>\<^isub>\<beta> u' \<and> r = App (Lam [x].t) u'"
+          then obtain u' where b1: "u \<longrightarrow>\<^isub>\<beta> u'" and b2: "r = App (Lam [x].t) u'" by blast
+          have "App (Lam [x].t) u'\<in>RED \<sigma>" using ih2 b1 as1 as2
+            apply(auto)
+            apply(drule_tac x="u'" in meta_spec)
+            apply(simp)
+            apply(drule meta_mp)
+            apply(subgoal_tac "CR2 \<tau>")
+            apply(unfold CR2_def)[1]
+            apply(drule_tac x="u" in spec)
+            apply(drule_tac x="u'" in spec)
+            apply(simp)
+            apply(simp add: RED_props)
+            apply(simp)
+            done
+          then have "r\<in>RED \<sigma>" using b2 by simp
+        }
+        moreover
+        { assume "r = t[x::=u]"
+          then have "r\<in>RED \<sigma>" using as1 as2 by auto
+        }
+        ultimately show "r \<in> RED \<sigma>" 
+          (* one wants to use the strong elimination principle; for this one 
+             has to know that x\<sharp>u *)
+        apply(cases) 
+        apply(auto simp add: lam.inject)
+        apply(drule beta_abs)
+        apply(auto)[1]
+        apply(auto simp add: alpha subst_rename)
+        done
     qed
     moreover 
     have "NEUT (App (Lam [x].t) u)" unfolding NEUT_def by (auto)
--- a/src/HOL/Nominal/Examples/Standardization.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/Examples/Standardization.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -715,19 +715,19 @@
       case (App ts t)
       show ?case
       proof
-	assume "Var t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> t'"
-	then obtain rs where "ts [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs"
-	  by (iprover dest: head_Var_reduction)
-	with App show False
-	  by (induct rs arbitrary: ts) (auto del: in_listspD)
+        assume "Var t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> t'"
+        then obtain rs where "ts [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs"
+          by (iprover dest: head_Var_reduction)
+        with App show False
+          by (induct rs arbitrary: ts) (auto del: in_listspD)
       qed
     next
       case (Abs t x)
       show ?case
       proof
-	assume "(Lam [x].t) \<rightarrow>\<^sub>\<beta> t'"
-	then show False using Abs
-	  by (cases rule: beta.strong_cases) (auto simp add: abs_fresh lam.inject alpha)
+        assume "(Lam [x].t) \<rightarrow>\<^sub>\<beta> t'"
+        then show False using Abs
+          by (cases rule: beta.strong_cases) (auto simp add: abs_fresh lam.inject alpha)
       qed
     qed
   qed
@@ -747,7 +747,7 @@
     proof (cases ts)
       case Nil
       from 2 have "\<forall>u'. \<not> u \<rightarrow>\<^sub>\<beta> u'"
-	by (auto intro: apps_preserves_beta)
+        by (auto intro: apps_preserves_beta)
       then have "NF u" by (rule 2)
       then have "NF (Lam [x].u)" by (rule NF.Abs)
       with Nil show ?thesis by simp
@@ -755,9 +755,9 @@
       case (Cons r rs)
       have "(Lam [x].u) \<degree> r \<rightarrow>\<^sub>\<beta> u[x::=r]" ..
       then have "(Lam [x].u) \<degree> r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> u[x::=r] \<degree>\<degree> rs"
-	by (rule apps_preserves_beta)
+        by (rule apps_preserves_beta)
       with Cons have "(Lam [x].u) \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> u[x::=r] \<degree>\<degree> rs"
-	by simp
+        by simp
       with 2 show ?thesis by iprover
     qed
   qed
--- a/src/HOL/Nominal/Examples/Weakening.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/Examples/Weakening.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,3 @@
-(* $Id$ *)
-
 theory Weakening 
   imports "../Nominal" 
 begin
@@ -180,7 +178,7 @@
     proof -
       have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact
       then have "[(c,x)]\<bullet>((x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2))" using fc0 fc1 
-	by (perm_simp add: eqvts calc_atm perm_fresh_fresh ty_fresh)
+        by (perm_simp add: eqvts calc_atm perm_fresh_fresh ty_fresh)
       then show "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2)" by (rule perm_boolE)
     qed
     moreover 
@@ -188,7 +186,7 @@
     proof -
       have "valid \<Gamma>2" by fact
       then show "valid ((x,T1)#([(c,x)]\<bullet>\<Gamma>2))" using fc1
-	by (auto intro!: v2 simp add: fresh_left calc_atm eqvts)
+        by (auto intro!: v2 simp add: fresh_left calc_atm eqvts)
     qed
     (* these two facts give us by induction hypothesis the following *)
     ultimately have "(x,T1)#([(c,x)]\<bullet>\<Gamma>2) \<turnstile> t : T2" using ih by simp 
--- a/src/HOL/Nominal/Nominal.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/Nominal.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -2659,7 +2659,7 @@
     obtain y::"'a" where fr: "y\<sharp>(c,pi\<bullet>Xs,As)" 
       apply(rule_tac at_exists_fresh[OF at, where x="(c,pi\<bullet>Xs,As)"])
       apply(auto simp add: supp_prod at_supp[OF at] at_fin_set_supp[OF at]
-	pt_supp_finite_pi[OF pt_fun_inst[OF at_pt_inst[OF at] pt_bool_inst at] at])
+        pt_supp_finite_pi[OF pt_fun_inst[OF at_pt_inst[OF at] pt_bool_inst at] at])
       done
     have "({y}\<union>(pi\<bullet>Xs))\<sharp>*c" using a1 fr by (simp add: fresh_star_def)
     moreover
@@ -2674,21 +2674,21 @@
     proof -
       have eq: "[(pi\<bullet>x,y)]\<bullet>(pi\<bullet>Xs) = (pi\<bullet>Xs)" 
       proof -
-	have "(pi\<bullet>x)\<sharp>(pi\<bullet>Xs)" using b d2 
-	  by(simp add: pt_fresh_bij[OF pt_fun_inst, OF at_pt_inst[OF at], 
+        have "(pi\<bullet>x)\<sharp>(pi\<bullet>Xs)" using b d2 
+          by(simp add: pt_fresh_bij[OF pt_fun_inst, OF at_pt_inst[OF at], 
             OF pt_bool_inst, OF at, OF at]
             at_fin_set_fresh[OF at])
-	moreover
-	have "y\<sharp>(pi\<bullet>Xs)" using fr by simp
-	ultimately show "[(pi\<bullet>x,y)]\<bullet>(pi\<bullet>Xs) = (pi\<bullet>Xs)" 
-	  by (simp add: pt_fresh_fresh[OF pt_fun_inst, 
+        moreover
+        have "y\<sharp>(pi\<bullet>Xs)" using fr by simp
+        ultimately show "[(pi\<bullet>x,y)]\<bullet>(pi\<bullet>Xs) = (pi\<bullet>Xs)" 
+          by (simp add: pt_fresh_fresh[OF pt_fun_inst, 
             OF at_pt_inst[OF at], OF pt_bool_inst, OF at, OF at])
       qed
       have "(((pi\<bullet>x,y)#pi)\<bullet>({x}\<union>Xs)) = ([(pi\<bullet>x,y)]\<bullet>(pi\<bullet>({x}\<union>Xs)))"
-	by (simp add: pt2[symmetric, OF pt_fun_inst, OF at_pt_inst[OF at], 
+        by (simp add: pt2[symmetric, OF pt_fun_inst, OF at_pt_inst[OF at], 
           OF pt_bool_inst, OF at])
       also have "\<dots> = {y}\<union>([(pi\<bullet>x,y)]\<bullet>(pi\<bullet>Xs))" 
-	by (simp only: union_eqvt perm_set_eq[OF at_pt_inst[OF at], OF at] at_calc[OF at])(auto)
+        by (simp only: union_eqvt perm_set_eq[OF at_pt_inst[OF at], OF at] at_calc[OF at])(auto)
       finally show "(((pi\<bullet>x,y)#pi)\<bullet>({x} \<union> Xs)) = {y}\<union>(pi\<bullet>Xs)" using eq by simp
     qed
     ultimately 
@@ -2795,30 +2795,30 @@
       assume a3: "(a::'x)\<sharp>h"
       show "h (a::'x) = h a0"
       proof (cases "a=a0")
-	case True thus "h (a::'x) = h a0" by simp
+        case True thus "h (a::'x) = h a0" by simp
       next
-	case False 
-	assume "a\<noteq>a0"
-	hence c1: "a\<notin>((supp a0)::'x set)" by  (simp add: fresh_def[symmetric] at_fresh[OF at])
-	have c2: "a\<notin>((supp h)::'x set)" using a3 by (simp add: fresh_def)
-	from c1 c2 have c3: "a\<notin>((supp h)\<union>((supp a0)::'x set))" by force
-	have f2: "finite ((supp a0)::'x set)" by (simp add: at_supp[OF at])
-	from f1 f2 have "((supp (h a0))::'x set)\<subseteq>((supp h)\<union>(supp a0))"
-	  by (simp add: pt_supp_fun_subset[OF ptb, OF pta, OF at])
-	hence "a\<notin>((supp (h a0))::'x set)" using c3 by force
-	hence "a\<sharp>(h a0)" by (simp add: fresh_def) 
-	with a2 have d1: "[(a0,a)]\<bullet>(h a0) = (h a0)" by (rule pt_fresh_fresh[OF pta, OF at])
-	from a1 a3 have d2: "[(a0,a)]\<bullet>h = h" by (rule pt_fresh_fresh[OF ptc, OF at])
-	from d1 have "h a0 = [(a0,a)]\<bullet>(h a0)" by simp
-	also have "\<dots>= ([(a0,a)]\<bullet>h)([(a0,a)]\<bullet>a0)" by (simp add: pt_fun_app_eq[OF ptb, OF at])
-	also have "\<dots> = h ([(a0,a)]\<bullet>a0)" using d2 by simp
-	also have "\<dots> = h a" by (simp add: at_calc[OF at])
-	finally show "h a = h a0" by simp
+        case False 
+        assume "a\<noteq>a0"
+        hence c1: "a\<notin>((supp a0)::'x set)" by  (simp add: fresh_def[symmetric] at_fresh[OF at])
+        have c2: "a\<notin>((supp h)::'x set)" using a3 by (simp add: fresh_def)
+        from c1 c2 have c3: "a\<notin>((supp h)\<union>((supp a0)::'x set))" by force
+        have f2: "finite ((supp a0)::'x set)" by (simp add: at_supp[OF at])
+        from f1 f2 have "((supp (h a0))::'x set)\<subseteq>((supp h)\<union>(supp a0))"
+          by (simp add: pt_supp_fun_subset[OF ptb, OF pta, OF at])
+        hence "a\<notin>((supp (h a0))::'x set)" using c3 by force
+        hence "a\<sharp>(h a0)" by (simp add: fresh_def) 
+        with a2 have d1: "[(a0,a)]\<bullet>(h a0) = (h a0)" by (rule pt_fresh_fresh[OF pta, OF at])
+        from a1 a3 have d2: "[(a0,a)]\<bullet>h = h" by (rule pt_fresh_fresh[OF ptc, OF at])
+        from d1 have "h a0 = [(a0,a)]\<bullet>(h a0)" by simp
+        also have "\<dots>= ([(a0,a)]\<bullet>h)([(a0,a)]\<bullet>a0)" by (simp add: pt_fun_app_eq[OF ptb, OF at])
+        also have "\<dots> = h ([(a0,a)]\<bullet>a0)" using d2 by simp
+        also have "\<dots> = h a" by (simp add: at_calc[OF at])
+        finally show "h a = h a0" by simp
       qed
     qed
   qed
 qed
-	    
+
 lemma freshness_lemma_unique:
   fixes h :: "'x\<Rightarrow>'a"
   assumes pt: "pt TYPE('a) TYPE('x)"
@@ -3090,44 +3090,44 @@
       have "(c=a) \<or> (c=b) \<or> (c\<noteq>a \<and> c\<noteq>b)" by blast
       moreover  --"case c=a"
       { have "nSome(x) = nSome([(a,b)]\<bullet>y)" using a2 by simp
-	also have "\<dots> = nSome([(b,a)]\<bullet>y)" by (simp, rule pt3[OF pt], rule at_ds5[OF at])
-	finally have "nSome(x) = nSome([(b,a)]\<bullet>y)" by simp
-	moreover
-	assume "c=a"
-	ultimately have "?LHS=?RHS" using a1 a3 by simp
+        also have "\<dots> = nSome([(b,a)]\<bullet>y)" by (simp, rule pt3[OF pt], rule at_ds5[OF at])
+        finally have "nSome(x) = nSome([(b,a)]\<bullet>y)" by simp
+        moreover
+        assume "c=a"
+        ultimately have "?LHS=?RHS" using a1 a3 by simp
       }
       moreover  -- "case c=b"
       { have a4: "y=[(a,b)]\<bullet>x" using a2 by (simp only: pt_swap_bij[OF pt, OF at])
-	hence "a\<sharp>([(a,b)]\<bullet>x)" using a3 by simp
-	hence "b\<sharp>x" by (simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])
-	moreover
-	assume "c=b"
-	ultimately have "?LHS=?RHS" using a1 a4 by simp
+        hence "a\<sharp>([(a,b)]\<bullet>x)" using a3 by simp
+        hence "b\<sharp>x" by (simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])
+        moreover
+        assume "c=b"
+        ultimately have "?LHS=?RHS" using a1 a4 by simp
       }
       moreover  -- "case c\<noteq>a \<and> c\<noteq>b"
       { assume a5: "c\<noteq>a \<and> c\<noteq>b"
-	moreover 
-	have "c\<sharp>x = c\<sharp>y" using a2 a5 by (force simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])
-	moreover 
-	have "c\<sharp>y \<longrightarrow> [(a,c)]\<bullet>x = [(b,c)]\<bullet>y" 
-	proof (intro strip)
-	  assume a6: "c\<sharp>y"
-	  have "[(a,c),(b,c),(a,c)] \<triangleq> [(a,b)]" using a1 a5 by (force intro: at_ds3[OF at])
-	  hence "[(a,c)]\<bullet>([(b,c)]\<bullet>([(a,c)]\<bullet>y)) = [(a,b)]\<bullet>y" 
-	    by (simp add: pt2[OF pt, symmetric] pt3[OF pt])
- 	  hence "[(a,c)]\<bullet>([(b,c)]\<bullet>y) = [(a,b)]\<bullet>y" using a3 a6 
-	    by (simp add: pt_fresh_fresh[OF pt, OF at])
-	  hence "[(a,c)]\<bullet>([(b,c)]\<bullet>y) = x" using a2 by simp
-	  hence "[(b,c)]\<bullet>y = [(a,c)]\<bullet>x" by (drule_tac pt_bij1[OF pt, OF at], simp)
-	  thus "[(a,c)]\<bullet>x = [(b,c)]\<bullet>y" by simp
-	qed
-	ultimately have "?LHS=?RHS" by simp
+        moreover 
+        have "c\<sharp>x = c\<sharp>y" using a2 a5 by (force simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])
+        moreover 
+        have "c\<sharp>y \<longrightarrow> [(a,c)]\<bullet>x = [(b,c)]\<bullet>y" 
+        proof (intro strip)
+          assume a6: "c\<sharp>y"
+          have "[(a,c),(b,c),(a,c)] \<triangleq> [(a,b)]" using a1 a5 by (force intro: at_ds3[OF at])
+          hence "[(a,c)]\<bullet>([(b,c)]\<bullet>([(a,c)]\<bullet>y)) = [(a,b)]\<bullet>y" 
+            by (simp add: pt2[OF pt, symmetric] pt3[OF pt])
+          hence "[(a,c)]\<bullet>([(b,c)]\<bullet>y) = [(a,b)]\<bullet>y" using a3 a6 
+            by (simp add: pt_fresh_fresh[OF pt, OF at])
+          hence "[(a,c)]\<bullet>([(b,c)]\<bullet>y) = x" using a2 by simp
+          hence "[(b,c)]\<bullet>y = [(a,c)]\<bullet>x" by (drule_tac pt_bij1[OF pt, OF at], simp)
+          thus "[(a,c)]\<bullet>x = [(b,c)]\<bullet>y" by simp
+        qed
+        ultimately have "?LHS=?RHS" by simp
       }
       ultimately show "?LHS = ?RHS" by blast
     qed
   qed
 qed
-	
+        
 (* alpha equivalence *)
 lemma abs_fun_eq: 
   fixes x  :: "'a"
@@ -3278,7 +3278,7 @@
     have "\<exists>c::'x. c\<sharp>(b,a,x,[a].x)" 
     proof (rule at_exists_fresh'[OF at], auto simp add: supp_prod at_supp[OF at] f)
       show "finite ((supp ([a].x))::'x set)" using f
-	by (simp add: abs_fun_finite_supp[OF pt, OF at])	
+        by (simp add: abs_fun_finite_supp[OF pt, OF at])        
     qed
     then obtain c where fr1: "c\<noteq>b"
                   and   fr2: "c\<noteq>a"
@@ -3309,7 +3309,7 @@
   have "\<exists>c::'x. c\<sharp>(b,a,x,[a].x)"
   proof (rule at_exists_fresh'[OF at], auto simp add: supp_prod at_supp[OF at] f)
     show "finite ((supp ([a].x))::'x set)" using f
-      by (simp add: abs_fun_finite_supp[OF pt, OF at])	
+      by (simp add: abs_fun_finite_supp[OF pt, OF at])  
   qed
   then obtain c where fr1: "b\<noteq>c"
                 and   fr2: "c\<noteq>a"
--- a/src/HOL/Nominal/nominal_atoms.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/nominal_atoms.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,4 +1,4 @@
-(*  title:      HOL/Nominal/nominal_atoms.ML
+(*  Title:      HOL/Nominal/nominal_atoms.ML
     Author:     Christian Urban and Stefan Berghofer, TU Muenchen
 
 Declaration of atom types to be used in nominal datatypes.
@@ -820,7 +820,7 @@
              (* produces a list of theorems of the form [t1 RS thm,..,tn RS thm] *) 
              fun instR thm thms = map (fn ti => ti RS thm) thms;
 
-	     (* takes a theorem thm and a list of theorems [(t1,m1),..,(tn,mn)]  *)
+             (* takes a theorem thm and a list of theorems [(t1,m1),..,(tn,mn)]  *)
              (* produces a list of theorems of the form [[t1,m1] MRS thm,..,[tn,mn] MRS thm] *) 
              fun instRR thm thms = map (fn (ti,mi) => [ti,mi] MRS thm) thms;
 
@@ -883,7 +883,7 @@
             ||>> add_thmss_string [(("alpha_fresh'", inst_pt_at [abs_fun_fresh']),[])]
             ||>> add_thmss_string [(("perm_swap", inst_pt_at [pt_swap_bij] @ inst_pt_at [pt_swap_bij']),[])]
             ||>> add_thmss_string 
-	      let val thms1 = inst_at at_swap_simps
+              let val thms1 = inst_at at_swap_simps
                   and thms2 = inst_dj [dj_perm_forget]
               in [(("swap_simps", thms1 @ thms2),[])] end 
             ||>> add_thmss_string 
--- a/src/HOL/Nominal/nominal_datatype.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -153,7 +153,7 @@
 
 fun projections rule =
   Project_Rule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
-  |> map (standard #> RuleCases.save rule);
+  |> map (Drule.standard #> RuleCases.save rule);
 
 val supp_prod = thm "supp_prod";
 val fresh_prod = thm "fresh_prod";
@@ -313,7 +313,7 @@
 
     val unfolded_perm_eq_thms =
       if length descr = length new_type_names then []
-      else map standard (List.drop (split_conj_thm
+      else map Drule.standard (List.drop (split_conj_thm
         (Goal.prove_global thy2 [] []
           (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
             (map (fn (c as (s, T), x) =>
@@ -333,7 +333,7 @@
 
     val perm_empty_thms = maps (fn a =>
       let val permT = mk_permT (Type (a, []))
-      in map standard (List.take (split_conj_thm
+      in map Drule.standard (List.take (split_conj_thm
         (Goal.prove_global thy2 [] []
           (augment_sort thy2 [pt_class_of thy2 a]
             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -365,7 +365,7 @@
         val pt_inst = pt_inst_of thy2 a;
         val pt2' = pt_inst RS pt2;
         val pt2_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a);
-      in List.take (map standard (split_conj_thm
+      in List.take (map Drule.standard (split_conj_thm
         (Goal.prove_global thy2 [] []
            (augment_sort thy2 [pt_class_of thy2 a]
              (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -400,7 +400,7 @@
         val pt3' = pt_inst RS pt3;
         val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
         val pt3_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a);
-      in List.take (map standard (split_conj_thm
+      in List.take (map Drule.standard (split_conj_thm
         (Goal.prove_global thy2 [] []
           (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
              (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
@@ -585,7 +585,7 @@
       (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
       (perm_indnames ~~ descr);
 
-    fun mk_perm_closed name = map (fn th => standard (th RS mp))
+    fun mk_perm_closed name = map (fn th => Drule.standard (th RS mp))
       (List.take (split_conj_thm (Goal.prove_global thy4 [] []
         (augment_sort thy4
           (pt_class_of thy4 name :: map (cp_class_of thy4 name) (dt_atoms \ name))
@@ -810,7 +810,7 @@
       let
         val rep_const = cterm_of thy
           (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
-        val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
+        val dist = Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
         val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
           ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax)
       in
@@ -874,7 +874,7 @@
           let
             val dist_thm = Goal.prove_global thy8 [] [] t (fn _ =>
               simp_tac (global_simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)
-          in dist_thm :: standard (dist_thm RS not_sym) ::
+          in dist_thm :: Drule.standard (dist_thm RS not_sym) ::
             prove_distinct_thms p (k, ts)
           end;
 
@@ -1089,7 +1089,7 @@
 
     val finite_supp_thms = map (fn atom =>
       let val atomT = Type (atom, [])
-      in map standard (List.take
+      in map Drule.standard (List.take
         (split_conj_thm (Goal.prove_global thy8 [] []
            (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
              (HOLogic.mk_Trueprop
@@ -1107,7 +1107,7 @@
 
     val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
 
-	(* Function to add both the simp and eqvt attributes *)
+        (* Function to add both the simp and eqvt attributes *)
         (* These two attributes are duplicated on all the types in the mutual nominal datatypes *)
 
     val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add];
@@ -1535,7 +1535,7 @@
           in
             (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
           end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
-        val ths = map (fn th => standard (th RS mp)) (split_conj_thm
+        val ths = map (fn th => Drule.standard (th RS mp)) (split_conj_thm
           (Goal.prove_global thy11 [] []
             (augment_sort thy1 pt_cp_sort
               (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
@@ -1567,7 +1567,7 @@
           (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
             (rec_fns ~~ rec_fn_Ts)
       in
-        map (fn th => standard (th RS mp)) (split_conj_thm
+        map (fn th => Drule.standard (th RS mp)) (split_conj_thm
           (Goal.prove_global thy11 []
             (map (augment_sort thy11 fs_cp_sort) fins)
             (augment_sort thy11 fs_cp_sort
@@ -1610,7 +1610,7 @@
             val y = Free ("y", U);
             val y' = Free ("y'", U)
           in
-            standard (Goal.prove (ProofContext.init thy11) []
+            Drule.standard (Goal.prove (ProofContext.init thy11) []
               (map (augment_sort thy11 fs_cp_sort)
                 (finite_prems @
                    [HOLogic.mk_Trueprop (R $ x $ y),
@@ -2055,7 +2055,7 @@
          ((Binding.name "rec_equiv'", flat rec_equiv_thms'), []),
          ((Binding.name "rec_fin_supp", flat rec_fin_supp_thms), []),
          ((Binding.name "rec_fresh", flat rec_fresh_thms), []),
-         ((Binding.name "rec_unique", map standard rec_unique_thms), []),
+         ((Binding.name "rec_unique", map Drule.standard rec_unique_thms), []),
          ((Binding.name "recs", rec_thms), [])] ||>
       Sign.parent_path ||>
       map_nominal_datatypes (fold Symtab.update dt_infos);
--- a/src/HOL/Nominal/nominal_permeq.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/nominal_permeq.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,9 @@
 (*  Title:      HOL/Nominal/nominal_permeq.ML
-    Authors:    Christian Urban, Julien Narboux, TU Muenchen
+    Author:     Christian Urban, TU Muenchen
+    Author:     Julien Narboux, TU Muenchen
 
-Methods for simplifying permutations and
-for analysing equations involving permutations.
+Methods for simplifying permutations and for analysing equations
+involving permutations. 
 *)
 
 (*
@@ -99,7 +100,7 @@
     (* constant or when (f x) is a permuation with two or more arguments     *)
     fun applicable_app t = 
           (case (strip_comb t) of
-	      (Const ("Nominal.perm",_),ts) => (length ts) >= 2
+              (Const ("Nominal.perm",_),ts) => (length ts) >= 2
             | (Const _,_) => false
             | _ => true)
   in
@@ -126,9 +127,9 @@
      fun applicable_fun t =
        (case (strip_comb t) of
           (Abs _ ,[]) => true
-	| (Const ("Nominal.perm",_),_) => false
+        | (Const ("Nominal.perm",_),_) => false
         | (Const _, _) => true
-	| _ => false)
+        | _ => false)
    in
      case redex of 
        (* case pi o f == (%x. pi o (f ((rev pi)o x))) *)     
@@ -159,14 +160,14 @@
 fun perm_simp stac ss = 
     let val simps = ["perm_swap","perm_fresh_fresh","perm_bij","perm_pi_simp","swap_simps"]
     in 
-	perm_simp_gen stac simps [] ss
+        perm_simp_gen stac simps [] ss
     end;
 
 fun eqvt_simp stac ss = 
     let val simps = ["perm_swap","perm_fresh_fresh","perm_pi_simp"]
-	val eqvts_thms = NominalThmDecls.get_eqvt_thms (Simplifier.the_context ss);
+        val eqvts_thms = NominalThmDecls.get_eqvt_thms (Simplifier.the_context ss);
     in 
-	perm_simp_gen stac simps eqvts_thms ss
+        perm_simp_gen stac simps eqvts_thms ss
     end;
 
 
@@ -252,15 +253,15 @@
 (* to 10 - this seems to be sufficient in most cases              *)
 fun perm_extend_simp_tac_i tactical ss =
   let fun perm_extend_simp_tac_aux tactical ss n = 
-	  if n=0 then K all_tac
-	  else DETERM o 
-	       (FIRST'[fn i => tactical ("splitting conjunctions on the rhs", rtac conjI i),
+          if n=0 then K all_tac
+          else DETERM o 
+               (FIRST'[fn i => tactical ("splitting conjunctions on the rhs", rtac conjI i),
                        fn i => tactical (perm_simp asm_full_simp_tac ss i),
-		       fn i => tactical (perm_compose_tac ss i),
-		       fn i => tactical (apply_cong_tac i), 
+                       fn i => tactical (perm_compose_tac ss i),
+                       fn i => tactical (apply_cong_tac i), 
                        fn i => tactical (unfold_perm_fun_def_tac i),
                        fn i => tactical (ext_fun_tac i)]
-		      THEN_ALL_NEW (TRY o (perm_extend_simp_tac_aux tactical ss (n-1))))
+                      THEN_ALL_NEW (TRY o (perm_extend_simp_tac_aux tactical ss (n-1))))
   in perm_extend_simp_tac_aux tactical ss 10 end;
 
 
@@ -329,10 +330,10 @@
 (* support of these free variables (op supports) the goal          *)
 fun fresh_guess_tac_i tactical ss i st =
     let 
-	val goal = List.nth(cprems_of st, i-1)
+        val goal = List.nth(cprems_of st, i-1)
         val fin_supp = dynamic_thms st ("fin_supp")
         val fresh_atm = dynamic_thms st ("fresh_atm")
-	val ss1 = ss addsimps [symmetric fresh_def,fresh_prod,fresh_unit,conj_absorb,not_false]@fresh_atm
+        val ss1 = ss addsimps [symmetric fresh_def,fresh_prod,fresh_unit,conj_absorb,not_false]@fresh_atm
         val ss2 = ss addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp
     in
       case Logic.strip_assums_concl (term_of goal) of
--- a/src/HOL/Nominal/nominal_thmdecls.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Nominal/nominal_thmdecls.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,12 +1,13 @@
-(* Authors: Julien Narboux and Christian Urban
-
-   This file introduces the infrastructure for the lemma
-   collection "eqvts".
+(*  Title:      HOL/Nominal/nominal_thmdecls.ML
+    Author:     Julien Narboux, TU Muenchen
+    Author:     Christian Urban, TU Muenchen
 
-   By attaching [eqvt] or [eqvt_force] to a lemma, it will get 
-   stored in a data-slot in the context. Possible modifiers
-   are [... add] and [... del] for adding and deleting, 
-   respectively, the lemma from the data-slot.
+Infrastructure for the lemma collection "eqvts".
+
+By attaching [eqvt] or [eqvt_force] to a lemma, it will get stored in
+a data-slot in the context. Possible modifiers are [... add] and
+[... del] for adding and deleting, respectively, the lemma from the
+data-slot.
 *)
 
 signature NOMINAL_THMDECLS =
@@ -58,11 +59,11 @@
   val perm_pi_simp = PureThy.get_thms ctxt "perm_pi_simp"
 in
   EVERY1 [tactic ("iffI applied", rtac @{thm iffI}),
-	  tactic ("remove pi with perm_boolE", dtac @{thm perm_boolE}),
+          tactic ("remove pi with perm_boolE", dtac @{thm perm_boolE}),
           tactic ("solve with orig_thm", etac orig_thm),
           tactic ("applies orig_thm instantiated with rev pi",
                      dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm)),
-	  tactic ("getting rid of the pi on the right", rtac @{thm perm_boolI}),
+          tactic ("getting rid of the pi on the right", rtac @{thm perm_boolI}),
           tactic ("getting rid of all remaining perms",
                      full_simp_tac (HOL_basic_ss addsimps perm_pi_simp))]
 end;
--- a/src/HOL/Old_Number_Theory/EulerFermat.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Old_Number_Theory/EulerFermat.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -142,7 +142,7 @@
   apply (cut_tac a = a and m = m in zcong_zless_unique, auto)
    apply (rule_tac [2] m = m in zcong_zless_imp_eq)
        apply (auto intro: Bnor_mem_zle Bnor_mem_zg zcong_trans
-	 order_less_imp_le norR_mem_unique_aux simp add: zcong_sym)
+         order_less_imp_le norR_mem_unique_aux simp add: zcong_sym)
   apply (rule_tac x = b in exI, safe)
   apply (rule Bnor_mem_if)
     apply (case_tac [2] "b = 0")
@@ -312,7 +312,7 @@
      apply (rule_tac [2] RRset2norRR_correct1)
        apply (rule_tac [5] RRset2norRR_inj)
         apply (auto intro: order_less_le [THEN iffD2]
-	   simp add: noX_is_RRset)
+           simp add: noX_is_RRset)
   apply (unfold noXRRset_def norRRset_def)
   apply (rule finite_imageI)
   apply (rule Bnor_fin)
--- a/src/HOL/Old_Number_Theory/IntPrimes.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Old_Number_Theory/IntPrimes.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -26,10 +26,10 @@
   "measure ((\<lambda>(m, n, r', r, s', s, t', t). nat r)
     :: int * int * int * int *int * int * int * int => nat)"
   "xzgcda (m, n, r', r, s', s, t', t) =
-	(if r \<le> 0 then (r', s', t')
-	 else xzgcda (m, n, r, r' mod r, 
-		      s, s' - (r' div r) * s, 
-		      t, t' - (r' div r) * t))"
+        (if r \<le> 0 then (r', s', t')
+         else xzgcda (m, n, r, r' mod r, 
+                      s, s' - (r' div r) * s, 
+                      t, t' - (r' div r) * t))"
 
 definition
   zprime :: "int \<Rightarrow> bool" where
--- a/src/HOL/Old_Number_Theory/Legacy_GCD.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Old_Number_Theory/Legacy_GCD.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,4 +1,4 @@
-(*  Title:      HOL/GCD.thy
+(*  Title:      HOL/Old_Number_Theory/Legacy_GCD.thy
     Author:     Christophe Tabacznyj and Lawrence C Paulson
     Copyright   1996  University of Cambridge
 *)
@@ -331,31 +331,31 @@
      moreover
      {assume db: "d=b"
        from prems have ?thesis apply simp
-	 apply (rule exI[where x = b], simp)
-	 apply (rule exI[where x = b])
-	by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
+         apply (rule exI[where x = b], simp)
+         apply (rule exI[where x = b])
+        by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
     moreover
     {assume db: "d < b" 
-	{assume "x=0" hence ?thesis  using prems by simp }
-	moreover
-	{assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
-	  
-	  from db have "d \<le> b - 1" by simp
-	  hence "d*b \<le> b*(b - 1)" by simp
-	  with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
-	  have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
-	  from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra
-	  hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
-	  hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" 
-	    by (simp only: diff_add_assoc[OF dble, of d, symmetric])
-	  hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
-	    by (simp only: diff_mult_distrib2 add_commute mult_ac)
-	  hence ?thesis using H(1,2)
-	    apply -
-	    apply (rule exI[where x=d], simp)
-	    apply (rule exI[where x="(b - 1) * y"])
-	    by (rule exI[where x="x*(b - 1) - d"], simp)}
-	ultimately have ?thesis by blast}
+        {assume "x=0" hence ?thesis  using prems by simp }
+        moreover
+        {assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
+          
+          from db have "d \<le> b - 1" by simp
+          hence "d*b \<le> b*(b - 1)" by simp
+          with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
+          have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
+          from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra
+          hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
+          hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" 
+            by (simp only: diff_add_assoc[OF dble, of d, symmetric])
+          hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
+            by (simp only: diff_mult_distrib2 add_commute mult_ac)
+          hence ?thesis using H(1,2)
+            apply -
+            apply (rule exI[where x=d], simp)
+            apply (rule exI[where x="(b - 1) * y"])
+            by (rule exI[where x="x*(b - 1) - d"], simp)}
+        ultimately have ?thesis by blast}
     ultimately have ?thesis by blast}
   ultimately have ?thesis by blast}
  ultimately show ?thesis by blast
--- a/src/HOL/Old_Number_Theory/Pocklington.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Old_Number_Theory/Pocklington.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -36,11 +36,11 @@
   {assume p0: "p\<noteq>0" "p\<noteq>1"
     {assume H: "?lhs"
       {fix m assume m: "m > 0" "m < p"
-	{assume "m=1" hence "coprime p m" by simp}
-	moreover
-	{assume "p dvd m" hence "p \<le> m" using dvd_imp_le m by blast with m(2)
-	  have "coprime p m" by simp}
-	ultimately have "coprime p m" using prime_coprime[OF H, of m] by blast}
+        {assume "m=1" hence "coprime p m" by simp}
+        moreover
+        {assume "p dvd m" hence "p \<le> m" using dvd_imp_le m by blast with m(2)
+          have "coprime p m" by simp}
+        ultimately have "coprime p m" using prime_coprime[OF H, of m] by blast}
       hence ?rhs using p0 by auto}
     moreover
     { assume H: "\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m"
@@ -50,8 +50,8 @@
       {assume "q = p" hence ?lhs using q(1) by blast}
       moreover
       {assume "q\<noteq>p" with qp have qplt: "q < p" by arith
-	from H[rule_format, of q] qplt q0 have "coprime p q" by arith
-	with coprime_prime[of p q q] q have False by simp hence ?lhs by blast}
+        from H[rule_format, of q] qplt q0 have "coprime p q" by arith
+        with coprime_prime[of p q q] q have False by simp hence ?lhs by blast}
       ultimately have ?lhs by blast}
     ultimately have ?thesis by blast}
   ultimately show ?thesis  by (cases"p=0 \<or> p=1", auto)
@@ -93,19 +93,19 @@
   {assume az: "a\<noteq>0"
     {assume xy: "x \<le> y" hence axy': "a*x \<le> a*y" by simp
       with axy cong_sub_cases[of "a*x" "a*y" n]  have "[a*(y - x) = 0] (mod n)"
-	by (simp only: if_True diff_mult_distrib2)
+        by (simp only: if_True diff_mult_distrib2)
       hence th: "n dvd a*(y -x)" by simp
       from coprime_divprod[OF th] an have "n dvd y - x"
-	by (simp add: coprime_commute)
+        by (simp add: coprime_commute)
       hence ?thesis using xy cong_sub_cases[of x y n] by simp}
     moreover
     {assume H: "\<not>x \<le> y" hence xy: "y \<le> x"  by arith
       from H az have axy': "\<not> a*x \<le> a*y" by auto
       with axy H cong_sub_cases[of "a*x" "a*y" n]  have "[a*(x - y) = 0] (mod n)"
-	by (simp only: if_False diff_mult_distrib2)
+        by (simp only: if_False diff_mult_distrib2)
       hence th: "n dvd a*(x - y)" by simp
       from coprime_divprod[OF th] an have "n dvd x - y"
-	by (simp add: coprime_commute)
+        by (simp add: coprime_commute)
       hence ?thesis using xy cong_sub_cases[of x y n] by simp}
     ultimately have ?thesis by blast}
   ultimately show ?thesis by blast
@@ -409,10 +409,10 @@
   {assume n: "n\<noteq>0" "n\<noteq>1"
     {fix m
       from n have "0 < m \<and> m <= n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
-	apply (cases "m = 0", simp_all)
-	apply (cases "m = 1", simp_all)
-	apply (cases "m = n", auto)
-	done }
+        apply (cases "m = 0", simp_all)
+        apply (cases "m = 1", simp_all)
+        apply (cases "m = n", auto)
+        done }
     hence ?thesis unfolding phi_def by simp}
   ultimately show ?thesis by auto
 qed
@@ -485,19 +485,19 @@
       hence ceq: "card ?S' = card ?S"
       using n finite_number_segment[of n] phi_another[OF n(2)] by simp
       {fix m assume m: "0 < m" "m < n" "\<not> coprime m n"
-	hence mS': "m \<notin> ?S'" by auto
-	have "insert m ?S' \<le> ?S" using m by auto
-	from m have "card (insert m ?S') \<le> card ?S"
-	  by - (rule card_mono[of ?S "insert m ?S'"], auto)
-	hence False
-	  unfolding card_insert_disjoint[of "?S'" m, OF fS' mS'] ceq
-	  by simp }
+        hence mS': "m \<notin> ?S'" by auto
+        have "insert m ?S' \<le> ?S" using m by auto
+        from m have "card (insert m ?S') \<le> card ?S"
+          by - (rule card_mono[of ?S "insert m ?S'"], auto)
+        hence False
+          unfolding card_insert_disjoint[of "?S'" m, OF fS' mS'] ceq
+          by simp }
       hence "\<forall>m. 0 <m \<and> m < n \<longrightarrow> coprime m n" by blast
       hence "prime n" unfolding prime using n by (simp add: coprime_commute)}
     moreover
     {assume H: "prime n"
       hence "?S = ?S'" unfolding prime using n
-	by (auto simp add: coprime_commute)
+        by (auto simp add: coprime_commute)
       hence "card ?S = card ?S'" by simp
       hence "\<phi> n = n - 1" unfolding phi_another[OF n(2)] by simp}
     ultimately have ?thesis using n by blast}
@@ -520,29 +520,29 @@
     have eq: "?f ` (?S (a*b)) = (?S a \<times> ?S b)"
     proof-
       {fix x assume x:"x \<in> ?S (a*b)"
-	hence x': "coprime x (a*b)" "x < a*b" by simp_all
-	hence xab: "coprime x a" "coprime x b" by (simp_all add: coprime_mul_eq)
-	from mod_less_divisor a b have xab':"x mod a < a" "x mod b < b" by auto
-	from xab xab' have "?f x \<in> (?S a \<times> ?S b)"
-	  by (simp add: coprime_mod[OF a(1)] coprime_mod[OF b(1)])}
+        hence x': "coprime x (a*b)" "x < a*b" by simp_all
+        hence xab: "coprime x a" "coprime x b" by (simp_all add: coprime_mul_eq)
+        from mod_less_divisor a b have xab':"x mod a < a" "x mod b < b" by auto
+        from xab xab' have "?f x \<in> (?S a \<times> ?S b)"
+          by (simp add: coprime_mod[OF a(1)] coprime_mod[OF b(1)])}
       moreover
       {fix x y assume x: "x \<in> ?S a" and y: "y \<in> ?S b"
-	hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all
-	from chinese_remainder_coprime_unique[OF ab a(1) b(1) x'(1) y'(1)]
-	obtain z where z: "coprime z (a * b)" "z < a * b" "[z = x] (mod a)"
-	  "[z = y] (mod b)" by blast
-	hence "(x,y) \<in> ?f ` (?S (a*b))"
-	  using y'(2) mod_less_divisor[of b y] x'(2) mod_less_divisor[of a x]
-	  by (auto simp add: image_iff modeq_def)}
+        hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all
+        from chinese_remainder_coprime_unique[OF ab a(1) b(1) x'(1) y'(1)]
+        obtain z where z: "coprime z (a * b)" "z < a * b" "[z = x] (mod a)"
+          "[z = y] (mod b)" by blast
+        hence "(x,y) \<in> ?f ` (?S (a*b))"
+          using y'(2) mod_less_divisor[of b y] x'(2) mod_less_divisor[of a x]
+          by (auto simp add: image_iff modeq_def)}
       ultimately show ?thesis by auto
     qed
     have finj: "inj_on ?f (?S (a*b))"
       unfolding inj_on_def
     proof(clarify)
       fix x y assume H: "coprime x (a * b)" "x < a * b" "coprime y (a * b)"
-	"y < a * b" "x mod a = y mod a" "x mod b = y mod b"
+        "y < a * b" "x mod a = y mod a" "x mod b = y mod b"
       hence cp: "coprime x a" "coprime x b" "coprime y a" "coprime y b"
-	by (simp_all add: coprime_mul_eq)
+        by (simp_all add: coprime_mul_eq)
       from chinese_remainder_coprime_unique[OF ab a(1) b(1) cp(3,4)] H
       show "x = y" unfolding modeq_def by blast
     qed
@@ -593,7 +593,7 @@
     {fix m assume m: "m \<in> ?S"
       hence "coprime m n" by simp
       with coprime_mul[of n a m] an have "coprime (a*m) n"
-	by (simp add: coprime_commute)}
+        by (simp add: coprime_commute)}
     hence Sn: "\<forall>m\<in> ?S. coprime (a*m) n " by blast
     from coprime_nproduct[OF fS, of n "\<lambda>m. m"] have nP:"coprime ?P n"
       by (simp add: coprime_commute)
@@ -601,7 +601,7 @@
     proof-
       let ?h = "\<lambda>m. m mod n"
       {fix m assume mS: "m\<in> ?S"
-	hence "?h m \<in> ?S" by simp}
+        hence "?h m \<in> ?S" by simp}
       hence hS: "?h ` ?S = ?S"by (auto simp add: image_iff)
       have "a\<noteq>0" using an n1 nz apply- apply (rule ccontr) by simp
       hence inj: "inj_on (op * a) ?S" unfolding inj_on_def by simp
@@ -609,41 +609,41 @@
       have eq0: "fold_image op * (?h \<circ> op * a) 1 {m. coprime m n \<and> m < n} =
      fold_image op * (\<lambda>m. m) 1 {m. coprime m n \<and> m < n}"
       proof (rule fold_image_eq_general[where h="?h o (op * a)"])
-	show "finite ?S" using fS .
+        show "finite ?S" using fS .
       next
-	{fix y assume yS: "y \<in> ?S" hence y: "coprime y n" "y < n" by simp_all
-	  from cong_solve_unique[OF an nz, of y]
-	  obtain x where x:"x < n" "[a * x = y] (mod n)" "\<forall>z. z < n \<and> [a * z = y] (mod n) \<longrightarrow> z=x" by blast
-	  from cong_coprime[OF x(2)] y(1)
-	  have xm: "coprime x n" by (simp add: coprime_mul_eq coprime_commute)
-	  {fix z assume "z \<in> ?S" "(?h \<circ> op * a) z = y"
-	    hence z: "coprime z n" "z < n" "(?h \<circ> op * a) z = y" by simp_all
-	    from x(3)[rule_format, of z] z(2,3) have "z=x"
-	      unfolding modeq_def mod_less[OF y(2)] by simp}
-	  with xm x(1,2) have "\<exists>!x. x \<in> ?S \<and> (?h \<circ> op * a) x = y"
-	    unfolding modeq_def mod_less[OF y(2)] by auto }
-	thus "\<forall>y\<in>{m. coprime m n \<and> m < n}.
+        {fix y assume yS: "y \<in> ?S" hence y: "coprime y n" "y < n" by simp_all
+          from cong_solve_unique[OF an nz, of y]
+          obtain x where x:"x < n" "[a * x = y] (mod n)" "\<forall>z. z < n \<and> [a * z = y] (mod n) \<longrightarrow> z=x" by blast
+          from cong_coprime[OF x(2)] y(1)
+          have xm: "coprime x n" by (simp add: coprime_mul_eq coprime_commute)
+          {fix z assume "z \<in> ?S" "(?h \<circ> op * a) z = y"
+            hence z: "coprime z n" "z < n" "(?h \<circ> op * a) z = y" by simp_all
+            from x(3)[rule_format, of z] z(2,3) have "z=x"
+              unfolding modeq_def mod_less[OF y(2)] by simp}
+          with xm x(1,2) have "\<exists>!x. x \<in> ?S \<and> (?h \<circ> op * a) x = y"
+            unfolding modeq_def mod_less[OF y(2)] by auto }
+        thus "\<forall>y\<in>{m. coprime m n \<and> m < n}.
        \<exists>!x. x \<in> {m. coprime m n \<and> m < n} \<and> ((\<lambda>m. m mod n) \<circ> op * a) x = y" by blast
       next
-	{fix x assume xS: "x\<in> ?S"
-	  hence x: "coprime x n" "x < n" by simp_all
-	  with an have "coprime (a*x) n"
-	    by (simp add: coprime_mul_eq[of n a x] coprime_commute)
-	  hence "?h (a*x) \<in> ?S" using nz
-	    by (simp add: coprime_mod[OF nz] mod_less_divisor)}
-	thus " \<forall>x\<in>{m. coprime m n \<and> m < n}.
+        {fix x assume xS: "x\<in> ?S"
+          hence x: "coprime x n" "x < n" by simp_all
+          with an have "coprime (a*x) n"
+            by (simp add: coprime_mul_eq[of n a x] coprime_commute)
+          hence "?h (a*x) \<in> ?S" using nz
+            by (simp add: coprime_mod[OF nz] mod_less_divisor)}
+        thus " \<forall>x\<in>{m. coprime m n \<and> m < n}.
        ((\<lambda>m. m mod n) \<circ> op * a) x \<in> {m. coprime m n \<and> m < n} \<and>
        ((\<lambda>m. m mod n) \<circ> op * a) x = ((\<lambda>m. m mod n) \<circ> op * a) x" by simp
       qed
       from nproduct_mod[OF fS nz, of "op * a"]
       have "[(setprod (op *a) ?S) = (setprod (?h o (op * a)) ?S)] (mod n)"
-	unfolding o_def
-	by (simp add: cong_commute)
+        unfolding o_def
+        by (simp add: cong_commute)
       also have "[setprod (?h o (op * a)) ?S = ?P ] (mod n)"
-	using eq0 fS an by (simp add: setprod_def modeq_def o_def)
+        using eq0 fS an by (simp add: setprod_def modeq_def o_def)
       finally show "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
-	unfolding cardfS mult_commute[of ?P "a^ (card ?S)"]
-	  nproduct_cmul[OF fS, symmetric] mult_1_right by simp
+        unfolding cardfS mult_commute[of ?P "a^ (card ?S)"]
+          nproduct_cmul[OF fS, symmetric] mult_1_right by simp
     qed
     from cong_mult_lcancel[OF nP Paphi] have ?thesis . }
   ultimately show ?thesis by blast
@@ -750,23 +750,23 @@
       let ?y = "a^ ((n - 1) div m * m)"
       note mdeq = mod_div_equality[of "(n - 1)" m]
       from coprime_exp[OF an(1)[unfolded coprime_commute[of a n]],
-	of "(n - 1) div m * m"]
+        of "(n - 1) div m * m"]
       have yn: "coprime ?y n" by (simp add: coprime_commute)
       have "?y mod n = (a^m)^((n - 1) div m) mod n"
-	by (simp add: algebra_simps power_mult)
+        by (simp add: algebra_simps power_mult)
       also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n"
-	using power_mod[of "a^m" n "(n - 1) div m"] by simp
+        using power_mod[of "a^m" n "(n - 1) div m"] by simp
       also have "\<dots> = 1" using m(3)[unfolded modeq_def onen] onen
-	by (simp add: power_Suc0)
+        by (simp add: power_Suc0)
       finally have th3: "?y mod n = 1"  .
       have th2: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)"
-	using an1[unfolded modeq_def onen] onen
-	  mod_div_equality[of "(n - 1)" m, symmetric]
-	by (simp add:power_add[symmetric] modeq_def th3 del: One_nat_def)
+        using an1[unfolded modeq_def onen] onen
+          mod_div_equality[of "(n - 1)" m, symmetric]
+        by (simp add:power_add[symmetric] modeq_def th3 del: One_nat_def)
       from cong_mult_lcancel[of ?y n "a^((n - 1) mod m)" 1, OF yn th2]
       have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)"  .
       from m(4)[rule_format, OF th0] nm1
-	less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1
+        less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1
       have False by blast }
     hence "(n - 1) mod m = 0" by auto
     then have mn: "m dvd n - 1" by presburger
@@ -924,18 +924,18 @@
     {assume pn: "prime n"
 
       from pn[unfolded prime_def] have "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
-	using n
-	apply (cases "n = 0 \<or> n=1",simp)
-	by (clarsimp, erule_tac x="p" in allE, auto)}
+        using n
+        apply (cases "n = 0 \<or> n=1",simp)
+        by (clarsimp, erule_tac x="p" in allE, auto)}
     moreover
     {assume H: "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
       from n have n1: "n > 1" by arith
       {fix m assume m: "m dvd n" "m\<noteq>1"
-	from prime_factor[OF m(2)] obtain p where
-	  p: "prime p" "p dvd m" by blast
-	from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast
-	with p(2) have "n dvd m"  by simp
-	hence "m=n"  using dvd_anti_sym[OF m(1)] by simp }
+        from prime_factor[OF m(2)] obtain p where
+          p: "prime p" "p dvd m" by blast
+        from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast
+        with p(2) have "n dvd m"  by simp
+        hence "m=n"  using dvd_anti_sym[OF m(1)] by simp }
       with n1 have "prime n"  unfolding prime_def by auto }
     ultimately have ?thesis using n by blast}
   ultimately       show ?thesis by auto
@@ -952,9 +952,9 @@
     {fix d assume d: "d dvd n" "d^2 \<le> n" and H: "\<forall>m. m dvd n \<longrightarrow> m=1 \<or> m=n"
       from H d have d1n: "d = 1 \<or> d=n" by blast
       {assume dn: "d=n"
-	have "n^2 > n*1" using n
-	  by (simp add: power2_eq_square mult_less_cancel1)
-	with dn d(2) have "d=1" by simp}
+        have "n^2 > n*1" using n
+          by (simp add: power2_eq_square mult_less_cancel1)
+        with dn d(2) have "d=1" by simp}
       with d1n have "d = 1" by blast  }
     moreover
     {fix d assume d: "d dvd n" and H: "\<forall>d'. d' dvd n \<and> d'^2 \<le> n \<longrightarrow> d' = 1"
@@ -963,15 +963,15 @@
       from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast
       from n dp e have ep:"e > 0" by simp
       have "d^2 \<le> n \<or> e^2 \<le> n" using dp ep
-	by (auto simp add: e power2_eq_square mult_le_cancel_left)
+        by (auto simp add: e power2_eq_square mult_le_cancel_left)
       moreover
       {assume h: "d^2 \<le> n"
-	from H[rule_format, of d] h d have "d = 1" by blast}
+        from H[rule_format, of d] h d have "d = 1" by blast}
       moreover
       {assume h: "e^2 \<le> n"
-	from e have "e dvd n" unfolding dvd_def by (simp add: mult_commute)
-	with H[rule_format, of e] h have "e=1" by simp
-	with e have "d = n" by simp}
+        from e have "e dvd n" unfolding dvd_def by (simp add: mult_commute)
+        with H[rule_format, of e] h have "e=1" by simp
+        with e have "d = n" by simp}
       ultimately have "d=1 \<or> d=n"  by blast}
     ultimately have ?thesis unfolding prime_def using np n(2) by blast}
   ultimately show ?thesis by auto
@@ -990,14 +990,14 @@
     moreover
     {assume H: ?rhs
       {fix d assume d: "d dvd n" "d^2 \<le> n" "d\<noteq>1"
-	from prime_factor[OF d(3)]
-	obtain p where p: "prime p" "p dvd d" by blast
-	from n have np: "n > 0" by arith
-	from d(1) n have "d \<noteq> 0" by - (rule ccontr, auto)
-	hence dp: "d > 0" by arith
-	from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2)
-	have "p^2 \<le> n" unfolding power2_eq_square by arith
-	with H n p(1) dvd_trans[OF p(2) d(1)] have False  by blast}
+        from prime_factor[OF d(3)]
+        obtain p where p: "prime p" "p dvd d" by blast
+        from n have np: "n > 0" by arith
+        from d(1) n have "d \<noteq> 0" by - (rule ccontr, auto)
+        hence dp: "d > 0" by arith
+        from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2)
+        have "p^2 \<le> n" unfolding power2_eq_square by arith
+        with H n p(1) dvd_trans[OF p(2) d(1)] have False  by blast}
       with n prime_divisor_sqrt  have ?lhs by auto}
     ultimately have ?thesis by blast }
   ultimately show ?thesis by (cases "n=0 \<or> n=1", auto)
@@ -1108,11 +1108,11 @@
     from one_le_power[OF a1] have ath: "1 \<le> a ^ ((n - 1) div p)" .
     {assume b0: "b = 0"
       from p(2) nqr have "(n - 1) mod p = 0"
-	apply (simp only: dvd_eq_mod_eq_0[symmetric]) by (rule dvd_mult2, simp)
+        apply (simp only: dvd_eq_mod_eq_0[symmetric]) by (rule dvd_mult2, simp)
       with mod_div_equality[of "n - 1" p]
       have "(n - 1) div p * p= n - 1" by auto
       hence eq: "(a^((n - 1) div p))^p = a^(n - 1)"
-	by (simp only: power_mult[symmetric])
+        by (simp only: power_mult[symmetric])
       from prime_ge_2[OF p(1)] have pS: "Suc (p - 1) = p" by arith
       from b(1) have d: "n dvd a^((n - 1) div p)" unfolding b0 cong_0_divides .
       from divides_rexp[OF d, of "p - 1"] pS eq cong_divides[OF an] n
@@ -1238,13 +1238,13 @@
     have th10: "a ^ ((n - 1) div p) mod n \<le> a ^ ((n - 1) div p)" .
     {assume "a ^ ((n - 1) div p) mod n = 0"
       then obtain s where s: "a ^ ((n - 1) div p) = n*s"
-	unfolding mod_eq_0_iff by blast
+        unfolding mod_eq_0_iff by blast
       hence eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp
       from qrn[symmetric] have qn1: "q dvd n - 1" unfolding dvd_def by auto
       from dvd_trans[OF p(2) qn1] div_mod_equality'[of "n - 1" p]
       have npp: "(n - 1) div p * p = n - 1" by (simp add: dvd_eq_mod_eq_0)
       with eq0 have "a^ (n - 1) = (n*s)^p"
-	by (simp add: power_mult[symmetric])
+        by (simp add: power_mult[symmetric])
       hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp
       also have "\<dots> = 0" by (simp add: mult_assoc)
       finally have False by simp }
--- a/src/HOL/Old_Number_Theory/Primes.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Old_Number_Theory/Primes.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -137,7 +137,7 @@
 
 fun fact :: "nat \<Rightarrow> nat" where
   "fact 0 = 1"
-| "fact (Suc n) = Suc n * fact n"	
+| "fact (Suc n) = Suc n * fact n"
 
 lemma fact_lt: "0 < fact n" by(induct n, simp_all)
 lemma fact_le: "fact n \<ge> 1" using fact_lt[of n] by simp 
@@ -297,7 +297,7 @@
       using z by auto 
     then have "a^n * x - b^n * y = ?g^n \<or> b^n * x - a^n * y = ?g^n"
       using z ab'' by (simp only: power_mult_distrib[symmetric] 
-	diff_mult_distrib2 mult_assoc[symmetric])
+        diff_mult_distrib2 mult_assoc[symmetric])
     hence  ?thesis by blast }
   ultimately show ?thesis by blast
 qed
@@ -615,7 +615,7 @@
       with prime_coprime[OF p, of b] b 
       have cpb: "coprime b p" using coprime_commute by blast 
       from coprime_exp[OF cpb] have pnb: "coprime (p^n) b" 
-	by (simp add: coprime_commute)
+        by (simp add: coprime_commute)
       from coprime_divprod[OF pnba pnb] have ?thesis by blast }
     moreover
     {assume pb: "p dvd b"
@@ -624,7 +624,7 @@
       with prime_coprime[OF p, of a] a
       have cpb: "coprime a p" using coprime_commute by blast 
       from coprime_exp[OF cpb] have pnb: "coprime (p^n) a" 
-	by (simp add: coprime_commute)
+        by (simp add: coprime_commute)
       from coprime_divprod[OF pab pnb] have ?thesis by blast }
     ultimately have ?thesis by blast}
   ultimately show ?thesis by blast
@@ -661,14 +661,14 @@
   {assume n: "n \<noteq> 0" then obtain m where m: "n = Suc m" by (cases n, auto)
     {assume c: "c = 0"
       with H(3) m H(2) have ?ths apply simp 
-	apply (cases "a=0", simp_all) 
-	apply (rule exI[where x="0"], simp)
-	apply (rule exI[where x="0"], simp)
-	done}
+        apply (cases "a=0", simp_all) 
+        apply (rule exI[where x="0"], simp)
+        apply (rule exI[where x="0"], simp)
+        done}
     moreover
     {assume "c=1" with H(3) power_one have "a*b = 1" by simp 
-	hence "a = 1 \<and> b = 1" by simp
-	hence ?ths 
+        hence "a = 1 \<and> b = 1" by simp
+        hence ?ths 
       apply -
       apply (rule exI[where x=1])
       apply (rule exI[where x=1])
@@ -690,7 +690,7 @@
       from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" b]]]
       have kb: "coprime k b" by (simp add: coprime_commute) 
       from H(3) l k pn0 have kbln: "k * b = l ^ n" 
-	by (auto simp add: power_mult_distrib)
+        by (auto simp add: power_mult_distrib)
       from H(1)[rule_format, OF lc kb kbln]
       obtain r s where rs: "k = r ^n" "b = s^n" by blast
       from k rs(1) have "a = (p*r)^n" by (simp add: power_mult_distrib)
@@ -705,7 +705,7 @@
       from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" a]]]
       have kb: "coprime k a" by (simp add: coprime_commute) 
       from H(3) l k pn0 n have kbln: "k * a = l ^ n" 
-	by (simp add: power_mult_distrib mult_commute)
+        by (simp add: power_mult_distrib mult_commute)
       from H(1)[rule_format, OF lc kb kbln]
       obtain r s where rs: "k = r ^n" "a = s^n" by blast
       from k rs(1) have "b = (p*r)^n" by (simp add: power_mult_distrib)
@@ -743,7 +743,7 @@
     {assume pp: "prime (p^Suc n)"
       hence "p = 1 \<or> p^n = 1" using prime_product[of p "p^n"] by simp
       with p have n: "n = 0" 
-	by (simp only: exp_eq_1 ) simp
+        by (simp only: exp_eq_1 ) simp
       with pp have "prime p \<and> Suc n = 1" by simp}
     moreover
     {assume n: "prime p \<and> Suc n = 1" hence "prime (p^Suc n)" by simp}
--- a/src/HOL/Old_Number_Theory/WilsonBij.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Old_Number_Theory/WilsonBij.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -155,7 +155,7 @@
                apply (rule_tac [7] inv_g_1 [THEN aux2])
                  apply (unfold zprime_def)
                  apply (auto intro: d22set_g_1 d22set_le
-		   aux1 aux2 aux3 aux4)
+                   aux1 aux2 aux3 aux4)
   done
 
 lemma inv_d22set_d22set:
--- a/src/HOL/Old_Number_Theory/WilsonRuss.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Old_Number_Theory/WilsonRuss.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -148,7 +148,7 @@
       apply (subst zcong_zmod [symmetric])
       apply (subst inv_inv_aux)
        apply (subgoal_tac [2]
-	 "zcong (a * a^(nat (p - 1) * nat (p - 3))) (a * 1) p")
+         "zcong (a * a^(nat (p - 1) * nat (p - 3))) (a * 1) p")
         apply (rule_tac [3] zcong_zmult)
          apply (rule_tac [4] zcong_zpower_zmult)
          apply (erule_tac [4] Little_Fermat)
@@ -263,7 +263,7 @@
   apply (subst setprod_insert)
     apply (tactic {* stac (thm "setprod_insert") 3 *})
       apply (subgoal_tac [5]
-	"zcong (a * inv p a * (\<Prod>x\<in> wset(a - 1, p). x)) (1 * 1) p")
+        "zcong (a * inv p a * (\<Prod>x\<in> wset(a - 1, p). x)) (1 * 1) p")
        prefer 5
        apply (simp add: zmult_assoc)
       apply (rule_tac [5] zcong_zmult)
--- a/src/HOL/Orderings.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Orderings.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -339,18 +339,18 @@
           (* exclude numeric types: linear arithmetic subsumes transitivity *)
           let val T = type_of t
           in
-	    T = HOLogic.natT orelse T = HOLogic.intT orelse T = HOLogic.realT
+            T = HOLogic.natT orelse T = HOLogic.intT orelse T = HOLogic.realT
           end;
-	fun rel (bin_op $ t1 $ t2) =
+        fun rel (bin_op $ t1 $ t2) =
               if excluded t1 then NONE
               else if Pattern.matches thy (eq, bin_op) then SOME (t1, "=", t2)
               else if Pattern.matches thy (le, bin_op) then SOME (t1, "<=", t2)
               else if Pattern.matches thy (less, bin_op) then SOME (t1, "<", t2)
               else NONE
-	  | rel _ = NONE;
-	fun dec (Const (@{const_name Not}, _) $ t) = (case rel t
-	      of NONE => NONE
-	       | SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2))
+          | rel _ = NONE;
+        fun dec (Const (@{const_name Not}, _) $ t) = (case rel t
+              of NONE => NONE
+               | SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2))
           | dec x = rel x;
       in dec t end
       | decomp thy _ = NONE;
--- a/src/HOL/PReal.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/PReal.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,9 @@
-(*  Title       : PReal.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : The positive reals as Dedekind sections of positive
-         rationals. Fundamentals of Abstract Analysis [Gleason- p. 121]
-                  provides some of the definitions.
+(*  Title:      HOL/PReal.thy
+    Author:     Jacques D. Fleuriot, University of Cambridge
+
+The positive reals as Dedekind sections of positive
+rationals. Fundamentals of Abstract Analysis [Gleason- p. 121]
+provides some of the definitions.
 *)
 
 header {* Positive real numbers *}
@@ -319,10 +319,10 @@
     show "x * ?f \<in> A"
     proof (rule preal_downwards_closed [OF A x])
       show "0 < x * ?f"
-	by (simp add: divide_inverse zero_less_mult_iff)
+        by (simp add: divide_inverse zero_less_mult_iff)
     next
       show "x * ?f < x"
-	by (insert mult_strict_left_mono [OF fless xpos], simp)
+        by (insert mult_strict_left_mono [OF fless xpos], simp)
     qed
   qed
 qed
@@ -399,16 +399,16 @@
     show "x * y \<notin> mult_set A B"
     proof -
       { fix u::rat and v::rat
-	      assume "u \<in> A" and "v \<in> B" and "x*y = u*v"
-	      moreover
-	      with prems have "u<x" and "v<y" by (blast dest: not_in_preal_ub)+
-	      moreover
-	      with prems have "0\<le>v"
-	        by (blast intro: preal_imp_pos [OF B]  order_less_imp_le prems)
-	      moreover
+              assume "u \<in> A" and "v \<in> B" and "x*y = u*v"
+              moreover
+              with prems have "u<x" and "v<y" by (blast dest: not_in_preal_ub)+
+              moreover
+              with prems have "0\<le>v"
+                by (blast intro: preal_imp_pos [OF B]  order_less_imp_le prems)
+              moreover
         from calculation
-	      have "u*v < x*y" by (blast intro: mult_strict_mono prems)
-	      ultimately have False by force }
+              have "u*v < x*y" by (blast intro: mult_strict_mono prems)
+              ultimately have False by force }
       thus ?thesis by (auto simp add: mult_set_def)
     qed
   qed
@@ -446,15 +446,15 @@
     show "\<exists>y'\<in>B. z = (z/y) * y'"
     proof
       show "z = (z/y)*y"
-	by (simp add: divide_inverse mult_commute [of y] mult_assoc
-		      order_less_imp_not_eq2)
+        by (simp add: divide_inverse mult_commute [of y] mult_assoc
+                      order_less_imp_not_eq2)
       show "y \<in> B" by fact
     qed
   next
     show "z/y \<in> A"
     proof (rule preal_downwards_closed [OF A x])
       show "0 < z/y"
-	by (simp add: zero_less_divide_iff)
+        by (simp add: zero_less_divide_iff)
       show "z/y < x" by (simp add: pos_divide_less_eq zless)
     qed
   qed
@@ -531,12 +531,12 @@
       proof (intro exI conjI)
         show "0 < x/v"
           by (simp add: zero_less_divide_iff xpos vpos)
-	show "x / v < 1"
+        show "x / v < 1"
           by (simp add: pos_divide_less_eq vpos xlessv)
         show "\<exists>v'\<in>A. x = (x / v) * v'"
         proof
           show "x = (x/v)*v"
-	    by (simp add: divide_inverse mult_assoc vpos
+            by (simp add: divide_inverse mult_assoc vpos
                           order_less_imp_not_eq2)
           show "v \<in> A" by fact
         qed
@@ -661,12 +661,12 @@
     show "inverse x \<notin> inverse_set A"
     proof -
       { fix y::rat 
-	assume ygt: "inverse x < y"
-	have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt])
-	have iyless: "inverse y < x" 
-	  by (simp add: inverse_less_imp_less [of x] ygt)
-	have "inverse y \<in> A"
-	  by (simp add: preal_downwards_closed [OF A x] iyless)}
+        assume ygt: "inverse x < y"
+        have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt])
+        have iyless: "inverse y < x" 
+          by (simp add: inverse_less_imp_less [of x] ygt)
+        have "inverse y \<in> A"
+          by (simp add: preal_downwards_closed [OF A x] iyless)}
      thus ?thesis by (auto simp add: inverse_set_def)
     qed
   qed
@@ -820,9 +820,9 @@
     proof -
       have "r + ?d < r + (r * ?d)/y" by (simp add: dless)
       also with ypos have "... = (r/y) * (y + ?d)"
-	by (simp only: algebra_simps divide_inverse, simp)
+        by (simp only: algebra_simps divide_inverse, simp)
       also have "... = r*x" using ypos
-	by (simp add: times_divide_eq_left) 
+        by (simp add: times_divide_eq_left) 
       finally show "r + ?d < r*x" .
     qed
     with r notin rdpos
--- a/src/HOL/RComplete.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/RComplete.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -196,11 +196,11 @@
 
       have "\<forall>pe \<in> ?pS. pe \<le> px"
       proof
-	fix pe
-	assume "pe \<in> ?pS"
-	hence "real_of_preal pe \<in> S" by simp
-	hence "real_of_preal pe \<le> x" using x_ub_S by simp
-	thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff)
+        fix pe
+        assume "pe \<in> ?pS"
+        hence "real_of_preal pe \<in> S" by simp
+        hence "real_of_preal pe \<le> x" using x_ub_S by simp
+        thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff)
       qed
 
       moreover have "?pS \<noteq> {}" using ps_in_pS by auto
@@ -259,13 +259,13 @@
         fix x
         assume "isUb (UNIV::real set) S x"
         hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)"
-	  using S_Ub_is_SHIFT_Ub by simp
+          using S_Ub_is_SHIFT_Ub by simp
         hence "t \<le> (x + (-X) + 1)"
-	  using t_is_Lub by (simp add: isLub_le_isUb)
+          using t_is_Lub by (simp add: isLub_le_isUb)
         hence "t + X + -1 \<le> x" by arith
       }
       then show "(t + X + -1) <=* Collect (isUb UNIV S)"
-	by (simp add: setgeI)
+        by (simp add: setgeI)
     next
       show "isUb UNIV S (t + X + -1)"
       proof -
--- a/src/HOL/Ring_and_Field.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Ring_and_Field.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,10 @@
-(*  Title:   HOL/Ring_and_Field.thy
-    Author:  Gertrud Bauer, Steven Obua, Tobias Nipkow, Lawrence C Paulson, and Markus Wenzel,
-             with contributions by Jeremy Avigad
+(*  Title:      HOL/Ring_and_Field.thy
+    Author:     Gertrud Bauer
+    Author:     Steven Obua
+    Author:     Tobias Nipkow
+    Author:     Lawrence C Paulson
+    Author:     Markus Wenzel
+    Author:     Jeremy Avigad
 *)
 
 header {* (Ordered) Rings and Fields *}
@@ -2206,11 +2210,11 @@
       apply (simp_all add: mulprts abs_prts)
       apply (insert prems)
       apply (auto simp add: 
-	algebra_simps 
-	iffD1[OF zero_le_iff_zero_nprt] iffD1[OF le_zero_iff_zero_pprt]
-	iffD1[OF le_zero_iff_pprt_id] iffD1[OF zero_le_iff_nprt_id])
-	apply(drule (1) mult_nonneg_nonpos[of a b], simp)
-	apply(drule (1) mult_nonneg_nonpos2[of b a], simp)
+        algebra_simps 
+        iffD1[OF zero_le_iff_zero_nprt] iffD1[OF le_zero_iff_zero_pprt]
+        iffD1[OF le_zero_iff_pprt_id] iffD1[OF zero_le_iff_nprt_id])
+        apply(drule (1) mult_nonneg_nonpos[of a b], simp)
+        apply(drule (1) mult_nonneg_nonpos2[of b a], simp)
       done
   next
     assume "~(0 <= a*b)"
--- a/src/HOL/SEQ.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SEQ.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,10 @@
-(*  Title       : SEQ.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Convergence of sequences and series
-    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
-    Additional contributions by Jeremy Avigad and Brian Huffman
+(*  Title:      HOL/SEQ.thy
+    Author:     Jacques D. Fleuriot, University of Cambridge
+    Author:     Lawrence C Paulson
+    Author:     Jeremy Avigad
+    Author:     Brian Huffman
+
+Convergence of sequences and series.
 *)
 
 header {* Sequences and Convergence *}
@@ -337,10 +338,10 @@
     have "\<bar>f (N+k) - l\<bar> < e"
     proof (induct k)
       case 0 show ?case using N
-	by simp   
+        by simp   
     next
       case (Suc k) thus ?case using N inc [of "N+k"]
-	by simp
+        by simp
     qed 
   } note 1 = this
   { fix n
@@ -549,11 +550,11 @@
     proof (simp add: LIMSEQ_iff)
       assume "\<forall>r>0. \<exists>no. \<forall>n\<ge>no. \<bar>f n - lim f\<bar> < r"
       hence "\<exists>no. \<forall>n\<ge>no. \<bar>f n - lim f\<bar> < lim f - x"
-	by (metis 0)
+        by (metis 0)
       from this obtain no where "\<forall>n\<ge>no. \<bar>f n - lim f\<bar> < lim f - x"
-	by blast
+        by blast
       thus "lim f \<le> x"
-	by (metis add_cancel_end add_minus_cancel diff_def linorder_linear 
+        by (metis add_cancel_end add_minus_cancel diff_def linorder_linear 
                   linorder_not_le minus_diff_eq abs_diff_less_iff fn_le) 
     qed
 qed
@@ -685,45 +686,45 @@
     hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
     {fix n
       have "?P (f (Suc n)) (f n)" 
-	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
-	using H apply - 
+        unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
+        using H apply - 
       apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
       unfolding order_le_less by blast 
     hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
   note fSuc = this
     {fix p q assume pq: "p \<ge> f q"
       have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
-	by (cases q, simp_all) }
+        by (cases q, simp_all) }
     note pqth = this
     {fix q
       have "f (Suc q) > f q" apply (induct q) 
-	using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
+        using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
     note fss = this
     from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
     {fix a b 
       have "f a \<le> f (a + b)"
       proof(induct b)
-	case 0 thus ?case by simp
+        case 0 thus ?case by simp
       next
-	case (Suc b)
-	from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
+        case (Suc b)
+        from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
       qed}
     note fmon0 = this
     have "monoseq (\<lambda>n. s (f n))" 
     proof-
       {fix n
-	have "s (f n) \<ge> s (f (Suc n))" 
-	proof(cases n)
-	  case 0
-	  assume n0: "n = 0"
-	  from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
-	  from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
-	next
-	  case (Suc m)
-	  assume m: "n = Suc m"
-	  from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
-	  from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
-	qed}
+        have "s (f n) \<ge> s (f (Suc n))" 
+        proof(cases n)
+          case 0
+          assume n0: "n = 0"
+          from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
+          from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
+        next
+          case (Suc m)
+          assume m: "n = Suc m"
+          from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
+          from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
+        qed}
       thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
     qed
     with th1 have ?thesis by blast}
@@ -750,26 +751,26 @@
     hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
     {fix n
       have "f n > N \<and> ?P (f (Suc n)) (f n)"
-	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
+        unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
       proof (induct n)
-	case 0 thus ?case
-	  using f0 N apply auto 
-	  apply (erule allE[where x="f 0"], clarsimp) 
-	  apply (rule_tac x="m" in exI, simp)
-	  by (subgoal_tac "f 0 \<noteq> m", auto)
+        case 0 thus ?case
+          using f0 N apply auto 
+          apply (erule allE[where x="f 0"], clarsimp) 
+          apply (rule_tac x="m" in exI, simp)
+          by (subgoal_tac "f 0 \<noteq> m", auto)
       next
-	case (Suc n)
-	from Suc.hyps have Nfn: "N < f n" by blast
-	from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
-	with Nfn have mN: "m > N" by arith
-	note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
-	
-	from key have th0: "f (Suc n) > N" by simp
-	from N[rule_format, OF th0]
-	obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
-	have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
-	hence "m' > f (Suc n)" using m'(1) by simp
-	with key m'(2) show ?case by auto
+        case (Suc n)
+        from Suc.hyps have Nfn: "N < f n" by blast
+        from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
+        with Nfn have mN: "m > N" by arith
+        note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
+        
+        from key have th0: "f (Suc n) > N" by simp
+        from N[rule_format, OF th0]
+        obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
+        have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
+        hence "m' > f (Suc n)" using m'(1) by simp
+        with key m'(2) show ?case by auto
       qed}
     note fSuc = this
     {fix n
--- a/src/HOL/SET-Protocol/Cardholder_Registration.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SET-Protocol/Cardholder_Registration.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,8 @@
-(*  Title:      HOL/Auth/SET/Cardholder_Registration
-    Authors:    Giampaolo Bella, Fabio Massacci, Lawrence C Paulson,
-                Piero Tramontano
+(*  Title:      HOL/SET-Protocol/Cardholder_Registration.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
+    Author:     Piero Tramontano
 *)
 
 header{*The SET Cardholder Registration Protocol*}
@@ -27,19 +29,19 @@
 
 KeyCryptKey_Cons:
       --{*Says is the only important case.
-	1st case: CR5, where KC3 encrypts KC2.
-	2nd case: any use of priEK C.
-	Revision 1.12 has a more complicated version with separate treatment of
-	  the dependency of KC1, KC2 and KC3 on priEK (CA i.)  Not needed since
-	  priEK C is never sent (and so can't be lost except at the start). *}
+        1st case: CR5, where KC3 encrypts KC2.
+        2nd case: any use of priEK C.
+        Revision 1.12 has a more complicated version with separate treatment of
+          the dependency of KC1, KC2 and KC3 on priEK (CA i.)  Not needed since
+          priEK C is never sent (and so can't be lost except at the start). *}
   "KeyCryptKey DK K (ev # evs) =
    (KeyCryptKey DK K evs |
     (case ev of
       Says A B Z =>
        ((\<exists>N X Y. A \<noteq> Spy &
-	         DK \<in> symKeys &
-		 Z = {|Crypt DK {|Agent A, Nonce N, Key K, X|}, Y|}) |
-	(\<exists>C. DK = priEK C))
+                 DK \<in> symKeys &
+                 Z = {|Crypt DK {|Agent A, Nonce N, Key K, X|}, Y|}) |
+        (\<exists>C. DK = priEK C))
     | Gets A' X => False
     | Notes A' X => False))"
 
@@ -63,7 +65,7 @@
     5th case: any use of @{term "priEK C"} (including CardSecret).
     NB the only Nonces we need to keep secret are CardSecret and NonceCCA.
     But we can't prove @{text Nonce_compromise} unless the relation covers ALL
-	nonces that the protocol keeps secret.
+        nonces that the protocol keeps secret.
   *}
   "KeyCryptNonce DK N (ev # evs) =
    (KeyCryptNonce DK N evs |
@@ -71,20 +73,20 @@
       Says A B Z =>
        A \<noteq> Spy &
        ((\<exists>X Y. DK \<in> symKeys &
-	       Z = (EXHcrypt DK X {|Agent A, Nonce N|} Y)) |
-	(\<exists>X Y. DK \<in> symKeys &
-	       Z = {|Crypt DK {|Agent A, Nonce N, X|}, Y|}) |
-	(\<exists>K i X Y.
-	  K \<in> symKeys &
+               Z = (EXHcrypt DK X {|Agent A, Nonce N|} Y)) |
+        (\<exists>X Y. DK \<in> symKeys &
+               Z = {|Crypt DK {|Agent A, Nonce N, X|}, Y|}) |
+        (\<exists>K i X Y.
+          K \<in> symKeys &
           Z = Crypt K {|sign (priSK (CA i)) {|Agent B, Nonce N, X|}, Y|} &
-	  (DK=K | KeyCryptKey DK K evs)) |
-	(\<exists>K C NC3 Y.
-	  K \<in> symKeys &
+          (DK=K | KeyCryptKey DK K evs)) |
+        (\<exists>K C NC3 Y.
+          K \<in> symKeys &
           Z = Crypt K
- 	        {|sign (priSK C) {|Agent B, Nonce NC3, Agent C, Nonce N|},
+                {|sign (priSK C) {|Agent B, Nonce NC3, Agent C, Nonce N|},
                   Y|} &
-	  (DK=K | KeyCryptKey DK K evs)) |
-	(\<exists>C. DK = priEK C))
+          (DK=K | KeyCryptKey DK K evs)) |
+        (\<exists>C. DK = priEK C))
     | Gets A' X => False
     | Notes A' X => False))"
 
@@ -96,28 +98,28 @@
 where
 
   Nil:    --{*Initial trace is empty*}
-	  "[] \<in> set_cr"
+          "[] \<in> set_cr"
 
 | Fake:    --{*The spy MAY say anything he CAN say.*}
-	   "[| evsf \<in> set_cr; X \<in> synth (analz (knows Spy evsf)) |]
-	    ==> Says Spy B X  # evsf \<in> set_cr"
+           "[| evsf \<in> set_cr; X \<in> synth (analz (knows Spy evsf)) |]
+            ==> Says Spy B X  # evsf \<in> set_cr"
 
 | Reception: --{*If A sends a message X to B, then B might receive it*}
-	     "[| evsr \<in> set_cr; Says A B X \<in> set evsr |]
+             "[| evsr \<in> set_cr; Says A B X \<in> set evsr |]
               ==> Gets B X  # evsr \<in> set_cr"
 
 | SET_CR1: --{*CardCInitReq: C initiates a run, sending a nonce to CCA*}
-	     "[| evs1 \<in> set_cr;  C = Cardholder k;  Nonce NC1 \<notin> used evs1 |]
-	      ==> Says C (CA i) {|Agent C, Nonce NC1|} # evs1 \<in> set_cr"
+             "[| evs1 \<in> set_cr;  C = Cardholder k;  Nonce NC1 \<notin> used evs1 |]
+              ==> Says C (CA i) {|Agent C, Nonce NC1|} # evs1 \<in> set_cr"
 
 | SET_CR2: --{*CardCInitRes: CA responds sending NC1 and its certificates*}
-	     "[| evs2 \<in> set_cr;
-		 Gets (CA i) {|Agent C, Nonce NC1|} \<in> set evs2 |]
-	      ==> Says (CA i) C
-		       {|sign (priSK (CA i)) {|Agent C, Nonce NC1|},
-			 cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
-			 cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
-		    # evs2 \<in> set_cr"
+             "[| evs2 \<in> set_cr;
+                 Gets (CA i) {|Agent C, Nonce NC1|} \<in> set evs2 |]
+              ==> Says (CA i) C
+                       {|sign (priSK (CA i)) {|Agent C, Nonce NC1|},
+                         cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
+                         cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
+                    # evs2 \<in> set_cr"
 
 | SET_CR3:
    --{*RegFormReq: C sends his PAN and a new nonce to CA.
@@ -135,8 +137,8 @@
     Nonce NC2 \<notin> used evs3;
     Key KC1 \<notin> used evs3; KC1 \<in> symKeys;
     Gets C {|sign (invKey SKi) {|Agent X, Nonce NC1|},
-	     cert (CA i) EKi onlyEnc (priSK RCA),
-	     cert (CA i) SKi onlySig (priSK RCA)|}
+             cert (CA i) EKi onlyEnc (priSK RCA),
+             cert (CA i) SKi onlySig (priSK RCA)|}
        \<in> set evs3;
     Says C (CA i) {|Agent C, Nonce NC1|} \<in> set evs3|]
  ==> Says C (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan C)))
@@ -154,9 +156,9 @@
     Gets (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan X)))
        \<in> set evs4 |]
   ==> Says (CA i) C
-	  {|sign (priSK (CA i)) {|Agent C, Nonce NC2, Nonce NCA|},
-	    cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
-	    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
+          {|sign (priSK (CA i)) {|Agent C, Nonce NC2, Nonce NCA|},
+            cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
+            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
        # evs4 \<in> set_cr"
 
 | SET_CR5:
@@ -177,10 +179,10 @@
          \<in> set evs5 |]
 ==> Says C (CA i)
          {|Crypt KC3
-	     {|Agent C, Nonce NC3, Key KC2, Key (pubSK C),
-	       Crypt (priSK C)
-	         (Hash {|Agent C, Nonce NC3, Key KC2,
-			 Key (pubSK C), Pan (pan C), Nonce CardSecret|})|},
+             {|Agent C, Nonce NC3, Key KC2, Key (pubSK C),
+               Crypt (priSK C)
+                 (Hash {|Agent C, Nonce NC3, Key KC2,
+                         Key (pubSK C), Pan (pan C), Nonce CardSecret|})|},
            Crypt EKi {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
     # Notes C {|Key KC2, Agent (CA i)|}
     # Notes C {|Key KC3, Agent (CA i)|}
@@ -204,15 +206,15 @@
       {|Crypt KC3 {|Agent C, Nonce NC3, Key KC2, Key cardSK,
                     Crypt (invKey cardSK)
                       (Hash {|Agent C, Nonce NC3, Key KC2,
-			      Key cardSK, Pan (pan C), Nonce CardSecret|})|},
+                              Key cardSK, Pan (pan C), Nonce CardSecret|})|},
         Crypt (pubEK (CA i)) {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
       \<in> set evs6 |]
 ==> Says (CA i) C
          (Crypt KC2
-	  {|sign (priSK (CA i))
-	         {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
-	    certC (pan C) cardSK (XOR(CardSecret,NonceCCA)) onlySig (priSK (CA i)),
-	    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
+          {|sign (priSK (CA i))
+                 {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
+            certC (pan C) cardSK (XOR(CardSecret,NonceCCA)) onlySig (priSK (CA i)),
+            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
       # Notes (CA i) (Key cardSK)
       # evs6 \<in> set_cr"
 
@@ -253,15 +255,15 @@
        set_cr.Nil 
         [THEN set_cr.SET_CR1 [of concl: C i NC1], 
          THEN Says_to_Gets, 
-	 THEN set_cr.SET_CR2 [of concl: i C NC1], 
-	 THEN Says_to_Gets,  
-	 THEN set_cr.SET_CR3 [of concl: C i KC1 _ NC2], 
-	 THEN Says_to_Gets,  
-	 THEN set_cr.SET_CR4 [of concl: i C NC2 NCA], 
-	 THEN Says_to_Gets,  
-	 THEN set_cr.SET_CR5 [of concl: C i KC3 NC3 KC2 CardSecret],
-	 THEN Says_to_Gets,  
-	 THEN set_cr.SET_CR6 [of concl: i C KC2]])
+         THEN set_cr.SET_CR2 [of concl: i C NC1], 
+         THEN Says_to_Gets,  
+         THEN set_cr.SET_CR3 [of concl: C i KC1 _ NC2], 
+         THEN Says_to_Gets,  
+         THEN set_cr.SET_CR4 [of concl: i C NC2 NCA], 
+         THEN Says_to_Gets,  
+         THEN set_cr.SET_CR5 [of concl: C i KC3 NC3 KC2 CardSecret],
+         THEN Says_to_Gets,  
+         THEN set_cr.SET_CR6 [of concl: i C KC2]])
 apply basic_possibility
 apply (simp_all (no_asm_simp) add: symKeys_neq_imp_neq)
 done
--- a/src/HOL/SET-Protocol/EventSET.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SET-Protocol/EventSET.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,7 @@
-(*  Title:      HOL/Auth/SET/EventSET
-    ID:         $Id$
-    Authors:     Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
+(*  Title:      HOL/SET-Protocol/EventSET.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
 *)
 
 header{*Theory of Events for SET*}
@@ -15,7 +16,7 @@
 text{*Message events*}
 datatype
   event = Says  agent agent msg
-	| Gets  agent	    msg
+        | Gets  agent       msg
         | Notes agent       msg
 
 
@@ -44,19 +45,19 @@
 knows_Cons:
     "knows A (ev # evs) =
        (if A = Spy then
-	(case ev of
-	   Says A' B X => insert X (knows Spy evs)
-	 | Gets A' X => knows Spy evs
-	 | Notes A' X  =>
-	     if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
-	else
-	(case ev of
-	   Says A' B X =>
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Gets A' X    =>
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Notes A' X    =>
-	     if A'=A then insert X (knows A evs) else knows A evs))"
+        (case ev of
+           Says A' B X => insert X (knows Spy evs)
+         | Gets A' X => knows Spy evs
+         | Notes A' X  =>
+             if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
+        else
+        (case ev of
+           Says A' B X =>
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Gets A' X    =>
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Notes A' X    =>
+             if A'=A then insert X (knows A evs) else knows A evs))"
 
 
 subsection{*Used Messages*}
@@ -70,10 +71,10 @@
 primrec
   used_Nil:  "used []         = (UN B. parts (initState B))"
   used_Cons: "used (ev # evs) =
-	         (case ev of
-		    Says A B X => parts {X} Un (used evs)
-         	  | Gets A X   => used evs
-		  | Notes A X  => parts {X} Un (used evs))"
+                 (case ev of
+                    Says A B X => parts {X} Un (used evs)
+                  | Gets A X   => used evs
+                  | Notes A X  => parts {X} Un (used evs))"
 
 
 
--- a/src/HOL/SET-Protocol/Merchant_Registration.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SET-Protocol/Merchant_Registration.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,7 @@
-(*  Title:      HOL/Auth/SET/Merchant_Registration
-    ID:         $Id$
-    Authors:    Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
+(*  Title:      HOL/SET-Protocol/Merchant_Registration.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
 *)
 
 header{*The SET Merchant Registration Protocol*}
@@ -18,21 +19,21 @@
 where
 
   Nil:    --{*Initial trace is empty*}
-	   "[] \<in> set_mr"
+           "[] \<in> set_mr"
 
 
 | Fake:    --{*The spy MAY say anything he CAN say.*}
-	   "[| evsf \<in> set_mr; X \<in> synth (analz (knows Spy evsf)) |]
-	    ==> Says Spy B X  # evsf \<in> set_mr"
-	
+           "[| evsf \<in> set_mr; X \<in> synth (analz (knows Spy evsf)) |]
+            ==> Says Spy B X  # evsf \<in> set_mr"
+        
 
 | Reception: --{*If A sends a message X to B, then B might receive it*}
-	     "[| evsr \<in> set_mr; Says A B X \<in> set evsr |]
+             "[| evsr \<in> set_mr; Says A B X \<in> set evsr |]
               ==> Gets B X  # evsr \<in> set_mr"
 
 
 | SET_MR1: --{*RegFormReq: M requires a registration form to a CA*}
- 	   "[| evs1 \<in> set_mr; M = Merchant k; Nonce NM1 \<notin> used evs1 |]
+           "[| evs1 \<in> set_mr; M = Merchant k; Nonce NM1 \<notin> used evs1 |]
             ==> Says M (CA i) {|Agent M, Nonce NM1|} # evs1 \<in> set_mr"
 
 
@@ -41,9 +42,9 @@
   "[| evs2 \<in> set_mr; Nonce NCA \<notin> used evs2;
       Gets (CA i) {|Agent M, Nonce NM1|} \<in> set evs2 |]
    ==> Says (CA i) M {|sign (priSK (CA i)) {|Agent M, Nonce NM1, Nonce NCA|},
-	               cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
+                       cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
                        cert (CA i) (pubSK (CA i)) onlySig (priSK RCA) |}
-	 # evs2 \<in> set_mr"
+         # evs2 \<in> set_mr"
 
 | SET_MR3:
          --{*CertReq: M submits the key pair to be certified.  The Notes
@@ -55,16 +56,16 @@
   "[| evs3 \<in> set_mr; M = Merchant k; Nonce NM2 \<notin> used evs3;
       Key KM1 \<notin> used evs3;  KM1 \<in> symKeys;
       Gets M {|sign (invKey SKi) {|Agent X, Nonce NM1, Nonce NCA|},
-	       cert (CA i) EKi onlyEnc (priSK RCA),
-	       cert (CA i) SKi onlySig (priSK RCA) |}
-	\<in> set evs3;
+               cert (CA i) EKi onlyEnc (priSK RCA),
+               cert (CA i) SKi onlySig (priSK RCA) |}
+        \<in> set evs3;
       Says M (CA i) {|Agent M, Nonce NM1|} \<in> set evs3 |]
    ==> Says M (CA i)
-	    {|Crypt KM1 (sign (priSK M) {|Agent M, Nonce NM2,
-					  Key (pubSK M), Key (pubEK M)|}),
-	      Crypt EKi (Key KM1)|}
-	 # Notes M {|Key KM1, Agent (CA i)|}
-	 # evs3 \<in> set_mr"
+            {|Crypt KM1 (sign (priSK M) {|Agent M, Nonce NM2,
+                                          Key (pubSK M), Key (pubEK M)|}),
+              Crypt EKi (Key KM1)|}
+         # Notes M {|Key KM1, Agent (CA i)|}
+         # evs3 \<in> set_mr"
 
 | SET_MR4:
          --{*CertRes: CA issues the certificates for merSK and merEK,
@@ -74,20 +75,20 @@
              CertRes shall be signed but not encrypted if the EE is a Merchant
              or Payment Gateway."-- Programmer's Guide, page 191.*}
     "[| evs4 \<in> set_mr; M = Merchant k;
-	merSK \<notin> symKeys;  merEK \<notin> symKeys;
-	Notes (CA i) (Key merSK) \<notin> set evs4;
-	Notes (CA i) (Key merEK) \<notin> set evs4;
-	Gets (CA i) {|Crypt KM1 (sign (invKey merSK)
-				 {|Agent M, Nonce NM2, Key merSK, Key merEK|}),
-		      Crypt (pubEK (CA i)) (Key KM1) |}
-	  \<in> set evs4 |]
+        merSK \<notin> symKeys;  merEK \<notin> symKeys;
+        Notes (CA i) (Key merSK) \<notin> set evs4;
+        Notes (CA i) (Key merEK) \<notin> set evs4;
+        Gets (CA i) {|Crypt KM1 (sign (invKey merSK)
+                                 {|Agent M, Nonce NM2, Key merSK, Key merEK|}),
+                      Crypt (pubEK (CA i)) (Key KM1) |}
+          \<in> set evs4 |]
     ==> Says (CA i) M {|sign (priSK(CA i)) {|Agent M, Nonce NM2, Agent(CA i)|},
-			cert  M      merSK    onlySig (priSK (CA i)),
-			cert  M      merEK    onlyEnc (priSK (CA i)),
-			cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
-	  # Notes (CA i) (Key merSK)
-	  # Notes (CA i) (Key merEK)
-	  # evs4 \<in> set_mr"
+                        cert  M      merSK    onlySig (priSK (CA i)),
+                        cert  M      merEK    onlyEnc (priSK (CA i)),
+                        cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
+          # Notes (CA i) (Key merSK)
+          # Notes (CA i) (Key merEK)
+          # evs4 \<in> set_mr"
 
 
 text{*Note possibility proofs are missing.*}
@@ -318,9 +319,9 @@
 
 lemma msg4_Says_imp_Notes:
  "[|Says (CA i) M {|sign (priSK (CA i)) {|Agent M, Nonce NM2, Agent (CA i)|},
-		    cert  M      merSK    onlySig (priSK (CA i)),
-		    cert  M      merEK    onlyEnc (priSK (CA i)),
-		    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
+                    cert  M      merSK    onlySig (priSK (CA i)),
+                    cert  M      merEK    onlyEnc (priSK (CA i)),
+                    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
     evs \<in> set_mr |]
   ==> Notes (CA i) (Key merSK) \<in> set evs
    &  Notes (CA i) (Key merEK) \<in> set evs"
@@ -333,13 +334,13 @@
   merSK uniquely identifies the other components, including merEK*}
 lemma merSK_unicity:
  "[|Says (CA i) M {|sign (priSK(CA i)) {|Agent M, Nonce NM2, Agent (CA i)|},
-		    cert  M      merSK    onlySig (priSK (CA i)),
-		    cert  M      merEK    onlyEnc (priSK (CA i)),
-		    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
+                    cert  M      merSK    onlySig (priSK (CA i)),
+                    cert  M      merEK    onlyEnc (priSK (CA i)),
+                    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
     Says (CA i) M' {|sign (priSK(CA i)) {|Agent M', Nonce NM2', Agent (CA i)|},
-		    cert  M'      merSK    onlySig (priSK (CA i)),
-		    cert  M'      merEK'    onlyEnc (priSK (CA i)),
-		    cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)|} \<in> set evs;
+                    cert  M'      merSK    onlySig (priSK (CA i)),
+                    cert  M'      merEK'    onlyEnc (priSK (CA i)),
+                    cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)|} \<in> set evs;
     evs \<in> set_mr |] ==> M=M' & NM2=NM2' & merEK=merEK'"
 apply (erule rev_mp)
 apply (erule rev_mp)
@@ -352,13 +353,13 @@
   merEK uniquely identifies the other components, including merSK*}
 lemma merEK_unicity:
  "[|Says (CA i) M {|sign (priSK(CA i)) {|Agent M, Nonce NM2, Agent (CA i)|},
-		    cert  M      merSK    onlySig (priSK (CA i)),
-		    cert  M      merEK    onlyEnc (priSK (CA i)),
-		    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
+                    cert  M      merSK    onlySig (priSK (CA i)),
+                    cert  M      merEK    onlyEnc (priSK (CA i)),
+                    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
     Says (CA i) M' {|sign (priSK(CA i)) {|Agent M', Nonce NM2', Agent (CA i)|},
-		     cert  M'      merSK'    onlySig (priSK (CA i)),
-		     cert  M'      merEK    onlyEnc (priSK (CA i)),
-		     cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)|} \<in> set evs;
+                     cert  M'      merSK'    onlySig (priSK (CA i)),
+                     cert  M'      merEK    onlyEnc (priSK (CA i)),
+                     cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)|} \<in> set evs;
     evs \<in> set_mr |] 
   ==> M=M' & NM2=NM2' & merSK=merSK'"
 apply (erule rev_mp)
--- a/src/HOL/SET-Protocol/MessageSET.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SET-Protocol/MessageSET.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,7 @@
-(*  Title:      HOL/Auth/SET/MessageSET
-    Authors:     Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
+(*  Title:      HOL/SET-Protocol/MessageSET.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
 *)
 
 header{*The Message Theory, Modified for SET*}
@@ -57,14 +59,14 @@
 
 text{*Messages*}
 datatype
-     msg = Agent  agent	    --{*Agent names*}
+     msg = Agent  agent     --{*Agent names*}
          | Number nat       --{*Ordinary integers, timestamps, ...*}
          | Nonce  nat       --{*Unguessable nonces*}
          | Pan    nat       --{*Unguessable Primary Account Numbers (??)*}
          | Key    key       --{*Crypto keys*}
-	 | Hash   msg       --{*Hashing*}
-	 | MPair  msg msg   --{*Compound messages*}
-	 | Crypt  key msg   --{*Encryption, public- or shared-key*}
+         | Hash   msg       --{*Hashing*}
+         | MPair  msg msg   --{*Compound messages*}
+         | Crypt  key msg   --{*Encryption, public- or shared-key*}
 
 
 (*Concrete syntax: messages appear as {|A,B,NA|}, etc...*)
@@ -82,10 +84,10 @@
 constdefs
   nat_of_agent :: "agent => nat"
    "nat_of_agent == agent_case (curry nat2_to_nat 0)
-			       (curry nat2_to_nat 1)
-			       (curry nat2_to_nat 2)
-			       (curry nat2_to_nat 3)
-			       (nat2_to_nat (4,0))"
+                               (curry nat2_to_nat 1)
+                               (curry nat2_to_nat 2)
+                               (curry nat2_to_nat 3)
+                               (nat2_to_nat (4,0))"
     --{*maps each agent to a unique natural number, for specifications*}
 
 text{*The function is indeed injective*}
@@ -871,8 +873,8 @@
     (Fake_insert_simp_tac ss 1
      THEN
      IF_UNSOLVED (Blast.depth_tac
-		  (cs addIs [@{thm analz_insertI},
-				   impOfSubs @{thm analz_subset_parts}]) 4 1))
+                  (cs addIs [@{thm analz_insertI},
+                                   impOfSubs @{thm analz_subset_parts}]) 4 1))
 
 fun spy_analz_tac (cs,ss) i =
   DETERM
--- a/src/HOL/SET-Protocol/PublicSET.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SET-Protocol/PublicSET.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,7 @@
-(*  Title:      HOL/Auth/SET/PublicSET
-    Authors:    Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
+(*  Title:      HOL/SET-Protocol/PublicSET.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
 *)
 
 header{*The Public-Key Theory, Modified for SET*}
@@ -72,33 +74,33 @@
   initState_CA:
     "initState (CA i)  =
        (if i=0 then Key ` ({priEK RCA, priSK RCA} Un
-			    pubEK ` (range CA) Un pubSK ` (range CA))
-	else {Key (priEK (CA i)), Key (priSK (CA i)),
-	      Key (pubEK (CA i)), Key (pubSK (CA i)),
-	      Key (pubEK RCA), Key (pubSK RCA)})"
+                            pubEK ` (range CA) Un pubSK ` (range CA))
+        else {Key (priEK (CA i)), Key (priSK (CA i)),
+              Key (pubEK (CA i)), Key (pubSK (CA i)),
+              Key (pubEK RCA), Key (pubSK RCA)})"
 
   initState_Cardholder:
     "initState (Cardholder i)  =    
        {Key (priEK (Cardholder i)), Key (priSK (Cardholder i)),
-	Key (pubEK (Cardholder i)), Key (pubSK (Cardholder i)),
-	Key (pubEK RCA), Key (pubSK RCA)}"
+        Key (pubEK (Cardholder i)), Key (pubSK (Cardholder i)),
+        Key (pubEK RCA), Key (pubSK RCA)}"
 
   initState_Merchant:
     "initState (Merchant i)  =    
        {Key (priEK (Merchant i)), Key (priSK (Merchant i)),
-	Key (pubEK (Merchant i)), Key (pubSK (Merchant i)),
-	Key (pubEK RCA), Key (pubSK RCA)}"
+        Key (pubEK (Merchant i)), Key (pubSK (Merchant i)),
+        Key (pubEK RCA), Key (pubSK RCA)}"
 
   initState_PG:
     "initState (PG i)  = 
        {Key (priEK (PG i)), Key (priSK (PG i)),
-	Key (pubEK (PG i)), Key (pubSK (PG i)),
-	Key (pubEK RCA), Key (pubSK RCA)}"
+        Key (pubEK (PG i)), Key (pubSK (PG i)),
+        Key (pubEK RCA), Key (pubSK RCA)}"
 (*>*)
   initState_Spy:
     "initState Spy = Key ` (invKey ` pubEK ` bad Un
-			    invKey ` pubSK ` bad Un
-			    range pubEK Un range pubSK)"
+                            invKey ` pubSK ` bad Un
+                            range pubEK Un range pubSK)"
 
 
 text{*Injective mapping from agents to PANs: an agent can have only one card*}
@@ -441,22 +443,22 @@
   redundancy!*}
 lemmas abbrev_simps [simp] =
     parts_insert_signCert analz_insert_signCert keysFor_insert_signCert
-    sign_def	 [THEN def_abbrev_simp_relation]
-    sign_def	 [THEN def_abbrev_simp_function]
+    sign_def     [THEN def_abbrev_simp_relation]
+    sign_def     [THEN def_abbrev_simp_function]
     signCert_def [THEN def_abbrev_simp_relation]
     signCert_def [THEN def_abbrev_simp_function]
-    certC_def	 [THEN def_abbrev_simp_relation]
-    certC_def	 [THEN def_abbrev_simp_function]
-    cert_def	 [THEN def_abbrev_simp_relation]
-    cert_def	 [THEN def_abbrev_simp_function]
-    EXcrypt_def	 [THEN def_abbrev_simp_relation]
-    EXcrypt_def	 [THEN def_abbrev_simp_function]
+    certC_def    [THEN def_abbrev_simp_relation]
+    certC_def    [THEN def_abbrev_simp_function]
+    cert_def     [THEN def_abbrev_simp_relation]
+    cert_def     [THEN def_abbrev_simp_function]
+    EXcrypt_def  [THEN def_abbrev_simp_relation]
+    EXcrypt_def  [THEN def_abbrev_simp_function]
     EXHcrypt_def [THEN def_abbrev_simp_relation]
     EXHcrypt_def [THEN def_abbrev_simp_function]
-    Enc_def	 [THEN def_abbrev_simp_relation]
-    Enc_def	 [THEN def_abbrev_simp_function]
-    EncB_def	 [THEN def_abbrev_simp_relation]
-    EncB_def	 [THEN def_abbrev_simp_function]
+    Enc_def      [THEN def_abbrev_simp_relation]
+    Enc_def      [THEN def_abbrev_simp_function]
+    EncB_def     [THEN def_abbrev_simp_relation]
+    EncB_def     [THEN def_abbrev_simp_function]
 
 
 subsubsection{*Elimination Rules for Controlled Rewriting *}
@@ -510,7 +512,7 @@
 
 lemmas analz_insert_simps = 
          analz_insert_subset_eq Un_upper2
-	 subset_insertI [THEN [2] subset_trans] 
+         subset_insertI [THEN [2] subset_trans] 
 
 
 subsection{*Freshness Lemmas*}
--- a/src/HOL/SET-Protocol/Purchase.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SET-Protocol/Purchase.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,7 @@
-(*  Title:      HOL/Auth/SET/Purchase
-    Authors:    Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
+(*  Title:      HOL/SET-Protocol/Purchase.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
 *)
 
 header{*Purchase Phase of SET*}
@@ -61,24 +63,24 @@
 where
 
   Nil:   --{*Initial trace is empty*}
-	 "[] \<in> set_pur"
+         "[] \<in> set_pur"
 
 | Fake:  --{*The spy MAY say anything he CAN say.*}
-	 "[| evsf \<in> set_pur;  X \<in> synth(analz(knows Spy evsf)) |]
-	  ==> Says Spy B X  # evsf \<in> set_pur"
+         "[| evsf \<in> set_pur;  X \<in> synth(analz(knows Spy evsf)) |]
+          ==> Says Spy B X  # evsf \<in> set_pur"
 
 
 | Reception: --{*If A sends a message X to B, then B might receive it*}
-	     "[| evsr \<in> set_pur;  Says A B X \<in> set evsr |]
-	      ==> Gets B X  # evsr \<in> set_pur"
+             "[| evsr \<in> set_pur;  Says A B X \<in> set evsr |]
+              ==> Gets B X  # evsr \<in> set_pur"
 
 | Start: 
       --{*Added start event which is out-of-band for SET: the Cardholder and
-	  the merchant agree on the amounts and uses @{text LID_M} as an
+          the merchant agree on the amounts and uses @{text LID_M} as an
           identifier.
-	  This is suggested by the External Interface Guide. The Programmer's
-	  Guide, in absence of @{text LID_M}, states that the merchant uniquely
-	  identifies the order out of some data contained in OrderDesc.*}
+          This is suggested by the External Interface Guide. The Programmer's
+          Guide, in absence of @{text LID_M}, states that the merchant uniquely
+          identifies the order out of some data contained in OrderDesc.*}
    "[|evsStart \<in> set_pur;
       Number LID_M \<notin> used evsStart;
       C = Cardholder k; M = Merchant i; P = PG j;
@@ -100,7 +102,7 @@
 
 | PInitRes:
      --{*Merchant replies with his own label XID and the encryption
-	 key certificate of his chosen Payment Gateway. Page 74 of Formal
+         key certificate of his chosen Payment Gateway. Page 74 of Formal
          Protocol Desc. We use @{text LID_M} to identify Cardholder*}
    "[|evsPIRes \<in> set_pur;
       Gets M {|Number LID_M, Nonce Chall_C|} \<in> set evsPIRes;
@@ -111,19 +113,19 @@
       Number XID \<notin> used evsPIRes;
       XID \<notin> range CardSecret; XID \<notin> range PANSecret|]
     ==> Says M C (sign (priSK M)
-		       {|Number LID_M, Number XID,
-			 Nonce Chall_C, Nonce Chall_M,
-			 cert P (pubEK P) onlyEnc (priSK RCA)|})
+                       {|Number LID_M, Number XID,
+                         Nonce Chall_C, Nonce Chall_M,
+                         cert P (pubEK P) onlyEnc (priSK RCA)|})
           # evsPIRes \<in> set_pur"
 
 | PReqUns:
       --{*UNSIGNED Purchase request (CardSecret = 0).
-	Page 79 of Formal Protocol Desc.
-	Merchant never sees the amount in clear. This holds of the real
-	protocol, where XID identifies the transaction. We omit
-	Hash{|Number XID, Nonce (CardSecret k)|} from PIHead because
-	the CardSecret is 0 and because AuthReq treated the unsigned case
-	very differently from the signed one anyway.*}
+        Page 79 of Formal Protocol Desc.
+        Merchant never sees the amount in clear. This holds of the real
+        protocol, where XID identifies the transaction. We omit
+        Hash{|Number XID, Nonce (CardSecret k)|} from PIHead because
+        the CardSecret is 0 and because AuthReq treated the unsigned case
+        very differently from the signed one anyway.*}
    "!!Chall_C Chall_M OrderDesc P PurchAmt XID evsPReqU.
     [|evsPReqU \<in> set_pur;
       C = Cardholder k; CardSecret k = 0;
@@ -133,24 +135,24 @@
       OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD,Nonce Chall_M|};
       PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M|};
       Gets C (sign (priSK M)
-		   {|Number LID_M, Number XID,
-		     Nonce Chall_C, Nonce Chall_M,
-		     cert P EKj onlyEnc (priSK RCA)|})
-	\<in> set evsPReqU;
+                   {|Number LID_M, Number XID,
+                     Nonce Chall_C, Nonce Chall_M,
+                     cert P EKj onlyEnc (priSK RCA)|})
+        \<in> set evsPReqU;
       Says C M {|Number LID_M, Nonce Chall_C|} \<in> set evsPReqU;
       Notes C {|Number LID_M, Transaction|} \<in> set evsPReqU |]
     ==> Says C M
-	     {|EXHcrypt KC1 EKj {|PIHead, Hash OIData|} (Pan (pan C)),
-	       OIData, Hash{|PIHead, Pan (pan C)|} |}
-	  # Notes C {|Key KC1, Agent M|}
-	  # evsPReqU \<in> set_pur"
+             {|EXHcrypt KC1 EKj {|PIHead, Hash OIData|} (Pan (pan C)),
+               OIData, Hash{|PIHead, Pan (pan C)|} |}
+          # Notes C {|Key KC1, Agent M|}
+          # evsPReqU \<in> set_pur"
 
 | PReqS:
       --{*SIGNED Purchase request.  Page 77 of Formal Protocol Desc.
           We could specify the equation
-	  @{term "PIReqSigned = {| PIDualSigned, OIDualSigned |}"}, since the
-	  Formal Desc. gives PIHead the same format in the unsigned case.
-	  However, there's little point, as P treats the signed and 
+          @{term "PIReqSigned = {| PIDualSigned, OIDualSigned |}"}, since the
+          Formal Desc. gives PIHead the same format in the unsigned case.
+          However, there's little point, as P treats the signed and 
           unsigned cases differently.*}
    "!!C Chall_C Chall_M EKj HOD KC2 LID_M M OIData
       OIDualSigned OrderDesc P PANData PIData PIDualSigned
@@ -162,22 +164,22 @@
       HOD = Hash{|Number OrderDesc, Number PurchAmt|};
       OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD, Nonce Chall_M|};
       PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
-		  Hash{|Number XID, Nonce (CardSecret k)|}|};
+                  Hash{|Number XID, Nonce (CardSecret k)|}|};
       PANData = {|Pan (pan C), Nonce (PANSecret k)|};
       PIData = {|PIHead, PANData|};
       PIDualSigned = {|sign (priSK C) {|Hash PIData, Hash OIData|},
-		       EXcrypt KC2 EKj {|PIHead, Hash OIData|} PANData|};
+                       EXcrypt KC2 EKj {|PIHead, Hash OIData|} PANData|};
       OIDualSigned = {|OIData, Hash PIData|};
       Gets C (sign (priSK M)
-		   {|Number LID_M, Number XID,
-		     Nonce Chall_C, Nonce Chall_M,
-		     cert P EKj onlyEnc (priSK RCA)|})
-	\<in> set evsPReqS;
+                   {|Number LID_M, Number XID,
+                     Nonce Chall_C, Nonce Chall_M,
+                     cert P EKj onlyEnc (priSK RCA)|})
+        \<in> set evsPReqS;
       Says C M {|Number LID_M, Nonce Chall_C|} \<in> set evsPReqS;
       Notes C {|Number LID_M, Transaction|} \<in> set evsPReqS |]
     ==> Says C M {|PIDualSigned, OIDualSigned|}
-	  # Notes C {|Key KC2, Agent M|}
-	  # evsPReqS \<in> set_pur"
+          # Notes C {|Key KC2, Agent M|}
+          # evsPReqS \<in> set_pur"
 
   --{*Authorization Request.  Page 92 of Formal Protocol Desc.
     Sent in response to Purchase Request.*}
@@ -187,20 +189,20 @@
        Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
        HOD = Hash{|Number OrderDesc, Number PurchAmt|};
        OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD,
-		  Nonce Chall_M|};
+                  Nonce Chall_M|};
        CardSecret k \<noteq> 0 -->
-	 P_I = {|sign (priSK C) {|HPIData, Hash OIData|}, encPANData|};
+         P_I = {|sign (priSK C) {|HPIData, Hash OIData|}, encPANData|};
        Gets M {|P_I, OIData, HPIData|} \<in> set evsAReq;
        Says M C (sign (priSK M) {|Number LID_M, Number XID,
-				  Nonce Chall_C, Nonce Chall_M,
-				  cert P EKj onlyEnc (priSK RCA)|})
-	 \<in> set evsAReq;
-	Notes M {|Number LID_M, Agent P, Transaction|}
-	   \<in> set evsAReq |]
+                                  Nonce Chall_C, Nonce Chall_M,
+                                  cert P EKj onlyEnc (priSK RCA)|})
+         \<in> set evsAReq;
+        Notes M {|Number LID_M, Agent P, Transaction|}
+           \<in> set evsAReq |]
     ==> Says M P
-	     (EncB (priSK M) KM (pubEK P)
-	       {|Number LID_M, Number XID, Hash OIData, HOD|}	P_I)
-	  # evsAReq \<in> set_pur"
+             (EncB (priSK M) KM (pubEK P)
+               {|Number LID_M, Number XID, Hash OIData, HOD|}   P_I)
+          # evsAReq \<in> set_pur"
 
   --{*Authorization Response has two forms: for UNSIGNED and SIGNED PIs.
     Page 99 of Formal Protocol Desc.
@@ -219,12 +221,12 @@
        PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M|};
        P_I = EXHcrypt KC1 EKj {|PIHead, HOIData|} (Pan (pan C));
        Gets P (EncB (priSK M) KM (pubEK P)
-	       {|Number LID_M, Number XID, HOIData, HOD|} P_I)
-	   \<in> set evsAResU |]
+               {|Number LID_M, Number XID, HOIData, HOD|} P_I)
+           \<in> set evsAResU |]
    ==> Says P M
-	    (EncB (priSK P) KP (pubEK M)
-	      {|Number LID_M, Number XID, Number PurchAmt|}
-	      authCode)
+            (EncB (priSK P) KP (pubEK M)
+              {|Number LID_M, Number XID, Number PurchAmt|}
+              authCode)
        # evsAResU \<in> set_pur"
 
 | AuthResS:
@@ -234,19 +236,19 @@
        Key KP \<notin> used evsAResS;  KP \<in> symKeys;
        CardSecret k \<noteq> 0;  KC2 \<in> symKeys;  KM \<in> symKeys;
        P_I = {|sign (priSK C) {|Hash PIData, HOIData|},
-	       EXcrypt KC2 (pubEK P) {|PIHead, HOIData|} PANData|};
+               EXcrypt KC2 (pubEK P) {|PIHead, HOIData|} PANData|};
        PANData = {|Pan (pan C), Nonce (PANSecret k)|};
        PIData = {|PIHead, PANData|};
        PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
-		  Hash{|Number XID, Nonce (CardSecret k)|}|};
+                  Hash{|Number XID, Nonce (CardSecret k)|}|};
        Gets P (EncB (priSK M) KM (pubEK P)
-		{|Number LID_M, Number XID, HOIData, HOD|}
-	       P_I)
-	   \<in> set evsAResS |]
+                {|Number LID_M, Number XID, HOIData, HOD|}
+               P_I)
+           \<in> set evsAResS |]
    ==> Says P M
-	    (EncB (priSK P) KP (pubEK M)
-	      {|Number LID_M, Number XID, Number PurchAmt|}
-	      authCode)
+            (EncB (priSK P) KP (pubEK M)
+              {|Number LID_M, Number XID, Number PurchAmt|}
+              authCode)
        # evsAResS \<in> set_pur"
 
 | PRes:
@@ -254,21 +256,21 @@
    "[| evsPRes \<in> set_pur;  KP \<in> symKeys;  M = Merchant i;
        Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
        Gets M (EncB (priSK P) KP (pubEK M)
-	      {|Number LID_M, Number XID, Number PurchAmt|}
-	      authCode)
-	  \<in> set evsPRes;
+              {|Number LID_M, Number XID, Number PurchAmt|}
+              authCode)
+          \<in> set evsPRes;
        Gets M {|Number LID_M, Nonce Chall_C|} \<in> set evsPRes;
        Says M P
-	    (EncB (priSK M) KM (pubEK P)
-	      {|Number LID_M, Number XID, Hash OIData, HOD|} P_I)
-	 \<in> set evsPRes;
+            (EncB (priSK M) KM (pubEK P)
+              {|Number LID_M, Number XID, Hash OIData, HOD|} P_I)
+         \<in> set evsPRes;
        Notes M {|Number LID_M, Agent P, Transaction|}
-	  \<in> set evsPRes
+          \<in> set evsPRes
       |]
    ==> Says M C
-	 (sign (priSK M) {|Number LID_M, Number XID, Nonce Chall_C,
-			   Hash (Number PurchAmt)|})
-	 # evsPRes \<in> set_pur"
+         (sign (priSK M) {|Number LID_M, Number XID, Nonce Chall_C,
+                           Hash (Number PurchAmt)|})
+         # evsPRes \<in> set_pur"
 
 
 specification (CardSecret PANSecret)
@@ -304,12 +306,12 @@
 lemma possibility_Uns:
     "[| CardSecret k = 0;
         C = Cardholder k;  M = Merchant i;
-	Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
-	KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
-	KC < KM; KM < KP;
-	Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
+        Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
+        KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
+        KC < KM; KM < KP;
+        Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
         Nonce Chall_M \<notin> used []; Chall_M \<notin> range CardSecret \<union> range PANSecret;
-	Chall_C < Chall_M; 
+        Chall_C < Chall_M; 
         Number LID_M \<notin> used []; LID_M \<notin> range CardSecret \<union> range PANSecret;
         Number XID \<notin> used []; XID \<notin> range CardSecret \<union> range PANSecret;
         LID_M < XID; XID < OrderDesc; OrderDesc < PurchAmt |] 
@@ -321,19 +323,19 @@
                   \<in> set evs" 
 apply (intro exI bexI)
 apply (rule_tac [2]
-	set_pur.Nil
+        set_pur.Nil
          [THEN set_pur.Start [of _ LID_M C k M i _ _ _ OrderDesc PurchAmt], 
-	  THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
+          THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
           THEN Says_to_Gets, 
-	  THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
+          THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
           THEN Says_to_Gets,
-	  THEN set_pur.PReqUns [of concl: C M KC],
+          THEN set_pur.PReqUns [of concl: C M KC],
           THEN Says_to_Gets, 
-	  THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
+          THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
           THEN Says_to_Gets, 
-	  THEN set_pur.AuthResUns [of concl: "PG j" M KP LID_M XID],
+          THEN set_pur.AuthResUns [of concl: "PG j" M KP LID_M XID],
           THEN Says_to_Gets, 
-	  THEN set_pur.PRes]) 
+          THEN set_pur.PRes]) 
 apply basic_possibility
 apply (simp_all add: used_Cons symKeys_neq_imp_neq) 
 done
@@ -341,12 +343,12 @@
 lemma possibility_S:
     "[| CardSecret k \<noteq> 0;
         C = Cardholder k;  M = Merchant i;
-	Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
-	KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
-	KC < KM; KM < KP;
-	Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
+        Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
+        KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
+        KC < KM; KM < KP;
+        Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
         Nonce Chall_M \<notin> used []; Chall_M \<notin> range CardSecret \<union> range PANSecret;
-	Chall_C < Chall_M; 
+        Chall_C < Chall_M; 
         Number LID_M \<notin> used []; LID_M \<notin> range CardSecret \<union> range PANSecret;
         Number XID \<notin> used []; XID \<notin> range CardSecret \<union> range PANSecret;
         LID_M < XID; XID < OrderDesc; OrderDesc < PurchAmt |] 
@@ -357,19 +359,19 @@
                \<in> set evs"
 apply (intro exI bexI)
 apply (rule_tac [2]
-	set_pur.Nil
+        set_pur.Nil
          [THEN set_pur.Start [of _ LID_M C k M i _ _ _ OrderDesc PurchAmt], 
-	  THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
+          THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
           THEN Says_to_Gets, 
-	  THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
+          THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
           THEN Says_to_Gets,
-	  THEN set_pur.PReqS [of concl: C M _ _ KC],
+          THEN set_pur.PReqS [of concl: C M _ _ KC],
           THEN Says_to_Gets, 
-	  THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
+          THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
           THEN Says_to_Gets, 
-	  THEN set_pur.AuthResS [of concl: "PG j" M KP LID_M XID],
+          THEN set_pur.AuthResS [of concl: "PG j" M KP LID_M XID],
           THEN Says_to_Gets, 
-	  THEN set_pur.PRes]) 
+          THEN set_pur.PRes]) 
 apply basic_possibility
 apply (auto simp add: used_Cons symKeys_neq_imp_neq) 
 done
@@ -761,7 +763,7 @@
 text{*When C receives PInitRes, he learns M's choice of P*}
 lemma C_verifies_PInitRes:
  "[| MsgPInitRes = {|Number LID_M, Number XID, Nonce Chall_C, Nonce Chall_M,
-	   cert P EKj onlyEnc (priSK RCA)|};
+           cert P EKj onlyEnc (priSK RCA)|};
      Crypt (priSK M) (Hash MsgPInitRes) \<in> parts (knows Spy evs);
      evs \<in> set_pur;  M \<notin> bad|]
   ==> \<exists>j trans.
@@ -826,14 +828,14 @@
       Crypt (priSK (PG j)) (Hash MsgAuthRes) \<in> parts (knows Spy evs);
       PG j \<notin> bad;  evs \<in> set_pur|]
    ==> \<exists>M KM KP HOIData HOD P_I.
-	Gets (PG j)
-	   (EncB (priSK M) KM (pubEK (PG j))
-		    {|Number LID_M, Number XID, HOIData, HOD|}
-		    P_I) \<in> set evs &
-	Says (PG j) M
-	     (EncB (priSK (PG j)) KP (pubEK M)
-	      {|Number LID_M, Number XID, Number PurchAmt|}
-	      authCode) \<in> set evs"
+        Gets (PG j)
+           (EncB (priSK M) KM (pubEK (PG j))
+                    {|Number LID_M, Number XID, HOIData, HOD|}
+                    P_I) \<in> set evs &
+        Says (PG j) M
+             (EncB (priSK (PG j)) KP (pubEK M)
+              {|Number LID_M, Number XID, Number PurchAmt|}
+              authCode) \<in> set evs"
 apply clarify
 apply (erule rev_mp)
 apply (erule set_pur.induct)
@@ -1030,9 +1032,9 @@
      Notes M {|Number LID_M, Agent P, extras|} \<in> set evs;
      M = Merchant i;  C = Cardholder k;  C \<notin> bad;  evs \<in> set_pur|]
   ==> \<exists>PIData PICrypt.
-	HPIData = Hash PIData &
-	Says C M {|{|sign (priSK C) MsgDualSign, PICrypt|}, OIData, Hash PIData|}
-	  \<in> set evs"
+        HPIData = Hash PIData &
+        Says C M {|{|sign (priSK C) MsgDualSign, PICrypt|}, OIData, Hash PIData|}
+          \<in> set evs"
 apply clarify
 apply (erule rev_mp)
 apply (erule rev_mp)
--- a/src/HOL/SetInterval.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SetInterval.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,7 +1,7 @@
 (*  Title:      HOL/SetInterval.thy
-    Author:     Tobias Nipkow and Clemens Ballarin
-                Additions by Jeremy Avigad in March 2004
-    Copyright   2000  TU Muenchen
+    Author:     Tobias Nipkow
+    Author:     Clemens Ballarin
+    Author:     Jeremy Avigad
 
 lessThan, greaterThan, atLeast, atMost and two-sided intervals
 *)
@@ -15,19 +15,19 @@
 context ord
 begin
 definition
-  lessThan    :: "'a => 'a set"	("(1{..<_})") where
+  lessThan    :: "'a => 'a set" ("(1{..<_})") where
   "{..<u} == {x. x < u}"
 
 definition
-  atMost      :: "'a => 'a set"	("(1{.._})") where
+  atMost      :: "'a => 'a set" ("(1{.._})") where
   "{..u} == {x. x \<le> u}"
 
 definition
-  greaterThan :: "'a => 'a set"	("(1{_<..})") where
+  greaterThan :: "'a => 'a set" ("(1{_<..})") where
   "{l<..} == {x. l<x}"
 
 definition
-  atLeast     :: "'a => 'a set"	("(1{_..})") where
+  atLeast     :: "'a => 'a set" ("(1{_..})") where
   "{l..} == {x. l\<le>x}"
 
 definition
--- a/src/HOL/SizeChange/Correctness.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SizeChange/Correctness.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/SCT_Theorem.thy
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 *)
 
@@ -450,19 +449,19 @@
     proof (cases rule:three_cases)
       assume "Suc l < j"
       with ths range show ?thesis 
-	unfolding is_fthread_def Ball_def
-	by simp
+        unfolding is_fthread_def Ball_def
+        by simp
     next
       assume "Suc l = j"
       hence "l < j" "\<theta>2 (Suc l) = \<theta>1 (Suc l)" by auto
       with ths range show ?thesis 
-	unfolding is_fthread_def Ball_def
-	by simp
+        unfolding is_fthread_def Ball_def
+        by simp
     next
       assume "j \<le> l"
       with ths range show ?thesis 
-	unfolding is_fthread_def Ball_def
-	by simp
+        unfolding is_fthread_def Ball_def
+        by simp
     qed arith
   qed
   moreover 
@@ -534,19 +533,19 @@
     proof (cases rule:three_cases)
       assume "Suc l < j"
       with ths2 range show ?thesis 
-	unfolding is_fthread_def Ball_def
-	by simp
+        unfolding is_fthread_def Ball_def
+        by simp
     next
       assume "Suc l = j"
       hence "l < j" "\<theta>2 (Suc l) = \<theta>1 (Suc l)" by auto
       with ths2 range show ?thesis 
-	unfolding is_fthread_def Ball_def
-	by simp
+        unfolding is_fthread_def Ball_def
+        by simp
     next
       assume "j \<le> l"
       with ths2 range show ?thesis 
-	unfolding is_fthread_def Ball_def
-	by simp
+        unfolding is_fthread_def Ball_def
+        by simp
     qed arith
   qed
   moreover
@@ -967,18 +966,18 @@
     proof (intro allI impI)
       fix i assume "n \<le> i"
       also have "i \<le> s i" 
-	using increasing_inc by auto
+        using increasing_inc by auto
       finally have "n \<le> s i" .
 
       with fr have "is_fthread \<theta> p (s i) (s (Suc i))"
-	unfolding is_fthread_def by auto
+        unfolding is_fthread_def by auto
       hence "has_fth p (s i) (s (Suc i)) (\<theta> (s i)) (\<theta> (s (Suc i)))"
-	unfolding has_fth_def by auto
+        unfolding has_fth_def by auto
       with less_imp_le[OF increasing_strict]
       have "eql (prod (p\<langle>s i,s (Suc i)\<rangle>)) (\<theta> (s i)) (\<theta> (s (Suc i)))"
-	by (simp add:Lemma7a)
+        by (simp add:Lemma7a)
       thus "eqlat (contract s p) ?c\<theta> i" unfolding contract_def
-	by auto
+        by auto
     qed
 
     show "\<exists>\<^sub>\<infinity>i. descat (contract s p) ?c\<theta> i"
@@ -988,30 +987,30 @@
 
       let ?K = "section_of s (max (s (Suc i)) n)"
       from `\<exists>\<^sub>\<infinity>i. descat p \<theta> i` obtain j
-	  where "s (Suc ?K) < j" "descat p \<theta> j"
-	unfolding INFM_nat by blast
+          where "s (Suc ?K) < j" "descat p \<theta> j"
+        unfolding INFM_nat by blast
       
       let ?L = "section_of s j"
       {
-	fix x assume r: "x \<in> section s ?L"
-	
-	have e1: "max (s (Suc i)) n < s (Suc ?K)" by (rule section_of2) simp
+        fix x assume r: "x \<in> section s ?L"
+        
+        have e1: "max (s (Suc i)) n < s (Suc ?K)" by (rule section_of2) simp
         note `s (Suc ?K) < j`
         also have "j < s (Suc ?L)"
           by (rule section_of2) simp
         finally have "Suc ?K \<le> ?L"
           by (simp add:increasing_bij)          
-	with increasing_weak have "s (Suc ?K) \<le> s ?L" by simp
-	with e1 r have "max (s (Suc i)) n < x" by simp
+        with increasing_weak have "s (Suc ?K) \<le> s ?L" by simp
+        with e1 r have "max (s (Suc i)) n < x" by simp
         
-	hence "(s (Suc i)) < x" "n < x" by auto
+        hence "(s (Suc i)) < x" "n < x" by auto
       }
       note range_est = this
       
       have "is_desc_fthread \<theta> p (s ?L) (s (Suc ?L))"
-	unfolding is_desc_fthread_def is_fthread_def
+        unfolding is_desc_fthread_def is_fthread_def
       proof
-	show "\<forall>m\<in>section s ?L. eqlat p \<theta> m"
+        show "\<forall>m\<in>section s ?L. eqlat p \<theta> m"
         proof 
           fix m assume "m\<in>section s ?L"
           with range_est(2) have "n < m" . 
@@ -1019,25 +1018,25 @@
         qed
 
         from in_section_of inc less_imp_le[OF `s (Suc ?K) < j`]
-	have "j \<in> section s ?L" .
+        have "j \<in> section s ?L" .
 
-	with `descat p \<theta> j`
-	show "\<exists>m\<in>section s ?L. descat p \<theta> m" ..
+        with `descat p \<theta> j`
+        show "\<exists>m\<in>section s ?L. descat p \<theta> m" ..
       qed
       
       with less_imp_le[OF increasing_strict]
       have a: "descat (contract s p) ?c\<theta> ?L"
-	unfolding contract_def Lemma7b[symmetric]
-	by (auto simp:Lemma7b[symmetric] has_desc_fth_def)
+        unfolding contract_def Lemma7b[symmetric]
+        by (auto simp:Lemma7b[symmetric] has_desc_fth_def)
       
       have "i < ?L"
       proof (rule classical)
-	assume "\<not> i < ?L" 
-	hence "s ?L < s (Suc i)" 
+        assume "\<not> i < ?L" 
+        hence "s ?L < s (Suc i)" 
           by (simp add:increasing_bij)
-	also have "\<dots> < s ?L"
-	  by (rule range_est(1)) (simp add:increasing_strict)
-	finally show ?thesis .
+        also have "\<dots> < s ?L"
+          by (rule range_est(1)) (simp add:increasing_strict)
+        finally show ?thesis .
       qed
       with a show "\<exists>l. i < l \<and> descat (contract s p) ?c\<theta> l"
         by blast
@@ -1185,7 +1184,7 @@
     have "f (set2pair {x, y}) \<in> S"
     proof (cases "x < y")
       case True hence "set2pair {x, y} = (x, y)"
-	by (rule set2pair_conv)
+        by (rule set2pair_conv)
       with True inS
       show ?thesis by simp
     next
@@ -1193,7 +1192,7 @@
       with neq have y_less: "y < x" by simp
       have x:"{x,y} = {y,x}" by auto
       with y_less have "set2pair {x, y} = (y, x)"
-	by (simp add:set2pair_conv)
+        by (simp add:set2pair_conv)
       with y_less inS
       show ?thesis by simp
     qed
@@ -1261,7 +1260,7 @@
       where k_pow: "A ^ k \<turnstile> n \<leadsto>\<^bsup>G\<^esup> n"
       and "0 < k"
       unfolding in_tcl by auto
-	
+        
     from power_induces_path k_pow
     obtain loop where loop_props:
       "has_fpath A loop"
@@ -1343,7 +1342,7 @@
       case base thus ?case by (simp add:sub_path_def)
     next
       case (step i) thus ?case
-	by (simp add:sub_path_def upt_rec[of i k] idemp)
+        by (simp add:sub_path_def upt_rec[of i k] idemp)
     qed
 
     with `i < j` `j < k` dfth Lemma7b[of i k ?cp p p]
@@ -1388,7 +1387,7 @@
     let ?q = "contract ?s p"
 
     note all_in_S[simp] = enumerate_in_set[OF `infinite S`]
-	from `infinite S` 
+        from `infinite S` 
     have inc[simp]: "increasing ?s" 
       unfolding increasing_def by (simp add:enumerate_mono)
     note increasing_bij[OF this, simp]
@@ -1398,7 +1397,7 @@
 
     from all_G G_struct 
     have all_H: "\<And>i. (snd (?q i)) = H"
-	  unfolding contract_def 
+          unfolding contract_def 
       by simp
     
     have loop: "(tcl A) \<turnstile> n \<leadsto>\<^bsup>H\<^esup> n"
@@ -1407,27 +1406,27 @@
       let ?i = "?s 0" and ?j = "?s (Suc 0)" and ?k = "?s (Suc (Suc 0))"
       
       have "pdesc (p\<langle>?i,?j\<rangle>) = G"
-		and "pdesc (p\<langle>?j,?k\<rangle>) = G"
-		and "pdesc (p\<langle>?i,?k\<rangle>) = G"
-		using all_G 
-		by auto
-	  
+                and "pdesc (p\<langle>?j,?k\<rangle>) = G"
+                and "pdesc (p\<langle>?i,?k\<rangle>) = G"
+                using all_G 
+                by auto
+          
       with G_struct 
       have "m = end_node (p\<langle>?i,?j\<rangle>)"
-				"n = fst (p\<langle>?j,?k\<rangle>)"
-				and Hs:	"prod (p\<langle>?i,?j\<rangle>) = H"
-				"prod (p\<langle>?j,?k\<rangle>) = H"
-				"prod (p\<langle>?i,?k\<rangle>) = H"
-		by auto
-			
+                                "n = fst (p\<langle>?j,?k\<rangle>)"
+                                and Hs: "prod (p\<langle>?i,?j\<rangle>) = H"
+                                "prod (p\<langle>?j,?k\<rangle>) = H"
+                                "prod (p\<langle>?i,?k\<rangle>) = H"
+                by auto
+                        
       hence "m = n" by simp
       thus "tcl A \<turnstile> n \<leadsto>\<^bsup>H\<^esup> n"
-		using G_struct `G \<in> dest_graph (tcl A)`
-		by (simp add:has_edge_def)
-	  
+                using G_struct `G \<in> dest_graph (tcl A)`
+                by (simp add:has_edge_def)
+          
       from sub_path_prod[of ?i ?j ?k p]      
       show "H * H = H"
-		unfolding Hs by simp
+                unfolding Hs by simp
     qed
     moreover have "\<And>k. \<not>dsc H k k"
     proof
@@ -1435,8 +1434,8 @@
       
       with all_H repeated_edge
       have "\<exists>\<theta>. is_desc_thread \<theta> ?q" by fast
-	  with inc have "\<exists>\<theta>. is_desc_thread \<theta> p"
-		by (subst contract_keeps_threads) 
+          with inc have "\<exists>\<theta>. is_desc_thread \<theta> p"
+                by (subst contract_keeps_threads) 
       with no_desc_th
       show False ..
     qed
--- a/src/HOL/SizeChange/Graphs.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SizeChange/Graphs.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/Graphs.thy
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 *)
 
@@ -710,7 +709,7 @@
   let ?p = "p\<langle>s i,s (Suc i)\<rangle>"
 
   from increasing_strict 
-	have "fst (p (s (Suc i))) = end_node ?p" by simp
+  have "fst (p (s (Suc i))) = end_node ?p" by simp
   moreover
   from increasing_strict[of s i "Suc i"] have "snd ?p \<noteq> []"
     by (simp add:sub_path_def)
--- a/src/HOL/SizeChange/Interpretation.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/SizeChange/Interpretation.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -27,27 +27,27 @@
   shows "\<exists>s. idseq R s x"
 proof -
   
-  have	"\<exists>f. \<forall>x. \<not>accp R x \<longrightarrow> R (f x) x \<and> \<not>accp R (f x)"
-	by (rule choice, auto simp:not_acc_smaller)
+  have "\<exists>f. \<forall>x. \<not>accp R x \<longrightarrow> R (f x) x \<and> \<not>accp R (f x)"
+    by (rule choice, auto simp:not_acc_smaller)
   
   then obtain f where
-	in_R: "\<And>x. \<not>accp R x \<Longrightarrow> R (f x) x"
-	and nia: "\<And>x. \<not>accp R x \<Longrightarrow> \<not>accp R (f x)"
-	by blast
+    in_R: "\<And>x. \<not>accp R x \<Longrightarrow> R (f x) x"
+    and nia: "\<And>x. \<not>accp R x \<Longrightarrow> \<not>accp R (f x)"
+    by blast
   
   let ?s = "\<lambda>i. (f ^^ i) x"
   
   {
-	fix i
-	have "\<not>accp R (?s i)"
-	  by (induct i) (auto simp:nia `\<not>accp R x`)
-	hence "R (f (?s i)) (?s i)"
-	  by (rule in_R)
+    fix i
+    have "\<not>accp R (?s i)"
+      by (induct i) (auto simp:nia `\<not>accp R x`)
+    hence "R (f (?s i)) (?s i)"
+      by (rule in_R)
   }
   
   hence "idseq R ?s x"
-	unfolding idseq_def
-	by auto
+    unfolding idseq_def
+    by auto
   
   thus ?thesis by auto
 qed
@@ -85,10 +85,10 @@
 proof -
   from idseq
   have a: "\<forall>i. \<exists>rd \<in> set rds. in_cdesc rd (s (Suc i)) (s i)"
-	by (auto simp:idseq_def intro:some_rd)
+    by (auto simp:idseq_def intro:some_rd)
   
   show ?thesis
-	by (rule choice) (insert a, blast)
+    by (rule choice) (insert a, blast)
 qed
 
 
@@ -145,34 +145,34 @@
   shows False
 proof -
   {
-	fix i j:: nat 
-	assume "n \<le> i"
-	assume "i \<le> j"
-	{
-	  fix k 
-	  have "s (i + k) \<le> s i"
-	  proof (induct k)
-		case 0 thus ?case by simp
-	  next
-		case (Suc k)
-		with leq[of "i + k"] `n \<le> i`
-		show ?case by simp
-	  qed
-	}
-	from this[of "j - i"] `n \<le> i` `i \<le> j`
-	have "s j \<le> s i" by auto
+    fix i j:: nat 
+    assume "n \<le> i"
+    assume "i \<le> j"
+    {
+      fix k 
+      have "s (i + k) \<le> s i"
+      proof (induct k)
+        case 0 thus ?case by simp
+      next
+        case (Suc k)
+        with leq[of "i + k"] `n \<le> i`
+        show ?case by simp
+      qed
+    }
+    from this[of "j - i"] `n \<le> i` `i \<le> j`
+    have "s j \<le> s i" by auto
   }
   note decr = this
   
   let ?min = "LEAST x. x \<in> range (\<lambda>i. s (n + i))"
   have "?min \<in> range (\<lambda>i. s (n + i))"
-	by (rule LeastI) auto
+    by (rule LeastI) auto
   then obtain k where min: "?min = s (n + k)" by auto
   
   from less 
   obtain k' where "n + k < k'"
-	and "s (Suc k') < s k'"
-	unfolding INFM_nat by auto
+    and "s (Suc k') < s k'"
+    unfolding INFM_nat by auto
   
   with decr[of "n + k" k'] min
   have "s (Suc k') < ?min" by auto
@@ -285,127 +285,127 @@
 proof (rule, rule classical)
   fix x
   assume "\<not> accp R x"
-  with non_acc_has_idseq	
+  with non_acc_has_idseq
   have "\<exists>s. idseq R s x" .
   then obtain s where "idseq R s x" ..
   hence "\<exists>cs. \<forall>i. cs i \<in> set RDs \<and>
-	in_cdesc (cs i) (s (Suc i)) (s i)"
-	unfolding R by (rule ex_cs) 
+    in_cdesc (cs i) (s (Suc i)) (s i)"
+    unfolding R by (rule ex_cs) 
   then obtain cs where
-	[simp]: "\<And>i. cs i \<in> set RDs"
-	  and ird[simp]: "\<And>i. in_cdesc (cs i) (s (Suc i)) (s i)"
-	by blast
+    [simp]: "\<And>i. cs i \<in> set RDs"
+      and ird[simp]: "\<And>i. in_cdesc (cs i) (s (Suc i)) (s i)"
+    by blast
   
   let ?cis = "\<lambda>i. index_of RDs (cs i)"
   have "\<forall>i. \<exists>G. (\<A> \<turnstile> ?cis i \<leadsto>\<^bsup>G\<^esup> (?cis (Suc i)))
-	\<and> approx G (RDs ! ?cis i) (RDs ! ?cis (Suc i)) 
-	(M ! ?cis i) (M ! ?cis (Suc i))" (is "\<forall>i. \<exists>G. ?P i G")
+    \<and> approx G (RDs ! ?cis i) (RDs ! ?cis (Suc i)) 
+    (M ! ?cis i) (M ! ?cis (Suc i))" (is "\<forall>i. \<exists>G. ?P i G")
   proof
-	fix i
-	let ?n = "?cis i" and ?n' = "?cis (Suc i)"
+    fix i
+    let ?n = "?cis i" and ?n' = "?cis (Suc i)"
     
-	have "in_cdesc (RDs ! ?n) (s (Suc i)) (s i)"
-	  "in_cdesc (RDs ! ?n') (s (Suc (Suc i))) (s (Suc i))"
-	  by (simp_all add:index_of_member)
-	with step_witness
- 	have "\<not> no_step (RDs ! ?n) (RDs ! ?n')" .
-	moreover have
-	  "?n < length RDs" 
-	  "?n' < length RDs"
-	  by (simp_all add:index_of_length[symmetric])
-	ultimately
-	obtain G
-	  where "\<A> \<turnstile> ?n \<leadsto>\<^bsup>G\<^esup> ?n'"
-	  and "approx G (RDs ! ?n) (RDs ! ?n') (M ! ?n) (M ! ?n')"
-	  using sound
-	  unfolding sound_int_def by auto
+    have "in_cdesc (RDs ! ?n) (s (Suc i)) (s i)"
+      "in_cdesc (RDs ! ?n') (s (Suc (Suc i))) (s (Suc i))"
+      by (simp_all add:index_of_member)
+    with step_witness
+    have "\<not> no_step (RDs ! ?n) (RDs ! ?n')" .
+    moreover have
+      "?n < length RDs" 
+      "?n' < length RDs"
+      by (simp_all add:index_of_length[symmetric])
+    ultimately
+    obtain G
+      where "\<A> \<turnstile> ?n \<leadsto>\<^bsup>G\<^esup> ?n'"
+      and "approx G (RDs ! ?n) (RDs ! ?n') (M ! ?n) (M ! ?n')"
+      using sound
+      unfolding sound_int_def by auto
     
-	thus "\<exists>G. ?P i G" by blast
+    thus "\<exists>G. ?P i G" by blast
   qed
   with choice
   have "\<exists>Gs. \<forall>i. ?P i (Gs i)" .
   then obtain Gs where 
-	A: "\<And>i. \<A> \<turnstile> ?cis i \<leadsto>\<^bsup>(Gs i)\<^esup> (?cis (Suc i))" 
-	and B: "\<And>i. approx (Gs i) (RDs ! ?cis i) (RDs ! ?cis (Suc i)) 
-	(M ! ?cis i) (M ! ?cis (Suc i))"
-	by blast
+    A: "\<And>i. \<A> \<turnstile> ?cis i \<leadsto>\<^bsup>(Gs i)\<^esup> (?cis (Suc i))" 
+    and B: "\<And>i. approx (Gs i) (RDs ! ?cis i) (RDs ! ?cis (Suc i)) 
+    (M ! ?cis i) (M ! ?cis (Suc i))"
+    by blast
   
   let ?p = "\<lambda>i. (?cis i, Gs i)"
   
   from A have "has_ipath \<A> ?p"
-	unfolding has_ipath_def
-	by auto
+    unfolding has_ipath_def
+    by auto
   
   with `SCT \<A>` SCT_def 
   obtain th where "is_desc_thread th ?p"
-	by auto
+    by auto
   
   then obtain n
-	where fr: "\<forall>i\<ge>n. eqlat ?p th i"
-	and inf: "\<exists>\<^sub>\<infinity>i. descat ?p th i"
-	unfolding is_desc_thread_def by auto
+    where fr: "\<forall>i\<ge>n. eqlat ?p th i"
+    and inf: "\<exists>\<^sub>\<infinity>i. descat ?p th i"
+    unfolding is_desc_thread_def by auto
   
   from B
   have approx:
-	"\<And>i. approx (Gs i) (cs i) (cs (Suc i)) 
-	(M ! ?cis i) (M ! ?cis (Suc i))"
-	by (simp add:index_of_member)
+    "\<And>i. approx (Gs i) (cs i) (cs (Suc i)) 
+    (M ! ?cis i) (M ! ?cis (Suc i))"
+    by (simp add:index_of_member)
   
   let ?seq = "\<lambda>i. (M ! ?cis i) (th i) (s i)"
   
   have "\<And>i. n < i \<Longrightarrow> ?seq (Suc i) \<le> ?seq i"
   proof -
-	fix i 
-	let ?q1 = "th i" and ?q2 = "th (Suc i)"
-	assume "n < i"
-	
-	with fr	have "eqlat ?p th i" by simp 
-	hence "dsc (Gs i) ?q1 ?q2 \<or> eqp (Gs i) ?q1 ?q2" 
+    fix i 
+    let ?q1 = "th i" and ?q2 = "th (Suc i)"
+    assume "n < i"
+    
+    with fr have "eqlat ?p th i" by simp 
+    hence "dsc (Gs i) ?q1 ?q2 \<or> eqp (Gs i) ?q1 ?q2" 
       by simp
-	thus "?seq (Suc i) \<le> ?seq i"
-	proof
-	  assume "dsc (Gs i) ?q1 ?q2"
-	  
-	  with approx
-	  have a:"decr (cs i) (cs (Suc i)) 
-		((M ! ?cis i) ?q1) ((M ! ?cis (Suc i)) ?q2)" 
-		unfolding approx_def by auto
+    thus "?seq (Suc i) \<le> ?seq i"
+    proof
+      assume "dsc (Gs i) ?q1 ?q2"
+      
+      with approx
+      have a:"decr (cs i) (cs (Suc i)) 
+        ((M ! ?cis i) ?q1) ((M ! ?cis (Suc i)) ?q2)" 
+        unfolding approx_def by auto
       
-	  show ?thesis
-		apply (rule less_imp_le)
-		apply (rule decr_in_cdesc[of _ "s (Suc i)" "s i"])
-		by (rule ird a)+
-	next
-	  assume "eqp (Gs i) ?q1 ?q2"
-	  
-	  with approx
-	  have a:"decreq (cs i) (cs (Suc i)) 
-		((M ! ?cis i) ?q1) ((M ! ?cis (Suc i)) ?q2)" 
-		unfolding approx_def by auto
+      show ?thesis
+        apply (rule less_imp_le)
+        apply (rule decr_in_cdesc[of _ "s (Suc i)" "s i"])
+        by (rule ird a)+
+    next
+      assume "eqp (Gs i) ?q1 ?q2"
       
-	  show ?thesis
-		apply (rule decreq_in_cdesc[of _ "s (Suc i)" "s i"])
-		by (rule ird a)+
-	qed
+      with approx
+      have a:"decreq (cs i) (cs (Suc i)) 
+        ((M ! ?cis i) ?q1) ((M ! ?cis (Suc i)) ?q2)" 
+        unfolding approx_def by auto
+      
+      show ?thesis
+        apply (rule decreq_in_cdesc[of _ "s (Suc i)" "s i"])
+        by (rule ird a)+
+    qed
   qed
   moreover have "\<exists>\<^sub>\<infinity>i. ?seq (Suc i) < ?seq i" unfolding INFM_nat
   proof 
-	fix i 
-	from inf obtain j where "i < j" and d: "descat ?p th j"
-	  unfolding INFM_nat by auto
-	let ?q1 = "th j" and ?q2 = "th (Suc j)"
-	from d have "dsc (Gs j) ?q1 ?q2" by auto
-	
-	with approx
-	have a:"decr (cs j) (cs (Suc j)) 
-	  ((M ! ?cis j) ?q1) ((M ! ?cis (Suc j)) ?q2)" 
-	  unfolding approx_def by auto
+    fix i 
+    from inf obtain j where "i < j" and d: "descat ?p th j"
+      unfolding INFM_nat by auto
+    let ?q1 = "th j" and ?q2 = "th (Suc j)"
+    from d have "dsc (Gs j) ?q1 ?q2" by auto
     
-	have "?seq (Suc j) < ?seq j"
-	  apply (rule decr_in_cdesc[of _ "s (Suc j)" "s j"])
-	  by (rule ird a)+
-	with `i < j` 
-	show "\<exists>j. i < j \<and> ?seq (Suc j) < ?seq j" by auto
+    with approx
+    have a:"decr (cs j) (cs (Suc j)) 
+      ((M ! ?cis j) ?q1) ((M ! ?cis (Suc j)) ?q2)" 
+      unfolding approx_def by auto
+    
+    have "?seq (Suc j) < ?seq j"
+      apply (rule decr_in_cdesc[of _ "s (Suc j)" "s j"])
+      by (rule ird a)+
+    with `i < j` 
+    show "\<exists>j. i < j \<and> ?seq (Suc j) < ?seq j" by auto
   qed
   ultimately have False
     by (rule no_inf_desc_nat_sequence[of "Suc n"]) simp
--- a/src/HOL/Statespace/DistinctTreeProver.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Statespace/DistinctTreeProver.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -132,15 +132,15 @@
       case (Some r')
       with Node.hyps
       have "set_of r' \<subseteq> set_of r"
-	by simp
+        by simp
       with l'_l Some x_l_Some del
       show ?thesis
-	by (auto split: split_if_asm)
+        by (auto split: split_if_asm)
     next
       case None
       with l'_l Some x_l_Some del
       show ?thesis
-	by (fastsimp split: split_if_asm)
+        by (fastsimp split: split_if_asm)
     qed
   next
     case None
@@ -150,15 +150,15 @@
       case (Some r')
       with Node.hyps
       have "set_of r' \<subseteq> set_of r"
-	by simp
+        by simp
       with Some x_l_None del
       show ?thesis
-	by (fastsimp split: split_if_asm)
+        by (fastsimp split: split_if_asm)
     next
       case None
       with x_l_None del
       show ?thesis
-	by (fastsimp split: split_if_asm)
+        by (fastsimp split: split_if_asm)
     qed
   qed
 qed
@@ -191,18 +191,18 @@
       case (Some r')
       from Node.hyps (2) [OF Some dist_r]
       have dist_r': "all_distinct r'"
-	by simp
+        by simp
       from delete_Some_set_of [OF Some]
       have "set_of r' \<subseteq> set_of r".
       
       with dist_l' dist_r' l'_l Some x_l_Some del d dist_l_r
       show ?thesis
-	by fastsimp
+        by fastsimp
     next
       case None
       with l'_l dist_l'  x_l_Some del d dist_l_r dist_r
       show ?thesis
-	by fastsimp
+        by fastsimp
     qed
   next
     case None
@@ -212,17 +212,17 @@
       case (Some r')
       with Node.hyps (2) [OF Some dist_r]
       have dist_r': "all_distinct r'"
-	by simp
+        by simp
       from delete_Some_set_of [OF Some]
       have "set_of r' \<subseteq> set_of r".
       with Some dist_r' x_l_None del dist_l d dist_l_r
       show ?thesis
-	by fastsimp
+        by fastsimp
     next
       case None
       with x_l_None del dist_l dist_r d dist_l_r
       show ?thesis
-	by (fastsimp split: split_if_asm)
+        by (fastsimp split: split_if_asm)
     qed
   qed
 qed
@@ -255,17 +255,17 @@
       case (Some r')
       from Node.hyps (2) [OF Some]
       obtain x_r: "x \<in> set_of r" "x \<notin> set_of r'"
-	by simp
+        by simp
       from x_r x_l Some x_l_Some del 
       show ?thesis
-	by (clarsimp split: split_if_asm)
+        by (clarsimp split: split_if_asm)
     next
       case None
       then have "x \<notin> set_of r"
-	by (simp add: delete_None_set_of_conv)
+        by (simp add: delete_None_set_of_conv)
       with x_l None x_l_Some del
       show ?thesis
-	by (clarsimp split: split_if_asm)
+        by (clarsimp split: split_if_asm)
     qed
   next
     case None
@@ -277,17 +277,17 @@
       case (Some r')
       from Node.hyps (2) [OF Some]
       obtain x_r: "x \<in> set_of r" "x \<notin> set_of r'"
-	by simp
+        by simp
       from x_r x_notin_l Some x_l_None del 
       show ?thesis
-	by (clarsimp split: split_if_asm)
+        by (clarsimp split: split_if_asm)
     next
       case None
       then have "x \<notin> set_of r"
-	by (simp add: delete_None_set_of_conv)
+        by (simp add: delete_None_set_of_conv)
       with None x_l_None x_notin_l del
       show ?thesis
-	by (clarsimp split: split_if_asm)
+        by (clarsimp split: split_if_asm)
     qed
   qed
 qed
@@ -324,23 +324,23 @@
       have t2''_t2': "set_of t\<^isub>2'' \<subseteq> set_of t\<^isub>2'" .
       show ?thesis
       proof (cases "subtract r t\<^isub>2''")
-	case (Some t\<^isub>2''')
-	from Node.hyps (2) [OF Some ] 
-	have "set_of t\<^isub>2''' \<subseteq> set_of t\<^isub>2''" .
-	with Some sub_l_Some del_x_Some sub t2''_t2' t2'_t2
-	show ?thesis
-	  by simp
+        case (Some t\<^isub>2''')
+        from Node.hyps (2) [OF Some ] 
+        have "set_of t\<^isub>2''' \<subseteq> set_of t\<^isub>2''" .
+        with Some sub_l_Some del_x_Some sub t2''_t2' t2'_t2
+        show ?thesis
+          by simp
       next
-	case None
-	with del_x_Some sub_l_Some sub
-	show ?thesis
-	  by simp
+        case None
+        with del_x_Some sub_l_Some sub
+        show ?thesis
+          by simp
       qed
     next
       case None
       with del_x_Some sub 
       show ?thesis
-	by simp
+        by simp
     qed
   next
     case None
@@ -374,23 +374,23 @@
       have t2''_t2': "set_of t\<^isub>2'' \<subseteq> set_of t\<^isub>2'" .
       show ?thesis
       proof (cases "subtract r t\<^isub>2''")
-	case (Some t\<^isub>2''')
-	from Node.hyps (2) [OF Some ] 
-	have r_t\<^isub>2'': "set_of r \<subseteq> set_of t\<^isub>2''" .
-	from Some sub_l_Some del_x_Some sub r_t\<^isub>2'' l_t2' t2'_t2 t2''_t2' x_t2
-	show ?thesis
-	  by auto
+        case (Some t\<^isub>2''')
+        from Node.hyps (2) [OF Some ] 
+        have r_t\<^isub>2'': "set_of r \<subseteq> set_of t\<^isub>2''" .
+        from Some sub_l_Some del_x_Some sub r_t\<^isub>2'' l_t2' t2'_t2 t2''_t2' x_t2
+        show ?thesis
+          by auto
       next
-	case None
-	with del_x_Some sub_l_Some sub
-	show ?thesis
-	  by simp
+        case None
+        with del_x_Some sub_l_Some sub
+        show ?thesis
+          by simp
       qed
     next
       case None
       with del_x_Some sub 
       show ?thesis
-	by simp
+        by simp
     qed
   next
     case None
@@ -420,24 +420,24 @@
       have dist_t2'': "all_distinct t\<^isub>2''" .
       show ?thesis
       proof (cases "subtract r t\<^isub>2''")
-	case (Some t\<^isub>2''')
-	from Node.hyps (2) [OF Some dist_t2''] 
-	have dist_t2''': "all_distinct t\<^isub>2'''" .
-	from Some sub_l_Some del_x_Some sub 
+        case (Some t\<^isub>2''')
+        from Node.hyps (2) [OF Some dist_t2''] 
+        have dist_t2''': "all_distinct t\<^isub>2'''" .
+        from Some sub_l_Some del_x_Some sub 
              dist_t2'''
-	show ?thesis
-	  by simp
+        show ?thesis
+          by simp
       next
-	case None
-	with del_x_Some sub_l_Some sub
-	show ?thesis
-	  by simp
+        case None
+        with del_x_Some sub_l_Some sub
+        show ?thesis
+          by simp
       qed
     next
       case None
       with del_x_Some sub 
       show ?thesis
-	by simp
+        by simp
     qed
   next
     case None
@@ -472,34 +472,34 @@
       have t2''_t2': "set_of t\<^isub>2'' \<subseteq> set_of t\<^isub>2'" .
       show ?thesis
       proof (cases "subtract r t\<^isub>2''")
-	case (Some t\<^isub>2''')
-	from Node.hyps (2) [OF Some] 
-	have dist_r_t2''': "set_of r \<inter> set_of t\<^isub>2''' = {}" .
-	from subtract_Some_set_of_res [OF Some]
-	have t2'''_t2'': "set_of t\<^isub>2''' \<subseteq> set_of t\<^isub>2''".
-	
-	from Some sub_l_Some del_x_Some sub t2'''_t2'' dist_l_t2'' dist_r_t2'''
+        case (Some t\<^isub>2''')
+        from Node.hyps (2) [OF Some] 
+        have dist_r_t2''': "set_of r \<inter> set_of t\<^isub>2''' = {}" .
+        from subtract_Some_set_of_res [OF Some]
+        have t2'''_t2'': "set_of t\<^isub>2''' \<subseteq> set_of t\<^isub>2''".
+        
+        from Some sub_l_Some del_x_Some sub t2'''_t2'' dist_l_t2'' dist_r_t2'''
              t2''_t2' t2'_t2 x_not_t2'
-	show ?thesis
-	  by auto
+        show ?thesis
+          by auto
       next
-	case None
-	with del_x_Some sub_l_Some sub
-	show ?thesis
-	  by simp
+        case None
+        with del_x_Some sub_l_Some sub
+        show ?thesis
+          by simp
       qed
     next
       case None
       with del_x_Some sub 
       show ?thesis
-	by simp
+        by simp
     qed
   next
     case None
     with sub show ?thesis by simp
   qed
 qed
-	
+        
 lemma subtract_Some_all_distinct:
   "\<And>t\<^isub>2 t. \<lbrakk>subtract t\<^isub>1 t\<^isub>2 = Some t; all_distinct t\<^isub>2\<rbrakk> \<Longrightarrow> all_distinct t\<^isub>1"
 proof (induct t\<^isub>1)
@@ -536,29 +536,29 @@
       have dist_l_t2'': "set_of l \<inter> set_of t\<^isub>2'' = {}".
       show ?thesis
       proof (cases "subtract r t\<^isub>2''")
-	case (Some t\<^isub>2''')
-	from Node.hyps (2) [OF Some dist_t2''] 
-	have dist_r: "all_distinct r" .
-	from subtract_Some_set_of [OF Some]
-	have r_t2'': "set_of r \<subseteq> set_of t\<^isub>2''" .
-	from subtract_Some_dist_res [OF Some]
-	have dist_r_t2''': "set_of r \<inter> set_of t\<^isub>2''' = {}".
+        case (Some t\<^isub>2''')
+        from Node.hyps (2) [OF Some dist_t2''] 
+        have dist_r: "all_distinct r" .
+        from subtract_Some_set_of [OF Some]
+        have r_t2'': "set_of r \<subseteq> set_of t\<^isub>2''" .
+        from subtract_Some_dist_res [OF Some]
+        have dist_r_t2''': "set_of r \<inter> set_of t\<^isub>2''' = {}".
 
-	from dist_l dist_r Some sub_l_Some del_x_Some r_t2'' l_t2' x_t2 x_not_t2' 
-	     t2''_t2' dist_l_t2'' dist_r_t2'''
-	show ?thesis
-	  by auto
+        from dist_l dist_r Some sub_l_Some del_x_Some r_t2'' l_t2' x_t2 x_not_t2' 
+             t2''_t2' dist_l_t2'' dist_r_t2'''
+        show ?thesis
+          by auto
       next
-	case None
-	with del_x_Some sub_l_Some sub
-	show ?thesis
-	  by simp
+        case None
+        with del_x_Some sub_l_Some sub
+        show ?thesis
+          by simp
       qed
     next
       case None
       with del_x_Some sub 
       show ?thesis
-	by simp
+        by simp
     qed
   next
     case None
--- a/src/HOL/Statespace/state_fun.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Statespace/state_fun.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -238,7 +238,7 @@
                       end
                | mk_updterm _ t = init_seed t;
 
-	     val ctxt = Simplifier.the_context ss |>
+             val ctxt = Simplifier.the_context ss |>
                         Config.map MetaSimplifier.simp_depth_limit (K 100);
              val ss1 = Simplifier.context ctxt ss';
              val ss2 = Simplifier.context ctxt 
@@ -285,10 +285,10 @@
                let val (_::nT::_) = binder_types lT;
                          (*  ('v => 'a) => 'n => ('n => 'v) => 'a *)
                    val x' = if not (loose_bvar1 (x,0))
-			    then Bound 1
+                            then Bound 1
                             else raise TERM ("",[x]);
                    val n' = if not (loose_bvar1 (n,0))
-			    then Bound 2
+                            then Bound 2
                             else raise TERM ("",[n]);
                    val sel' = lo $ d $ n' $ s;
                   in (Const ("op =",Teq)$sel'$x',hd (binder_types Teq),nT,swap) end;
@@ -316,7 +316,7 @@
                   val prop = list_all ([("n",nT),("x",eT)],
                               Logic.mk_equals (Const ("Ex",Tex)$Abs(s,T,eq),
                                                HOLogic.true_const));
-                  val thm = standard (prove prop);
+                  val thm = Drule.standard (prove prop);
                   val thm' = if swap then swap_ex_eq OF [thm] else thm
              in SOME thm' end
              handle TERM _ => NONE)
@@ -372,7 +372,7 @@
   let
      val (lookup_ss,ex_lookup_ss,simprocs_active) = StateFunData.get ctxt;
      val (lookup_ss', ex_lookup_ss') = 
-	   (case (concl_of thm) of
+           (case (concl_of thm) of
             (_$((Const ("Ex",_)$_))) => (lookup_ss, ex_lookup_ss addsimps [thm])
             | _ => (lookup_ss addsimps [thm], ex_lookup_ss))
      fun activate_simprocs ctxt =
--- a/src/HOL/Statespace/state_space.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Statespace/state_space.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -288,7 +288,7 @@
        | _ =>
         let
           val {declinfo,distinctthm=tt,silent} = (NameSpaceData.get ctxt);
-	  val all_names = comps_of_thm thm;
+          val all_names = comps_of_thm thm;
           fun upd name tt =
                (case (Symtab.lookup tt name) of
                  SOME dthm => if sorted_subset (op =) (comps_of_thm dthm) all_names
@@ -365,7 +365,7 @@
       | rename (NONE::rs)  (x::xs) = x::rename rs xs
       | rename (SOME r::rs) ((x,T)::xs) = (r,T)::rename rs xs;
     val {args,parents,components,...} =
-	      the (Symtab.lookup (StateSpaceData.get ctxt) pname);
+              the (Symtab.lookup (StateSpaceData.get ctxt) pname);
     val inst = map fst args ~~ Ts;
     val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst);
     val parent_comps =
@@ -450,7 +450,7 @@
          NONE => []
        | SOME s =>
           let
-	    val fx = Element.Fixes [(Binding.name s,SOME (string_of_typ stateT),NoSyn)];
+            val fx = Element.Fixes [(Binding.name s,SOME (string_of_typ stateT),NoSyn)];
             val cs = Element.Constrains
                        (map (fn (n,T) =>  (n,string_of_typ T))
                          ((map (fn (n,_) => (n,nameT)) all_comps) @
@@ -610,7 +610,7 @@
            Syntax.const "StateFun.lookup"$Syntax.free (project_name T)$Syntax.free n$s
        | NONE =>
            if get_silent (Context.Proof ctxt)
-	   then Syntax.const "StateFun.lookup" $ Syntax.const "undefined" $ Syntax.free n $ s
+           then Syntax.const "StateFun.lookup" $ Syntax.const "undefined" $ Syntax.free n $ s
            else raise TERM ("StateSpace.gen_lookup_tr: component " ^ n ^ " not defined",[]));
 
 fun lookup_tr ctxt [s,Free (n,_)] = gen_lookup_tr ctxt s n;
--- a/src/HOL/TLA/Memory/RPCParameters.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/TLA/Memory/RPCParameters.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -40,11 +40,11 @@
 
 defs
   IsLegalRcvArg_def: "IsLegalRcvArg ra ==
-		         case ra of (memcall m) => True
-		                  | (othercall v) => False"
+                         case ra of (memcall m) => True
+                                  | (othercall v) => False"
   RPCRelayArg_def:   "RPCRelayArg ra ==
-		         case ra of (memcall m) => m
-		                  | (othercall v) => undefined"
+                         case ra of (memcall m) => m
+                                  | (othercall v) => undefined"
 
 lemmas [simp] = RFNoMemVal NotAResultNotRF OKNotRF BANotRF
   NotAResultNotRF [symmetric] OKNotRF [symmetric] BANotRF [symmetric]
--- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -244,8 +244,10 @@
         val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
         val rep_const = cterm_of thy
           (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT));
-        val cong' = standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
-        val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
+        val cong' =
+          Drule.standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
+        val dist =
+          Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
         val (thy', defs', eqns', _) = Library.foldl ((make_constr_def tname T) (length constrs))
           ((Sign.add_path tname thy, defs, [], 1), constrs ~~ constr_syntax)
       in
@@ -520,7 +522,7 @@
               let
                 val dist_thm = SkipProof.prove_global thy5 [] [] t (fn _ =>
                   EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
-              in dist_thm :: standard (dist_thm RS not_sym) :: prove ts end;
+              in dist_thm :: Drule.standard (dist_thm RS not_sym) :: prove ts end;
       in prove ts end;
 
     val distinct_thms = map2 (prove_distinct_thms)
--- a/src/HOL/Tools/Function/size.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/Function/size.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -198,7 +198,7 @@
 
     fun prove_size_eqs p size_fns size_ofp simpset =
       maps (fn (((_, (_, _, constrs)), size_const), T) =>
-        map (fn constr => standard (SkipProof.prove ctxt [] []
+        map (fn constr => Drule.standard (SkipProof.prove ctxt [] []
           (gen_mk_size_eq p (AList.lookup op = (new_type_names ~~ size_fns))
              size_ofp size_const T constr)
           (fn _ => simp_tac simpset 1))) constrs)
--- a/src/HOL/Tools/TFL/post.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/TFL/post.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -7,8 +7,6 @@
 
 signature TFL =
 sig
-  val tgoalw: theory -> thm list -> thm list -> thm list
-  val tgoal: theory -> thm list -> thm list
   val define_i: bool -> theory -> claset -> simpset -> thm list -> thm list -> xstring ->
     term -> term list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list}
   val define: bool -> theory -> claset -> simpset -> thm list -> thm list -> xstring ->
@@ -33,19 +31,6 @@
       (List.foldr (fn (th,A) => gen_union (op aconv) (prems_of th, A)) [] rules);
 
 (*---------------------------------------------------------------------------
- * Finds the termination conditions in (highly massaged) definition and
- * puts them into a goalstack.
- *--------------------------------------------------------------------------*)
-fun tgoalw thy defs rules =
-  case termination_goals rules of
-      [] => error "tgoalw: no termination conditions to prove"
-    | L  => OldGoals.goalw_cterm defs
-              (Thm.cterm_of thy
-                        (HOLogic.mk_Trueprop(USyntax.list_mk_conj L)));
-
-fun tgoal thy = tgoalw thy [];
-
-(*---------------------------------------------------------------------------
  * Three postprocessors are applied to the definition.  It
  * attempts to prove wellfoundedness of the given relation, simplifies the
  * non-proved termination conditions, and finally attempts to prove the
@@ -144,7 +129,7 @@
 
 (*lcp: put a theorem into Isabelle form, using meta-level connectives*)
 val meta_outer =
-  curry_rule o standard o
+  curry_rule o Drule.standard o
   rule_by_tactic (REPEAT (FIRSTGOAL (resolve_tac [allI, impI, conjI] ORELSE' etac conjE)));
 
 (*Strip off the outer !P*)
@@ -166,7 +151,7 @@
                {f = f, R = R, rules = rules,
                 full_pats_TCs = full_pats_TCs,
                 TCs = TCs}
-       val rules' = map (standard o ObjectLogic.rulify_no_asm)
+       val rules' = map (Drule.standard o ObjectLogic.rulify_no_asm)
                         (R.CONJUNCTS rules)
          in  {induct = meta_outer (ObjectLogic.rulify_no_asm (induction RS spec')),
         rules = ListPair.zip(rules', rows),
@@ -195,7 +180,7 @@
     | solve_eq (th, [a], i) = [(a, i)]
     | solve_eq (th, splitths as (_ :: _), i) = 
       (writeln "Proving unsplit equation...";
-      [((standard o ObjectLogic.rulify_no_asm)
+      [((Drule.standard o ObjectLogic.rulify_no_asm)
           (CaseSplit.splitto splitths th), i)])
       (* if there's an error, pretend nothing happened with this definition 
          We should probably print something out so that the user knows...? *)
@@ -252,7 +237,7 @@
  in (theory,
      (*return the conjoined induction rule and recursion equations,
        with assumptions remaining to discharge*)
-     standard (induction RS (rules RS conjI)))
+     Drule.standard (induction RS (rules RS conjI)))
  end
 
 fun defer thy congs fid seqs =
--- a/src/HOL/Tools/arith_data.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/arith_data.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -77,7 +77,7 @@
 fun prove_conv tacs ctxt (_: thm list) = prove_conv_nohyps tacs ctxt;
 
 fun prove_conv2 expand_tac norm_tac ss tu = (*FIXME avoid standard*)
-  mk_meta_eq (standard (Goal.prove (Simplifier.the_context ss) [] []
+  mk_meta_eq (Drule.standard (Goal.prove (Simplifier.the_context ss) [] []
       (HOLogic.mk_Trueprop (HOLogic.mk_eq tu))
     (K (EVERY [expand_tac, norm_tac ss]))));
 
--- a/src/HOL/Tools/choice_specification.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/choice_specification.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -75,7 +75,7 @@
 fun add_specification axiomatic cos arg =
     arg |> apsnd Thm.freezeT
         |> proc_exprop axiomatic cos
-        |> apsnd standard
+        |> apsnd Drule.standard
 
 
 (* Collect all intances of constants in term *)
@@ -188,7 +188,7 @@
                     in
                         args |> apsnd (remove_alls frees)
                              |> apsnd undo_imps
-                             |> apsnd standard
+                             |> apsnd Drule.standard
                              |> Thm.theory_attributes (map (Attrib.attribute thy) atts)
                              |> add_final
                              |> Library.swap
--- a/src/HOL/Tools/cnf_funcs.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/cnf_funcs.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -2,33 +2,34 @@
     Author:     Alwen Tiu, QSL Team, LORIA (http://qsl.loria.fr)
     Author:     Tjark Weber, TU Muenchen
 
-  FIXME: major overlaps with the code in meson.ML
+FIXME: major overlaps with the code in meson.ML
+
+Functions and tactics to transform a formula into Conjunctive Normal
+Form (CNF).
 
-  Description:
-  This file contains functions and tactics to transform a formula into
-  Conjunctive Normal Form (CNF).
-  A formula in CNF is of the following form:
+A formula in CNF is of the following form:
 
-      (x11 | x12 | ... | x1n) & ... & (xm1 | xm2 | ... | xmk)
-      False
-      True
+    (x11 | x12 | ... | x1n) & ... & (xm1 | xm2 | ... | xmk)
+    False
+    True
 
-  where each xij is a literal (a positive or negative atomic Boolean term),
-  i.e. the formula is a conjunction of disjunctions of literals, or
-  "False", or "True".
+where each xij is a literal (a positive or negative atomic Boolean
+term), i.e. the formula is a conjunction of disjunctions of literals,
+or "False", or "True".
 
-  A (non-empty) disjunction of literals is referred to as "clause".
+A (non-empty) disjunction of literals is referred to as "clause".
 
-  For the purpose of SAT proof reconstruction, we also make use of another
-  representation of clauses, which we call the "raw clauses".
-  Raw clauses are of the form
+For the purpose of SAT proof reconstruction, we also make use of
+another representation of clauses, which we call the "raw clauses".
+Raw clauses are of the form
+
+    [..., x1', x2', ..., xn'] |- False ,
 
-      [..., x1', x2', ..., xn'] |- False ,
+where each xi is a literal, and each xi' is the negation normal form
+of ~xi.
 
-  where each xi is a literal, and each xi' is the negation normal form of ~xi.
-
-  Literals are successively removed from the hyps of raw clauses by resolution
-  during SAT proof reconstruction.
+Literals are successively removed from the hyps of raw clauses by
+resolution during SAT proof reconstruction.
 *)
 
 signature CNF =
--- a/src/HOL/Tools/inductive_codegen.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/inductive_codegen.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -477,7 +477,7 @@
       (fn NONE => "X" | SOME k' => string_of_int k')
         (ks @ [SOME k]))) arities));
 
-fun prep_intrs intrs = map (rename_term o #prop o rep_thm o standard) intrs;
+fun prep_intrs intrs = map (rename_term o #prop o rep_thm o Drule.standard) intrs;
 
 fun constrain cs [] = []
   | constrain cs ((s, xs) :: ys) = (s, case AList.lookup (op =) cs (s : string) of
--- a/src/HOL/Tools/meson.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/meson.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -266,21 +266,21 @@
   fun signed_nclauses b (Const("Trueprop",_) $ t) = signed_nclauses b t
     | signed_nclauses b (Const("Not",_) $ t) = signed_nclauses (not b) t
     | signed_nclauses b (Const("op &",_) $ t $ u) =
-	if b then sum (signed_nclauses b t) (signed_nclauses b u)
-	     else prod (signed_nclauses b t) (signed_nclauses b u)
+        if b then sum (signed_nclauses b t) (signed_nclauses b u)
+             else prod (signed_nclauses b t) (signed_nclauses b u)
     | signed_nclauses b (Const("op |",_) $ t $ u) =
-	if b then prod (signed_nclauses b t) (signed_nclauses b u)
-	     else sum (signed_nclauses b t) (signed_nclauses b u)
+        if b then prod (signed_nclauses b t) (signed_nclauses b u)
+             else sum (signed_nclauses b t) (signed_nclauses b u)
     | signed_nclauses b (Const("op -->",_) $ t $ u) =
-	if b then prod (signed_nclauses (not b) t) (signed_nclauses b u)
-	     else sum (signed_nclauses (not b) t) (signed_nclauses b u)
+        if b then prod (signed_nclauses (not b) t) (signed_nclauses b u)
+             else sum (signed_nclauses (not b) t) (signed_nclauses b u)
     | signed_nclauses b (Const("op =", Type ("fun", [T, _])) $ t $ u) =
-	if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*)
-	    if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u))
-			  (prod (signed_nclauses (not b) u) (signed_nclauses b t))
-		 else sum (prod (signed_nclauses b t) (signed_nclauses b u))
-			  (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u))
-	else 1
+        if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*)
+            if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u))
+                          (prod (signed_nclauses (not b) u) (signed_nclauses b t))
+                 else sum (prod (signed_nclauses b t) (signed_nclauses b u))
+                          (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u))
+        else 1
     | signed_nclauses b (Const("Ex", _) $ Abs (_,_,t)) = signed_nclauses b t
     | signed_nclauses b (Const("All",_) $ Abs (_,_,t)) = signed_nclauses b t
     | signed_nclauses _ _ = 1; (* literal *)
@@ -347,9 +347,9 @@
           | _ => nodups ctxt th :: ths  (*no work to do*)
       and cnf_nil th = cnf_aux (th,[])
       val cls = 
-	    if too_many_clauses (SOME ctxt) (concl_of th)
-	    then (trace_msg (fn () => "cnf is ignoring: " ^ Display.string_of_thm ctxt th); ths)
-	    else cnf_aux (th,ths)
+            if too_many_clauses (SOME ctxt) (concl_of th)
+            then (trace_msg (fn () => "cnf is ignoring: " ^ Display.string_of_thm ctxt th); ths)
+            else cnf_aux (th,ths)
   in  (cls, !ctxtr)  end;
 
 fun make_cnf skoths th ctxt = cnf skoths ctxt (th, []);
--- a/src/HOL/Tools/numeral_simprocs.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/numeral_simprocs.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -210,7 +210,7 @@
 
 (*push the unary minus down: - x * y = x * - y *)
 val minus_mult_eq_1_to_2 =
-    [@{thm mult_minus_left}, @{thm minus_mult_right}] MRS trans |> standard;
+    [@{thm mult_minus_left}, @{thm minus_mult_right}] MRS trans |> Drule.standard;
 
 (*to extract again any uncancelled minuses*)
 val minus_from_mult_simps =
--- a/src/HOL/Tools/polyhash.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/polyhash.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -106,10 +106,10 @@
 
 datatype ('key, 'data) hash_table = 
     HT of {hashVal   : 'key -> int,
-	   sameKey   : 'key * 'key -> bool,
-	   not_found : exn,
-	   table     : ('key, 'data) bucket_t Array.array Unsynchronized.ref,
-	   n_items   : int Unsynchronized.ref}
+           sameKey   : 'key * 'key -> bool,
+           not_found : exn,
+           table     : ('key, 'data) bucket_t Array.array Unsynchronized.ref,
+           n_items   : int Unsynchronized.ref}
 
 local
 (*
@@ -123,8 +123,8 @@
 
   (* find smallest power of 2 (>= 32) that is >= n *)
     fun roundUp n = 
-	let fun f i = if (i >= n) then i else f (lshift_ i 1)
-	in f 32 end
+        let fun f i = if (i >= n) then i else f (lshift_ i 1)
+        in f 32 end
 end;
 
   (* Create a new table; the int is a size hint and the exception
@@ -132,255 +132,255 @@
    *)
     fun mkTable (hashVal, sameKey) (sizeHint, notFound) = HT{
             hashVal=hashVal,
-	    sameKey=sameKey,
-	    not_found = notFound,
-	    table = Unsynchronized.ref (Array.array(roundUp sizeHint, NIL)),
-	    n_items = Unsynchronized.ref 0
-	  };
+            sameKey=sameKey,
+            not_found = notFound,
+            table = Unsynchronized.ref (Array.array(roundUp sizeHint, NIL)),
+            n_items = Unsynchronized.ref 0
+          };
 
   (* conditionally grow a table *)
     fun growTable (HT{table, n_items, ...}) = let
-	    val arr = !table
-	    val sz = Array.length arr
-	    in
-	      if (!n_items >= sz)
-		then let
-		  val newSz = sz+sz
-		  val newArr = Array.array (newSz, NIL)
-		  fun copy NIL = ()
-		    | copy (B(h, key, v, rest)) = let
-			val indx = index (h, newSz)
-			in
-			  Array.update (newArr, indx,
-			    B(h, key, v, Array.sub(newArr, indx)));
-			  copy rest
-			end
-		  fun bucket n = (copy (Array.sub(arr, n)); bucket (n+1))
-		  in
-		    (bucket 0) handle _ => ();  (* FIXME avoid handle _ *)
-		    table := newArr
-		  end
-		else ()
-	    end (* growTable *);
+            val arr = !table
+            val sz = Array.length arr
+            in
+              if (!n_items >= sz)
+                then let
+                  val newSz = sz+sz
+                  val newArr = Array.array (newSz, NIL)
+                  fun copy NIL = ()
+                    | copy (B(h, key, v, rest)) = let
+                        val indx = index (h, newSz)
+                        in
+                          Array.update (newArr, indx,
+                            B(h, key, v, Array.sub(newArr, indx)));
+                          copy rest
+                        end
+                  fun bucket n = (copy (Array.sub(arr, n)); bucket (n+1))
+                  in
+                    (bucket 0) handle _ => ();  (* FIXME avoid handle _ *)
+                    table := newArr
+                  end
+                else ()
+            end (* growTable *);
 
   (* Insert an item.  If the key already has an item associated with it,
    * then the old item is discarded.
    *)
     fun insert (tbl as HT{hashVal, sameKey, table, n_items, ...}) (key, item) =
-	let
-	  val arr = !table
-	  val sz = Array.length arr
-	  val hash = hashVal key
-	  val indx = index (hash, sz)
-	  fun look NIL = (
-		Array.update(arr, indx, B(hash, key, item, Array.sub(arr, indx)));
-		Unsynchronized.inc n_items;
-		growTable tbl;
-		NIL)
-	    | look (B(h, k, v, r)) = if ((hash = h) andalso sameKey(key, k))
-		then B(hash, key, item, r)
-		else (case (look r)
-		   of NIL => NIL
-		    | rest => B(h, k, v, rest)
-		  (* end case *))
-	  in
-	    case (look (Array.sub (arr, indx)))
-	     of NIL => ()
-	      | b => Array.update(arr, indx, b)
-	  end;
+        let
+          val arr = !table
+          val sz = Array.length arr
+          val hash = hashVal key
+          val indx = index (hash, sz)
+          fun look NIL = (
+                Array.update(arr, indx, B(hash, key, item, Array.sub(arr, indx)));
+                Unsynchronized.inc n_items;
+                growTable tbl;
+                NIL)
+            | look (B(h, k, v, r)) = if ((hash = h) andalso sameKey(key, k))
+                then B(hash, key, item, r)
+                else (case (look r)
+                   of NIL => NIL
+                    | rest => B(h, k, v, rest)
+                  (* end case *))
+          in
+            case (look (Array.sub (arr, indx)))
+             of NIL => ()
+              | b => Array.update(arr, indx, b)
+          end;
 
   (* Insert an item if not there already; if it is there already, 
      then return the old data value and leave the table unmodified..
    *)
     fun peekInsert (tbl as HT{hashVal, sameKey, table, n_items, ...}) (key, item) =
-	let val arr = !table
-	    val sz = Array.length arr
-	    val hash = hashVal key
-	    val indx = index (hash, sz)
-	    fun look NIL = 
-		(Array.update(arr, indx, B(hash, key, item, 
-					   Array.sub(arr, indx)));
-		 Unsynchronized.inc n_items;
-		 growTable tbl;
-		 NONE)
-	      | look (B(h, k, v, r)) = 
-		if hash = h andalso sameKey(key, k) then SOME v
-		else look r
-	in
-	    look (Array.sub (arr, indx))
-	end;
+        let val arr = !table
+            val sz = Array.length arr
+            val hash = hashVal key
+            val indx = index (hash, sz)
+            fun look NIL = 
+                (Array.update(arr, indx, B(hash, key, item, 
+                                           Array.sub(arr, indx)));
+                 Unsynchronized.inc n_items;
+                 growTable tbl;
+                 NONE)
+              | look (B(h, k, v, r)) = 
+                if hash = h andalso sameKey(key, k) then SOME v
+                else look r
+        in
+            look (Array.sub (arr, indx))
+        end;
 
   (* find an item, the table's exception is raised if the item doesn't exist *)
     fun find (HT{hashVal, sameKey, table, not_found, ...}) key = let
-	  val arr = !table
-	  val sz = Array.length arr
-	  val hash = hashVal key
-	  val indx = index (hash, sz)
-	  fun look NIL = raise not_found
-	    | look (B(h, k, v, r)) = if ((hash = h) andalso sameKey(key, k))
-		then v
-		else look r
-	  in
-	    look (Array.sub (arr, indx))
-	  end;
+          val arr = !table
+          val sz = Array.length arr
+          val hash = hashVal key
+          val indx = index (hash, sz)
+          fun look NIL = raise not_found
+            | look (B(h, k, v, r)) = if ((hash = h) andalso sameKey(key, k))
+                then v
+                else look r
+          in
+            look (Array.sub (arr, indx))
+          end;
 
   (* look for an item, return NONE if the item doesn't exist *)
     fun peek (HT{hashVal, sameKey, table, ...}) key = let
-	  val arr = !table
-	  val sz = Array.length arr
-	  val hash = hashVal key
-	  val indx = index (hash, sz)
-	  fun look NIL = NONE
-	    | look (B(h, k, v, r)) = if ((hash = h) andalso sameKey(key, k))
-		then SOME v
-		else look r
-	  in
-	    look (Array.sub (arr, indx))
-	  end;
+          val arr = !table
+          val sz = Array.length arr
+          val hash = hashVal key
+          val indx = index (hash, sz)
+          fun look NIL = NONE
+            | look (B(h, k, v, r)) = if ((hash = h) andalso sameKey(key, k))
+                then SOME v
+                else look r
+          in
+            look (Array.sub (arr, indx))
+          end;
 
   (* Remove an item.  The table's exception is raised if
    * the item doesn't exist.
    *)
     fun remove (HT{hashVal, sameKey, not_found, table, n_items}) key = let
-	  val arr = !table
-	  val sz = Array.length arr
-	  val hash = hashVal key
-	  val indx = index (hash, sz)
-	  fun look NIL = raise not_found
-	    | look (B(h, k, v, r)) = if ((hash = h) andalso sameKey(key, k))
-		then (v, r)
-		else let val (item, r') = look r in (item, B(h, k, v, r')) end
-	  val (item, bucket) = look (Array.sub (arr, indx))
-	  in
-	    Array.update (arr, indx, bucket);
-	    n_items := !n_items - 1;
-	    item
-	  end (* remove *);
+          val arr = !table
+          val sz = Array.length arr
+          val hash = hashVal key
+          val indx = index (hash, sz)
+          fun look NIL = raise not_found
+            | look (B(h, k, v, r)) = if ((hash = h) andalso sameKey(key, k))
+                then (v, r)
+                else let val (item, r') = look r in (item, B(h, k, v, r')) end
+          val (item, bucket) = look (Array.sub (arr, indx))
+          in
+            Array.update (arr, indx, bucket);
+            n_items := !n_items - 1;
+            item
+          end (* remove *);
 
   (* Return the number of items in the table *)
    fun numItems (HT{n_items, ...}) = !n_items
 
   (* return a list of the items in the table *)
     fun listItems (HT{table = Unsynchronized.ref arr, n_items, ...}) = let
-	  fun f (_, l, 0) = l
-	    | f (~1, l, _) = l
-	    | f (i, l, n) = let
-		fun g (NIL, l, n) = f (i-1, l, n)
-		  | g (B(_, k, v, r), l, n) = g(r, (k, v)::l, n-1)
-		in
-		  g (Array.sub(arr, i), l, n)
-		end
-	  in
-	    f ((Array.length arr) - 1, [], !n_items)
-	  end (* listItems *);
+          fun f (_, l, 0) = l
+            | f (~1, l, _) = l
+            | f (i, l, n) = let
+                fun g (NIL, l, n) = f (i-1, l, n)
+                  | g (B(_, k, v, r), l, n) = g(r, (k, v)::l, n-1)
+                in
+                  g (Array.sub(arr, i), l, n)
+                end
+          in
+            f ((Array.length arr) - 1, [], !n_items)
+          end (* listItems *);
 
   (* Apply a function to the entries of the table *)
     fun apply f (HT{table, ...}) = let
-	  fun appF NIL = ()
-	    | appF (B(_, key, item, rest)) = (
-		f (key, item);
-		appF rest)
-	  val arr = !table
-	  val sz = Array.length arr
-	  fun appToTbl i = if (i < sz)
-		then (appF (Array.sub (arr, i)); appToTbl(i+1))
-		else ()
-	  in
-	    appToTbl 0
-	  end (* apply *);
+          fun appF NIL = ()
+            | appF (B(_, key, item, rest)) = (
+                f (key, item);
+                appF rest)
+          val arr = !table
+          val sz = Array.length arr
+          fun appToTbl i = if (i < sz)
+                then (appF (Array.sub (arr, i)); appToTbl(i+1))
+                else ()
+          in
+            appToTbl 0
+          end (* apply *);
 
   (* Map a table to a new table that has the same keys and exception *)
     fun map f (HT{hashVal, sameKey, table, n_items, not_found}) = let
-	  fun mapF NIL = NIL
-	    | mapF (B(hash, key, item, rest)) =
-		B(hash, key, f (key, item), mapF rest)
-	  val arr = !table
-	  val sz = Array.length arr
-	  val newArr = Array.array (sz, NIL)
-	  fun mapTbl i = if (i < sz)
-		then (
-		  Array.update(newArr, i, mapF (Array.sub(arr, i)));
-		  mapTbl (i+1))
-		else ()
-	  in
-	    mapTbl 0;
-	    HT{hashVal=hashVal,
-	       sameKey=sameKey,
-	       table = Unsynchronized.ref newArr, 
-	       n_items = Unsynchronized.ref(!n_items), 
-	       not_found = not_found}
-	  end (* transform *);
+          fun mapF NIL = NIL
+            | mapF (B(hash, key, item, rest)) =
+                B(hash, key, f (key, item), mapF rest)
+          val arr = !table
+          val sz = Array.length arr
+          val newArr = Array.array (sz, NIL)
+          fun mapTbl i = if (i < sz)
+                then (
+                  Array.update(newArr, i, mapF (Array.sub(arr, i)));
+                  mapTbl (i+1))
+                else ()
+          in
+            mapTbl 0;
+            HT{hashVal=hashVal,
+               sameKey=sameKey,
+               table = Unsynchronized.ref newArr, 
+               n_items = Unsynchronized.ref(!n_items), 
+               not_found = not_found}
+          end (* transform *);
 
   (* remove any hash table items that do not satisfy the given
    * predicate.
    *)
     fun filter pred (HT{table, n_items, not_found, ...}) = let
-	  fun filterP NIL = NIL
-	    | filterP (B(hash, key, item, rest)) = if (pred(key, item))
-		then B(hash, key, item, filterP rest)
-		else filterP rest
-	  val arr = !table
-	  val sz = Array.length arr
-	  fun filterTbl i = if (i < sz)
-		then (
-		  Array.update (arr, i, filterP (Array.sub (arr, i)));
-		  filterTbl (i+1))
-		else ()
-	  in
-	    filterTbl 0
-	  end (* filter *);
+          fun filterP NIL = NIL
+            | filterP (B(hash, key, item, rest)) = if (pred(key, item))
+                then B(hash, key, item, filterP rest)
+                else filterP rest
+          val arr = !table
+          val sz = Array.length arr
+          fun filterTbl i = if (i < sz)
+                then (
+                  Array.update (arr, i, filterP (Array.sub (arr, i)));
+                  filterTbl (i+1))
+                else ()
+          in
+            filterTbl 0
+          end (* filter *);
 
   (* Map a table to a new table that has the same keys, exception,
      hash function, and equality function *)
 
     fun transform f (HT{hashVal, sameKey, table, n_items, not_found}) = let
-	  fun mapF NIL = NIL
-	    | mapF (B(hash, key, item, rest)) = B(hash, key, f item, mapF rest)
-	  val arr = !table
-	  val sz = Array.length arr
-	  val newArr = Array.array (sz, NIL)
-	  fun mapTbl i = if (i < sz)
-		then (
-		  Array.update(newArr, i, mapF (Array.sub(arr, i)));
-		  mapTbl (i+1))
-		else ()
-	  in
-	    mapTbl 0;
-	    HT{hashVal=hashVal, 
-	       sameKey=sameKey, 
-	       table = Unsynchronized.ref newArr, 
-	       n_items = Unsynchronized.ref(!n_items), 
-	       not_found = not_found}
-	  end (* transform *);
+          fun mapF NIL = NIL
+            | mapF (B(hash, key, item, rest)) = B(hash, key, f item, mapF rest)
+          val arr = !table
+          val sz = Array.length arr
+          val newArr = Array.array (sz, NIL)
+          fun mapTbl i = if (i < sz)
+                then (
+                  Array.update(newArr, i, mapF (Array.sub(arr, i)));
+                  mapTbl (i+1))
+                else ()
+          in
+            mapTbl 0;
+            HT{hashVal=hashVal, 
+               sameKey=sameKey, 
+               table = Unsynchronized.ref newArr, 
+               n_items = Unsynchronized.ref(!n_items), 
+               not_found = not_found}
+          end (* transform *);
 
   (* Create a copy of a hash table *)
     fun copy (HT{hashVal, sameKey, table, n_items, not_found}) = let
-	  val arr = !table
-	  val sz = Array.length arr
-	  val newArr = Array.array (sz, NIL)
-	  fun mapTbl i = (
-		Array.update (newArr, i, Array.sub(arr, i));
-		mapTbl (i+1))
-	  in
-	    (mapTbl 0) handle _ => ();  (* FIXME avoid handle _ *)
-	    HT{hashVal=hashVal, 
-	       sameKey=sameKey,
-	       table = Unsynchronized.ref newArr, 
-	       n_items = Unsynchronized.ref(!n_items), 
-	       not_found = not_found}
-	  end (* copy *);
+          val arr = !table
+          val sz = Array.length arr
+          val newArr = Array.array (sz, NIL)
+          fun mapTbl i = (
+                Array.update (newArr, i, Array.sub(arr, i));
+                mapTbl (i+1))
+          in
+            (mapTbl 0) handle _ => ();  (* FIXME avoid handle _ *)
+            HT{hashVal=hashVal, 
+               sameKey=sameKey,
+               table = Unsynchronized.ref newArr, 
+               n_items = Unsynchronized.ref(!n_items), 
+               not_found = not_found}
+          end (* copy *);
 
   (* returns a list of the sizes of the various buckets.  This is to
    * allow users to gauge the quality of their hashing function.
    *)
     fun bucketSizes (HT{table = Unsynchronized.ref arr, ...}) = let
-	  fun len (NIL, n) = n
-	    | len (B(_, _, _, r), n) = len(r, n+1)
-	  fun f (~1, l) = l
-	    | f (i, l) = f (i-1, len (Array.sub (arr, i), 0) :: l)
-	  in
-	    f ((Array.length arr)-1, [])
-	  end
+          fun len (NIL, n) = n
+            | len (B(_, _, _, r), n) = len(r, n+1)
+          fun f (~1, l) = l
+            | f (i, l) = f (i-1, len (Array.sub (arr, i), 0) :: l)
+          in
+            f ((Array.length arr)-1, [])
+          end
 
    (*Added by lcp.
       This is essentially the  described in Compilers:
--- a/src/HOL/Tools/record.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/record.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1045,7 +1045,7 @@
         we varify the proposition manually here.*)
   else
     let val prf = Goal.prove (ProofContext.init thy) [] asms prop tac
-    in if stndrd then standard prf else prf end;
+    in if stndrd then Drule.standard prf else prf end;
 
 fun quick_and_dirty_prf noopt opt () =
   if ! record_quick_and_dirty_sensitive andalso ! quick_and_dirty
@@ -1097,7 +1097,7 @@
           if is_sel_upd_pair thy acc upd
           then o_eq_dest
           else o_eq_id_dest;
-      in standard (othm RS dest) end;
+      in Drule.standard (othm RS dest) end;
   in map get_simp upd_funs end;
 
 fun get_updupd_simp thy defset intros_tac u u' comp =
@@ -1118,7 +1118,7 @@
             REPEAT_DETERM (intros_tac 1),
             TRY (simp_tac (HOL_ss addsimps [id_apply]) 1)]);
     val dest = if comp then o_eq_dest_lhs else o_eq_dest;
-  in standard (othm RS dest) end;
+  in Drule.standard (othm RS dest) end;
 
 fun get_updupd_simps thy term defset intros_tac =
   let
@@ -1312,7 +1312,8 @@
             val ss = get_sel_upd_defs thy;
             val uathm = get_upd_acc_cong_thm upd acc thy ss;
           in
-            [standard (uathm RS updacc_noopE), standard (uathm RS updacc_noop_compE)]
+           [Drule.standard (uathm RS updacc_noopE),
+            Drule.standard (uathm RS updacc_noop_compE)]
           end;
 
         (*If f is constant then (f o g) = f. we know that K_skeleton
@@ -1755,7 +1756,7 @@
       to prove other theorems. We haven't given names to the accessors
       f, g etc yet however, so we generate an ext structure with
       free variables as all arguments and allow the introduction tactic to
-      operate on it as far as it can. We then use standard to convert
+      operate on it as far as it can. We then use Drule.standard to convert
       the free variables into unifiable variables and unify them with
       (roughly) the definition of the accessor.*)
     fun surject_prf () =
@@ -1766,8 +1767,8 @@
           simp_tac (HOL_basic_ss addsimps [ext_def]) 1 THEN
           REPEAT_ALL_NEW intros_tac 1;
         val tactic2 = REPEAT (rtac surject_assistI 1 THEN rtac refl 1);
-        val [halfway] = Seq.list_of (tactic1 start);    (* FIXME Seq.lift_of ?? *)
-        val [surject] = Seq.list_of (tactic2 (standard halfway));    (* FIXME Seq.lift_of ?? *)
+        val [halfway] = Seq.list_of (tactic1 start);  (* FIXME Seq.lift_of ?? *)
+        val [surject] = Seq.list_of (tactic2 (Drule.standard halfway));  (* FIXME Seq.lift_of ?? *)
       in
         surject
       end;
@@ -2168,14 +2169,14 @@
     fun sel_convs_prf () =
       map (prove_simp false ss (sel_defs @ accessor_thms)) sel_conv_props;
     val sel_convs = timeit_msg "record sel_convs proof:" sel_convs_prf;
-    fun sel_convs_standard_prf () = map standard sel_convs
+    fun sel_convs_standard_prf () = map Drule.standard sel_convs
     val sel_convs_standard =
       timeit_msg "record sel_convs_standard proof:" sel_convs_standard_prf;
 
     fun upd_convs_prf () =
       map (prove_simp false ss (upd_defs @ updator_thms)) upd_conv_props;
     val upd_convs = timeit_msg "record upd_convs proof:" upd_convs_prf;
-    fun upd_convs_standard_prf () = map standard upd_convs
+    fun upd_convs_standard_prf () = map Drule.standard upd_convs
     val upd_convs_standard =
       timeit_msg "record upd_convs_standard proof:" upd_convs_standard_prf;
 
@@ -2183,7 +2184,7 @@
       let
         val symdefs = map symmetric (sel_defs @ upd_defs);
         val fold_ss = HOL_basic_ss addsimps symdefs;
-        val ua_congs = map (standard o simplify fold_ss) upd_acc_cong_assists;
+        val ua_congs = map (Drule.standard o simplify fold_ss) upd_acc_cong_assists;
       in (ua_congs RL [updacc_foldE], ua_congs RL [updacc_unfoldE]) end;
     val (fold_congs, unfold_congs) =
       timeit_msg "record upd fold/unfold congs:" get_upd_acc_congs;
@@ -2217,7 +2218,7 @@
             [(cterm_of defs_thy Pvar, cterm_of defs_thy
               (lambda w (HOLogic.imp $ HOLogic.mk_eq (r0, w) $ C)))]
             induct_scheme;
-        in standard (ObjectLogic.rulify (mp OF [ind, refl])) end;
+        in Drule.standard (ObjectLogic.rulify (mp OF [ind, refl])) end;
 
     fun cases_scheme_prf_noopt () =
       prove_standard [] cases_scheme_prop
@@ -2262,7 +2263,7 @@
             rtac (prop_subst OF [surjective]) 1,
             REPEAT (etac meta_allE 1), atac 1]);
     val split_meta = timeit_msg "record split_meta proof:" split_meta_prf;
-    fun split_meta_standardise () = standard split_meta;
+    fun split_meta_standardise () = Drule.standard split_meta;
     val split_meta_standard =
       timeit_msg "record split_meta standard:" split_meta_standardise;
 
@@ -2287,7 +2288,7 @@
           |> equal_elim (symmetric split_meta') (*!!r. P r*)
           |> meta_to_obj_all                    (*All r. P r*)
           |> implies_intr cr                    (* 2 ==> 1 *)
-     in standard (thr COMP (thl COMP iffI)) end;
+     in Drule.standard (thr COMP (thl COMP iffI)) end;
 
     fun split_object_prf_noopt () =
       prove_standard [] split_object_prop
--- a/src/HOL/Tools/res_atp.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/res_atp.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,4 +1,4 @@
-(*  Author: Jia Meng, Cambridge University Computer Laboratory, NICTA *)
+(*  Author:     Jia Meng, Cambridge University Computer Laboratory, NICTA *)
 
 signature RES_ATP =
 sig
@@ -156,17 +156,17 @@
 
 fun count_axiom_consts theory_const thy ((thm,_), tab) = 
   let fun count_const (a, T, tab) =
-	let val (c, cts) = const_with_typ thy (a,T)
-	in  (*Two-dimensional table update. Constant maps to types maps to count.*)
-	    Symtab.map_default (c, CTtab.empty) 
-	                       (CTtab.map_default (cts,0) (fn n => n+1)) tab
-	end
+        let val (c, cts) = const_with_typ thy (a,T)
+        in  (*Two-dimensional table update. Constant maps to types maps to count.*)
+            Symtab.map_default (c, CTtab.empty) 
+                               (CTtab.map_default (cts,0) (fn n => n+1)) tab
+        end
       fun count_term_consts (Const(a,T), tab) = count_const(a,T,tab)
-	| count_term_consts (Free(a,T), tab) = count_const(a,T,tab)
-	| count_term_consts (t $ u, tab) =
-	    count_term_consts (t, count_term_consts (u, tab))
-	| count_term_consts (Abs(_,_,t), tab) = count_term_consts (t, tab)
-	| count_term_consts (_, tab) = tab
+        | count_term_consts (Free(a,T), tab) = count_const(a,T,tab)
+        | count_term_consts (t $ u, tab) =
+            count_term_consts (t, count_term_consts (u, tab))
+        | count_term_consts (Abs(_,_,t), tab) = count_term_consts (t, tab)
+        | count_term_consts (_, tab) = tab
   in  count_term_consts (const_prop_of theory_const thm, tab)  end;
 
 
@@ -189,7 +189,7 @@
     let val rel = filter (uni_mem gctyps) consts_typs
         val rel_weight = List.foldl (add_ct_weight ctab) 0.0 rel
     in
-	rel_weight / (rel_weight + real (length consts_typs - length rel))
+        rel_weight / (rel_weight + real (length consts_typs - length rel))
     end;
     
 (*Multiplies out to a list of pairs: 'a * 'b list -> ('a * 'b) list -> ('a * 'b) list*)
@@ -210,17 +210,17 @@
 (*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*)
 fun defines thy (thm,(name,n)) gctypes =
     let val tm = prop_of thm
-	fun defs lhs rhs =
+        fun defs lhs rhs =
             let val (rator,args) = strip_comb lhs
-		val ct = const_with_typ thy (dest_ConstFree rator)
+                val ct = const_with_typ thy (dest_ConstFree rator)
             in  forall is_Var args andalso uni_mem gctypes ct andalso
                 Term.add_vars rhs [] subset Term.add_vars lhs []
             end
-	    handle ConstFree => false
+            handle ConstFree => false
     in    
-	case tm of Const ("Trueprop",_) $ (Const("op =",_) $ lhs $ rhs) => 
-		   defs lhs rhs 
-		 | _ => false
+        case tm of Const ("Trueprop",_) $ (Const("op =",_) $ lhs $ rhs) => 
+                   defs lhs rhs 
+                 | _ => false
     end;
 
 type annotd_cls = (thm * (string * int)) * ((string * const_typ list) list);
@@ -238,40 +238,40 @@
           val accepted = List.take (cls, max_new)
       in
         ResAxioms.trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^ 
-		       ", exceeds the limit of " ^ Int.toString (max_new)));
+                       ", exceeds the limit of " ^ Int.toString (max_new)));
         ResAxioms.trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
         ResAxioms.trace_msg (fn () => "Actually passed: " ^
           space_implode ", " (map (fn (((_,(name,_)),_),_) => name) accepted));
 
-	(map #1 accepted, map #1 (List.drop (cls, max_new)))
+        (map #1 accepted, map #1 (List.drop (cls, max_new)))
       end
   end;
 
 fun relevant_clauses max_new thy ctab p rel_consts =
   let fun relevant ([],_) [] = [] : (thm * (string * int)) list  (*Nothing added this iteration*)
-	| relevant (newpairs,rejects) [] =
-	    let val (newrels,more_rejects) = take_best max_new newpairs
-		val new_consts = maps #2 newrels
-		val rel_consts' = List.foldl add_const_typ_table rel_consts new_consts
-		val newp = p + (1.0-p) / convergence
-	    in
+        | relevant (newpairs,rejects) [] =
+            let val (newrels,more_rejects) = take_best max_new newpairs
+                val new_consts = maps #2 newrels
+                val rel_consts' = List.foldl add_const_typ_table rel_consts new_consts
+                val newp = p + (1.0-p) / convergence
+            in
               ResAxioms.trace_msg (fn () => ("relevant this iteration: " ^ Int.toString (length newrels)));
-	       (map #1 newrels) @ 
-	       (relevant_clauses max_new thy ctab newp rel_consts' (more_rejects@rejects))
-	    end
-	| relevant (newrels,rejects) ((ax as (clsthm as (_,(name,n)),consts_typs)) :: axs) =
-	    let val weight = clause_weight ctab rel_consts consts_typs
-	    in
-	      if p <= weight orelse (follow_defs andalso defines thy clsthm rel_consts)
-	      then (ResAxioms.trace_msg (fn () => (name ^ " clause " ^ Int.toString n ^ 
-	                                    " passes: " ^ Real.toString weight));
-	            relevant ((ax,weight)::newrels, rejects) axs)
-	      else relevant (newrels, ax::rejects) axs
-	    end
+               (map #1 newrels) @ 
+               (relevant_clauses max_new thy ctab newp rel_consts' (more_rejects@rejects))
+            end
+        | relevant (newrels,rejects) ((ax as (clsthm as (_,(name,n)),consts_typs)) :: axs) =
+            let val weight = clause_weight ctab rel_consts consts_typs
+            in
+              if p <= weight orelse (follow_defs andalso defines thy clsthm rel_consts)
+              then (ResAxioms.trace_msg (fn () => (name ^ " clause " ^ Int.toString n ^ 
+                                            " passes: " ^ Real.toString weight));
+                    relevant ((ax,weight)::newrels, rejects) axs)
+              else relevant (newrels, ax::rejects) axs
+            end
     in  ResAxioms.trace_msg (fn () => ("relevant_clauses, current pass mark = " ^ Real.toString p));
         relevant ([],[]) 
     end;
-	
+        
 fun relevance_filter max_new theory_const thy axioms goals = 
  if run_relevance_filter andalso pass_mark >= 0.1
  then
--- a/src/HOL/Tools/res_hol_clause.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Tools/res_hol_clause.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -345,13 +345,13 @@
 fun add_decls (t_full, cma, cnh) (CombConst(c,tp,tvars), (funcs,preds)) =
       if c = "equal" then (addtypes tvars funcs, preds)
       else
-	let val arity = min_arity_of cma c
-	    val ntys = if not t_full then length tvars else 0
-	    val addit = Symtab.update(c, arity+ntys)
-	in
-	    if needs_hBOOL cnh c then (addtypes tvars (addit funcs), preds)
-	    else (addtypes tvars funcs, addit preds)
-	end
+        let val arity = min_arity_of cma c
+            val ntys = if not t_full then length tvars else 0
+            val addit = Symtab.update(c, arity+ntys)
+        in
+            if needs_hBOOL cnh c then (addtypes tvars (addit funcs), preds)
+            else (addtypes tvars funcs, addit preds)
+        end
   | add_decls _ (CombVar(_,ctp), (funcs,preds)) =
       (RC.add_foltype_funcs (ctp,funcs), preds)
   | add_decls params (CombApp(P,Q),decls) = add_decls params (P,add_decls params (Q,decls));
--- a/src/HOL/Transcendental.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Transcendental.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
-(*  Title       : Transcendental.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998,1999 University of Cambridge
-                  1999,2001 University of Edinburgh
-    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+(*  Title:      HOL/Transcendental.thy
+    Author:     Jacques D. Fleuriot, University of Cambridge, University of Edinburgh
+    Author:     Lawrence C Paulson
 *)
 
 header{*Power Series, Transcendental Functions etc.*}
@@ -243,20 +241,20 @@
       assume "n \<ge> (max (2 * f_no) (2 * g_no))" hence "n \<ge> 2 * f_no" and "n \<ge> 2 * g_no" by auto
       have "norm (?Sa n - l) < r"
       proof (cases "even n")
-	case True from even_nat_div_two_times_two[OF this]
-	have n_eq: "2 * (n div 2) = n" unfolding numeral_2_eq_2[symmetric] by auto
-	with `n \<ge> 2 * f_no` have "n div 2 \<ge> f_no" by auto
-	from f[OF this]
-	show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost .
+        case True from even_nat_div_two_times_two[OF this]
+        have n_eq: "2 * (n div 2) = n" unfolding numeral_2_eq_2[symmetric] by auto
+        with `n \<ge> 2 * f_no` have "n div 2 \<ge> f_no" by auto
+        from f[OF this]
+        show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost .
       next
-	case False hence "even (n - 1)" using even_num_iff odd_pos by auto 
-	from even_nat_div_two_times_two[OF this]
-	have n_eq: "2 * ((n - 1) div 2) = n - 1" unfolding numeral_2_eq_2[symmetric] by auto
-	hence range_eq: "n - 1 + 1 = n" using odd_pos[OF False] by auto
-
-	from n_eq `n \<ge> 2 * g_no` have "(n - 1) div 2 \<ge> g_no" by auto
-	from g[OF this]
-	show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost range_eq .
+        case False hence "even (n - 1)" using even_num_iff odd_pos by auto 
+        from even_nat_div_two_times_two[OF this]
+        have n_eq: "2 * ((n - 1) div 2) = n - 1" unfolding numeral_2_eq_2[symmetric] by auto
+        hence range_eq: "n - 1 + 1 = n" using odd_pos[OF False] by auto
+
+        from n_eq `n \<ge> 2 * g_no` have "(n - 1) div 2 \<ge> g_no" by auto
+        from g[OF this]
+        show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost range_eq .
       qed
     }
     thus "\<exists> no. \<forall> n \<ge> no. norm (?Sa n - l) < r" by blast
@@ -698,7 +696,7 @@
 
     { fix n
       have "\<bar> ?diff (n + ?N) x \<bar> \<le> L (n + ?N) * \<bar> (x0 + x) - x0 \<bar> / \<bar> x \<bar>" 
-	using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] unfolding abs_divide .
+        using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] unfolding abs_divide .
       hence "\<bar> ( \<bar> ?diff (n + ?N) x \<bar>) \<bar> \<le> L (n + ?N)" using `x \<noteq> 0` by auto
     } note L_ge = summable_le2[OF allI[OF this] ign[OF `summable L`]]
     from order_trans[OF summable_rabs[OF conjunct1[OF L_ge]] L_ge[THEN conjunct2]]
@@ -713,8 +711,8 @@
       also have "S \<le> S'" using `S \<le> S'` .
       also have "S' \<le> ?s n" unfolding S'_def 
       proof (rule Min_le_iff[THEN iffD2])
-	have "?s n \<in> (?s ` {0..<?N}) \<and> ?s n \<le> ?s n" using `n \<in> { 0 ..< ?N}` by auto
-	thus "\<exists> a \<in> (?s ` {0..<?N}). a \<le> ?s n" by blast
+        have "?s n \<in> (?s ` {0..<?N}) \<and> ?s n \<le> ?s n" using `n \<in> { 0 ..< ?N}` by auto
+        thus "\<exists> a \<in> (?s ` {0..<?N}). a \<le> ?s n" by blast
       qed auto
       finally have "\<bar> x \<bar> < ?s n" .
 
@@ -754,48 +752,48 @@
     proof (rule DERIV_series')
       show "summable (\<lambda> n. \<bar>f n * real (Suc n) * R'^n\<bar>)"
       proof -
-	have "(R' + R) / 2 < R" and "0 < (R' + R) / 2" using `0 < R'` `0 < R` `R' < R` by auto
-	hence in_Rball: "(R' + R) / 2 \<in> {-R <..< R}" using `R' < R` by auto
-	have "norm R' < norm ((R' + R) / 2)" using `0 < R'` `0 < R` `R' < R` by auto
-	from powser_insidea[OF converges[OF in_Rball] this] show ?thesis by auto
+        have "(R' + R) / 2 < R" and "0 < (R' + R) / 2" using `0 < R'` `0 < R` `R' < R` by auto
+        hence in_Rball: "(R' + R) / 2 \<in> {-R <..< R}" using `R' < R` by auto
+        have "norm R' < norm ((R' + R) / 2)" using `0 < R'` `0 < R` `R' < R` by auto
+        from powser_insidea[OF converges[OF in_Rball] this] show ?thesis by auto
       qed
       { fix n x y assume "x \<in> {-R' <..< R'}" and "y \<in> {-R' <..< R'}"
-	show "\<bar>?f x n - ?f y n\<bar> \<le> \<bar>f n * real (Suc n) * R'^n\<bar> * \<bar>x-y\<bar>"
-	proof -
-	  have "\<bar>f n * x ^ (Suc n) - f n * y ^ (Suc n)\<bar> = (\<bar>f n\<bar> * \<bar>x-y\<bar>) * \<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar>" 
-	    unfolding right_diff_distrib[symmetric] lemma_realpow_diff_sumr2 abs_mult by auto
-	  also have "\<dots> \<le> (\<bar>f n\<bar> * \<bar>x-y\<bar>) * (\<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>)" 
-	  proof (rule mult_left_mono)
-	    have "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> (\<Sum>p = 0..<Suc n. \<bar>x ^ p * y ^ (n - p)\<bar>)" by (rule setsum_abs)
-	    also have "\<dots> \<le> (\<Sum>p = 0..<Suc n. R' ^ n)"
-	    proof (rule setsum_mono)
-	      fix p assume "p \<in> {0..<Suc n}" hence "p \<le> n" by auto
-	      { fix n fix x :: real assume "x \<in> {-R'<..<R'}"
-		hence "\<bar>x\<bar> \<le> R'"  by auto
-		hence "\<bar>x^n\<bar> \<le> R'^n" unfolding power_abs by (rule power_mono, auto)
-	      } from mult_mono[OF this[OF `x \<in> {-R'<..<R'}`, of p] this[OF `y \<in> {-R'<..<R'}`, of "n-p"]] `0 < R'`
-	      have "\<bar>x^p * y^(n-p)\<bar> \<le> R'^p * R'^(n-p)" unfolding abs_mult by auto
-	      thus "\<bar>x^p * y^(n-p)\<bar> \<le> R'^n" unfolding power_add[symmetric] using `p \<le> n` by auto
-	    qed
-	    also have "\<dots> = real (Suc n) * R' ^ n" unfolding setsum_constant card_atLeastLessThan real_of_nat_def by auto
-	    finally show "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> \<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>" unfolding abs_real_of_nat_cancel abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF `0 < R'`]]] .
-	    show "0 \<le> \<bar>f n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult[symmetric] by auto
-	  qed
-	  also have "\<dots> = \<bar>f n * real (Suc n) * R' ^ n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult real_mult_assoc[symmetric] by algebra
-	  finally show ?thesis .
-	qed }
+        show "\<bar>?f x n - ?f y n\<bar> \<le> \<bar>f n * real (Suc n) * R'^n\<bar> * \<bar>x-y\<bar>"
+        proof -
+          have "\<bar>f n * x ^ (Suc n) - f n * y ^ (Suc n)\<bar> = (\<bar>f n\<bar> * \<bar>x-y\<bar>) * \<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar>" 
+            unfolding right_diff_distrib[symmetric] lemma_realpow_diff_sumr2 abs_mult by auto
+          also have "\<dots> \<le> (\<bar>f n\<bar> * \<bar>x-y\<bar>) * (\<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>)" 
+          proof (rule mult_left_mono)
+            have "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> (\<Sum>p = 0..<Suc n. \<bar>x ^ p * y ^ (n - p)\<bar>)" by (rule setsum_abs)
+            also have "\<dots> \<le> (\<Sum>p = 0..<Suc n. R' ^ n)"
+            proof (rule setsum_mono)
+              fix p assume "p \<in> {0..<Suc n}" hence "p \<le> n" by auto
+              { fix n fix x :: real assume "x \<in> {-R'<..<R'}"
+                hence "\<bar>x\<bar> \<le> R'"  by auto
+                hence "\<bar>x^n\<bar> \<le> R'^n" unfolding power_abs by (rule power_mono, auto)
+              } from mult_mono[OF this[OF `x \<in> {-R'<..<R'}`, of p] this[OF `y \<in> {-R'<..<R'}`, of "n-p"]] `0 < R'`
+              have "\<bar>x^p * y^(n-p)\<bar> \<le> R'^p * R'^(n-p)" unfolding abs_mult by auto
+              thus "\<bar>x^p * y^(n-p)\<bar> \<le> R'^n" unfolding power_add[symmetric] using `p \<le> n` by auto
+            qed
+            also have "\<dots> = real (Suc n) * R' ^ n" unfolding setsum_constant card_atLeastLessThan real_of_nat_def by auto
+            finally show "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> \<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>" unfolding abs_real_of_nat_cancel abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF `0 < R'`]]] .
+            show "0 \<le> \<bar>f n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult[symmetric] by auto
+          qed
+          also have "\<dots> = \<bar>f n * real (Suc n) * R' ^ n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult real_mult_assoc[symmetric] by algebra
+          finally show ?thesis .
+        qed }
       { fix n show "DERIV (\<lambda> x. ?f x n) x0 :> (?f' x0 n)"
-	  by (auto intro!: DERIV_intros simp del: power_Suc) }
+          by (auto intro!: DERIV_intros simp del: power_Suc) }
       { fix x assume "x \<in> {-R' <..< R'}" hence "R' \<in> {-R <..< R}" and "norm x < norm R'" using assms `R' < R` by auto
-	have "summable (\<lambda> n. f n * x^n)"
-	proof (rule summable_le2[THEN conjunct1, OF _ powser_insidea[OF converges[OF `R' \<in> {-R <..< R}`] `norm x < norm R'`]], rule allI)
-	  fix n
-	  have le: "\<bar>f n\<bar> * 1 \<le> \<bar>f n\<bar> * real (Suc n)" by (rule mult_left_mono, auto)
-	  show "\<bar>f n * x ^ n\<bar> \<le> norm (f n * real (Suc n) * x ^ n)" unfolding real_norm_def abs_mult
-	    by (rule mult_right_mono, auto simp add: le[unfolded mult_1_right])
-	qed
-	from this[THEN summable_mult2[where c=x], unfolded real_mult_assoc, unfolded real_mult_commute]
-	show "summable (?f x)" by auto }
+        have "summable (\<lambda> n. f n * x^n)"
+        proof (rule summable_le2[THEN conjunct1, OF _ powser_insidea[OF converges[OF `R' \<in> {-R <..< R}`] `norm x < norm R'`]], rule allI)
+          fix n
+          have le: "\<bar>f n\<bar> * 1 \<le> \<bar>f n\<bar> * real (Suc n)" by (rule mult_left_mono, auto)
+          show "\<bar>f n * x ^ n\<bar> \<le> norm (f n * real (Suc n) * x ^ n)" unfolding real_norm_def abs_mult
+            by (rule mult_right_mono, auto simp add: le[unfolded mult_1_right])
+        qed
+        from this[THEN summable_mult2[where c=x], unfolded real_mult_assoc, unfolded real_mult_commute]
+        show "summable (?f x)" by auto }
       show "summable (?f' x0)" using converges[OF `x0 \<in> {-R <..< R}`] .
       show "x0 \<in> {-R' <..< R'}" using `x0 \<in> {-R' <..< R'}` .
     qed
@@ -1234,9 +1232,9 @@
     proof (rule DERIV_power_series')
       show "x - 1 \<in> {- 1<..<1}" and "(0 :: real) < 1" using `0 < x` `x < 2` by auto
       { fix x :: real assume "x \<in> {- 1<..<1}" hence "norm (-x) < 1" by auto
-	show "summable (\<lambda>n. -1 ^ n * (1 / real (n + 1)) * real (Suc n) * x ^ n)"
+        show "summable (\<lambda>n. -1 ^ n * (1 / real (n + 1)) * real (Suc n) * x ^ n)"
           unfolding One_nat_def
-	  by (auto simp del: power_mult_distrib simp add: power_mult_distrib[symmetric] summable_geometric[OF `norm (-x) < 1`])
+          by (auto simp del: power_mult_distrib simp add: power_mult_distrib[symmetric] summable_geometric[OF `norm (-x) < 1`])
       }
     qed
     hence "DERIV (\<lambda>x. suminf (?f x)) (x - 1) :> suminf (?f' x)" unfolding One_nat_def by auto
@@ -2172,7 +2170,7 @@
 txt{*Now, simulate TRYALL*}
 apply (rule_tac [!] DERIV_tan asm_rl)
 apply (auto dest!: DERIV_unique [OF _ DERIV_tan]
-	    simp add: cos_gt_zero_pi [THEN less_imp_neq, THEN not_sym]) 
+            simp add: cos_gt_zero_pi [THEN less_imp_neq, THEN not_sym]) 
 done
 
 lemma tan_monotone: assumes "- (pi / 2) < y" and "y < x" and "x < pi / 2"
@@ -2749,10 +2747,10 @@
     { fix n fix x :: real assume "0 \<le> x" and "x \<le> 1"
       have "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \<le> 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)"
       proof (rule mult_mono)
-	show "1 / real (Suc (Suc n * 2)) \<le> 1 / real (Suc (n * 2))" by (rule frac_le) simp_all
-	show "0 \<le> 1 / real (Suc (n * 2))" by auto
-	show "x ^ Suc (Suc n * 2) \<le> x ^ Suc (n * 2)" by (rule power_decreasing) (simp_all add: `0 \<le> x` `x \<le> 1`)
-	show "0 \<le> x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: `0 \<le> x`)
+        show "1 / real (Suc (Suc n * 2)) \<le> 1 / real (Suc (n * 2))" by (rule frac_le) simp_all
+        show "0 \<le> 1 / real (Suc (n * 2))" by auto
+        show "x ^ Suc (Suc n * 2) \<le> x ^ Suc (n * 2)" by (rule power_decreasing) (simp_all add: `0 \<le> x` `x \<le> 1`)
+        show "0 \<le> x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: `0 \<le> x`)
       qed
     } note mono = this
     
@@ -2849,7 +2847,7 @@
 
       let ?S = "\<Sum> n. (-1)^n * x'^(2 * n)"
       show "summable (\<lambda> n. ?f n * real (Suc n) * x'^n)" unfolding if_eq
-	by (rule sums_summable[where l="0 + ?S"], rule sums_if, rule sums_zero, rule summable_sums, rule summable_Integral[OF `\<bar>x'\<bar> < 1`])
+        by (rule sums_summable[where l="0 + ?S"], rule sums_if, rule sums_zero, rule summable_sums, rule summable_Integral[OF `\<bar>x'\<bar> < 1`])
     }
   qed auto
   thus ?thesis unfolding Int_eq arctan_eq .
@@ -2880,23 +2878,23 @@
       hence "\<bar>x\<bar> < r" by auto
       show "suminf (?c x) - arctan x = suminf (?c a) - arctan a"
       proof (rule DERIV_isconst2[of "a" "b"])
-	show "a < b" and "a \<le> x" and "x \<le> b" using `a < b` `a \<le> x` `x \<le> b` by auto
-	have "\<forall> x. -r < x \<and> x < r \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
-	proof (rule allI, rule impI)
-	  fix x assume "-r < x \<and> x < r" hence "\<bar>x\<bar> < r" by auto
-	  hence "\<bar>x\<bar> < 1" using `r < 1` by auto
-	  have "\<bar> - (x^2) \<bar> < 1" using less_one_imp_sqr_less_one[OF `\<bar>x\<bar> < 1`] by auto
-	  hence "(\<lambda> n. (- (x^2)) ^ n) sums (1 / (1 - (- (x^2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums)
-	  hence "(?c' x) sums (1 / (1 - (- (x^2))))" unfolding power_mult_distrib[symmetric] power_mult nat_mult_commute[of _ 2] by auto
-	  hence suminf_c'_eq_geom: "inverse (1 + x^2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto
-	  have "DERIV (\<lambda> x. suminf (?c x)) x :> (inverse (1 + x^2))" unfolding suminf_c'_eq_geom
-	    by (rule DERIV_arctan_suminf[OF `0 < r` `r < 1` `\<bar>x\<bar> < r`])
-	  from DERIV_add_minus[OF this DERIV_arctan]
-	  show "DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0" unfolding diff_minus by auto
-	qed
-	hence DERIV_in_rball: "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0" using `-r < a` `b < r` by auto
-	thus "\<forall> y. a < y \<and> y < b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0" using `\<bar>x\<bar> < r` by auto
-	show "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> isCont (\<lambda> x. suminf (?c x) - arctan x) y" using DERIV_in_rball DERIV_isCont by auto
+        show "a < b" and "a \<le> x" and "x \<le> b" using `a < b` `a \<le> x` `x \<le> b` by auto
+        have "\<forall> x. -r < x \<and> x < r \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
+        proof (rule allI, rule impI)
+          fix x assume "-r < x \<and> x < r" hence "\<bar>x\<bar> < r" by auto
+          hence "\<bar>x\<bar> < 1" using `r < 1` by auto
+          have "\<bar> - (x^2) \<bar> < 1" using less_one_imp_sqr_less_one[OF `\<bar>x\<bar> < 1`] by auto
+          hence "(\<lambda> n. (- (x^2)) ^ n) sums (1 / (1 - (- (x^2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums)
+          hence "(?c' x) sums (1 / (1 - (- (x^2))))" unfolding power_mult_distrib[symmetric] power_mult nat_mult_commute[of _ 2] by auto
+          hence suminf_c'_eq_geom: "inverse (1 + x^2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto
+          have "DERIV (\<lambda> x. suminf (?c x)) x :> (inverse (1 + x^2))" unfolding suminf_c'_eq_geom
+            by (rule DERIV_arctan_suminf[OF `0 < r` `r < 1` `\<bar>x\<bar> < r`])
+          from DERIV_add_minus[OF this DERIV_arctan]
+          show "DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0" unfolding diff_minus by auto
+        qed
+        hence DERIV_in_rball: "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0" using `-r < a` `b < r` by auto
+        thus "\<forall> y. a < y \<and> y < b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0" using `\<bar>x\<bar> < r` by auto
+        show "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> isCont (\<lambda> x. suminf (?c x) - arctan x) y" using DERIV_in_rball DERIV_isCont by auto
       qed
     qed
     
@@ -2909,10 +2907,10 @@
     next
       case False hence "0 < \<bar>x\<bar>" and "- \<bar>x\<bar> < \<bar>x\<bar>" by auto
       have "suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>) = suminf (?c 0) - arctan 0"
-	by (rule suminf_eq_arctan_bounded[where x="0" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>", symmetric], auto simp add: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>`)
+        by (rule suminf_eq_arctan_bounded[where x="0" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>", symmetric], auto simp add: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>`)
       moreover
       have "suminf (?c x) - arctan x = suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>)"
-	by (rule suminf_eq_arctan_bounded[where x="x" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>"], auto simp add: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>`)
+        by (rule suminf_eq_arctan_bounded[where x="x" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>"], auto simp add: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>`)
       ultimately 
       show ?thesis using suminf_arctan_zero by auto
     qed
@@ -2929,35 +2927,35 @@
       have "0 < (1 :: real)" by auto
       moreover
       { fix x :: real assume "0 < x" and "x < 1" hence "\<bar>x\<bar> \<le> 1" and "\<bar>x\<bar> < 1" by auto
-	from `0 < x` have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)" by auto
-	note bounds = mp[OF arctan_series_borders(2)[OF `\<bar>x\<bar> \<le> 1`] this, unfolded when_less_one[OF `\<bar>x\<bar> < 1`, symmetric], THEN spec]
-	have "0 < 1 / real (n*2+1) * x^(n*2+1)" by (rule mult_pos_pos, auto simp only: zero_less_power[OF `0 < x`], auto)
-	hence a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)" by (rule abs_of_pos)
+        from `0 < x` have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)" by auto
+        note bounds = mp[OF arctan_series_borders(2)[OF `\<bar>x\<bar> \<le> 1`] this, unfolded when_less_one[OF `\<bar>x\<bar> < 1`, symmetric], THEN spec]
+        have "0 < 1 / real (n*2+1) * x^(n*2+1)" by (rule mult_pos_pos, auto simp only: zero_less_power[OF `0 < x`], auto)
+        hence a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)" by (rule abs_of_pos)
         have "?diff x n \<le> ?a x n"
-	proof (cases "even n")
-	  case True hence sgn_pos: "(-1)^n = (1::real)" by auto
-	  from `even n` obtain m where "2 * m = n" unfolding even_mult_two_ex by auto
-	  from bounds[of m, unfolded this atLeastAtMost_iff]
-	  have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n + 1. (?c x i)) - (\<Sum>i = 0..<n. (?c x i))" by auto
-	  also have "\<dots> = ?c x n" unfolding One_nat_def by auto
-	  also have "\<dots> = ?a x n" unfolding sgn_pos a_pos by auto
-	  finally show ?thesis .
-	next
-	  case False hence sgn_neg: "(-1)^n = (-1::real)" by auto
-	  from `odd n` obtain m where m_def: "2 * m + 1 = n" unfolding odd_Suc_mult_two_ex by auto
-	  hence m_plus: "2 * (m + 1) = n + 1" by auto
-	  from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2]
-	  have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n. (?c x i)) - (\<Sum>i = 0..<n+1. (?c x i))" by auto
-	  also have "\<dots> = - ?c x n" unfolding One_nat_def by auto
-	  also have "\<dots> = ?a x n" unfolding sgn_neg a_pos by auto
-	  finally show ?thesis .
-	qed
+        proof (cases "even n")
+          case True hence sgn_pos: "(-1)^n = (1::real)" by auto
+          from `even n` obtain m where "2 * m = n" unfolding even_mult_two_ex by auto
+          from bounds[of m, unfolded this atLeastAtMost_iff]
+          have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n + 1. (?c x i)) - (\<Sum>i = 0..<n. (?c x i))" by auto
+          also have "\<dots> = ?c x n" unfolding One_nat_def by auto
+          also have "\<dots> = ?a x n" unfolding sgn_pos a_pos by auto
+          finally show ?thesis .
+        next
+          case False hence sgn_neg: "(-1)^n = (-1::real)" by auto
+          from `odd n` obtain m where m_def: "2 * m + 1 = n" unfolding odd_Suc_mult_two_ex by auto
+          hence m_plus: "2 * (m + 1) = n + 1" by auto
+          from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2]
+          have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n. (?c x i)) - (\<Sum>i = 0..<n+1. (?c x i))" by auto
+          also have "\<dots> = - ?c x n" unfolding One_nat_def by auto
+          also have "\<dots> = ?a x n" unfolding sgn_neg a_pos by auto
+          finally show ?thesis .
+        qed
         hence "0 \<le> ?a x n - ?diff x n" by auto
       }
       hence "\<forall> x \<in> { 0 <..< 1 }. 0 \<le> ?a x n - ?diff x n" by auto
       moreover have "\<And>x. isCont (\<lambda> x. ?a x n - ?diff x n) x"
-	unfolding real_diff_def divide_inverse
-	by (auto intro!: isCont_add isCont_rabs isCont_ident isCont_minus isCont_arctan isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum)
+        unfolding real_diff_def divide_inverse
+        by (auto intro!: isCont_add isCont_rabs isCont_ident isCont_minus isCont_arctan isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum)
       ultimately have "0 \<le> ?a 1 n - ?diff 1 n" by (rule LIM_less_bound)
       hence "?diff 1 n \<le> ?a 1 n" by auto
     }
@@ -2969,7 +2967,7 @@
       fix r :: real assume "0 < r"
       obtain N :: nat where N_I: "\<And> n. N \<le> n \<Longrightarrow> ?a 1 n < r" using LIMSEQ_D[OF `?a 1 ----> 0` `0 < r`] by auto
       { fix n assume "N \<le> n" from `?diff 1 n \<le> ?a 1 n` N_I[OF this]
-	have "norm (?diff 1 n - 0) < r" by auto }
+        have "norm (?diff 1 n - 0) < r" by auto }
       thus "\<exists> N. \<forall> n \<ge> N. norm (?diff 1 n - 0) < r" by blast
     qed
     from this[unfolded LIMSEQ_rabs_zero real_diff_def add_commute[of "arctan 1"], THEN LIMSEQ_add_const, of "- arctan 1", THEN LIMSEQ_minus]
--- a/src/HOL/UNITY/Comp.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Comp.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,16 +1,11 @@
 (*  Title:      HOL/UNITY/Comp.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1998  University of Cambridge
+    Author:     Sidi Ehmety
 
-Composition
+Composition.
+
 From Chandy and Sanders, "Reasoning About Program Composition",
 Technical Report 2000-003, University of Florida, 2000.
-
-Revised by Sidi Ehmety on January  2001
-
-Added: a strong form of the \<subseteq> relation (component_of) and localize
-
 *)
 
 header{*Composition: Basic Primitives*}
@@ -46,7 +41,7 @@
 
   localize  :: "('a=>'b) => 'a program => 'a program"
   "localize v F == mk_program(Init F, Acts F,
-			      AllowedActs F \<inter> (\<Union>G \<in> preserves v. Acts G))"
+                              AllowedActs F \<inter> (\<Union>G \<in> preserves v. Acts G))"
 
   funPair      :: "['a => 'b, 'a => 'c, 'a] => 'b * 'c"
   "funPair f g == %x. (f x, g x)"
--- a/src/HOL/UNITY/Comp/AllocBase.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Comp/AllocBase.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
-(*  Title:      HOL/UNITY/AllocBase.thy
-    ID:         $Id$
+(*  Title:      HOL/UNITY/Comp/AllocBase.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
-
 *)
 
 header{*Common Declarations for Chandy and Charpentier's Allocator*}
@@ -87,7 +85,7 @@
      "bag_of (sublist l (A Un B)) + bag_of (sublist l (A Int B)) =  
       bag_of (sublist l A) + bag_of (sublist l B)"
 apply (subgoal_tac "A Int B Int {..<length l} =
-	            (A Int {..<length l}) Int (B Int {..<length l}) ")
+                    (A Int {..<length l}) Int (B Int {..<length l}) ")
 apply (simp add: bag_of_sublist Int_Un_distrib2 setsum_Un_Int, blast)
 done
 
--- a/src/HOL/UNITY/Comp/AllocImpl.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Comp/AllocImpl.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      HOL/UNITY/AllocImpl
-    ID:         $Id$
+(*  Title:      HOL/UNITY/Comp/AllocImpl.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 *)
@@ -75,12 +74,12 @@
   (*spec (13)*)
   merge_follows :: "('a,'b) merge_d program set"
     "merge_follows ==
-	 (\<Inter>i \<in> lessThan Nclients. Increasing (sub i o merge.In))
-	 guarantees
-	 (\<Inter>i \<in> lessThan Nclients.
-	  (%s. sublist (merge.Out s)
+         (\<Inter>i \<in> lessThan Nclients. Increasing (sub i o merge.In))
+         guarantees
+         (\<Inter>i \<in> lessThan Nclients.
+          (%s. sublist (merge.Out s)
                        {k. k < size(merge.iOut s) & merge.iOut s! k = i})
-	  Fols (sub i o merge.In))"
+          Fols (sub i o merge.In))"
 
   (*spec: preserves part*)
   merge_preserves :: "('a,'b) merge_d program set"
@@ -90,7 +89,7 @@
   merge_allowed_acts :: "('a,'b) merge_d program set"
     "merge_allowed_acts ==
        {F. AllowedActs F =
-	    insert Id (UNION (preserves (funPair merge.Out merge.iOut)) Acts)}"
+            insert Id (UNION (preserves (funPair merge.Out merge.iOut)) Acts)}"
 
   merge_spec :: "('a,'b) merge_d program set"
     "merge_spec == merge_increasing Int merge_eqOut Int merge_bounded Int
@@ -101,12 +100,12 @@
   (*spec (14)*)
   distr_follows :: "('a,'b) distr_d program set"
     "distr_follows ==
-	 Increasing distr.In Int Increasing distr.iIn Int
-	 Always {s. \<forall>elt \<in> set (distr.iIn s). elt < Nclients}
-	 guarantees
-	 (\<Inter>i \<in> lessThan Nclients.
-	  (sub i o distr.Out) Fols
-	  (%s. sublist (distr.In s)
+         Increasing distr.In Int Increasing distr.iIn Int
+         Always {s. \<forall>elt \<in> set (distr.iIn s). elt < Nclients}
+         guarantees
+         (\<Inter>i \<in> lessThan Nclients.
+          (sub i o distr.Out) Fols
+          (%s. sublist (distr.In s)
                        {k. k < size(distr.iIn s) & distr.iIn s ! k = i}))"
 
   distr_allowed_acts :: "('a,'b) distr_d program set"
@@ -125,18 +124,18 @@
   (*spec (19)*)
   alloc_safety :: "'a allocState_d program set"
     "alloc_safety ==
-	 Increasing rel
+         Increasing rel
          guarantees  Always {s. tokens (giv s) \<le> NbT + tokens (rel s)}"
 
   (*spec (20)*)
   alloc_progress :: "'a allocState_d program set"
     "alloc_progress ==
-	 Increasing ask Int Increasing rel Int
+         Increasing ask Int Increasing rel Int
          Always {s. \<forall>elt \<in> set (ask s). elt \<le> NbT}
          Int
          (\<Inter>h. {s. h \<le> giv s & h pfixGe (ask s)}
-		 LeadsTo
-	         {s. tokens h \<le> tokens (rel s)})
+                 LeadsTo
+                 {s. tokens h \<le> tokens (rel s)})
          guarantees  (\<Inter>h. {s. h \<le> ask s} LeadsTo {s. h pfixLe giv s})"
 
   (*spec: preserves part*)
@@ -171,52 +170,52 @@
 
 #    {*spec (9.1)*}
 #    network_ask :: "'a systemState program set
-#	"network_ask == \<Inter>i \<in> lessThan Nclients.
-#			    Increasing (ask o sub i o client)
-#			    guarantees[ask]
-#			    (ask  Fols (ask o sub i o client))"
+#       "network_ask == \<Inter>i \<in> lessThan Nclients.
+#                           Increasing (ask o sub i o client)
+#                           guarantees[ask]
+#                           (ask  Fols (ask o sub i o client))"
 
 #    {*spec (9.2)*}
 #    network_giv :: "'a systemState program set
-#	"network_giv == \<Inter>i \<in> lessThan Nclients.
-#			    Increasing giv
-#			    guarantees[giv o sub i o client]
-#			    ((giv o sub i o client) Fols giv)"
+#       "network_giv == \<Inter>i \<in> lessThan Nclients.
+#                           Increasing giv
+#                           guarantees[giv o sub i o client]
+#                           ((giv o sub i o client) Fols giv)"
 
 #    {*spec (9.3)*}
 #    network_rel :: "'a systemState program set
-#	"network_rel == \<Inter>i \<in> lessThan Nclients.
-#			    Increasing (rel o sub i o client)
-#			    guarantees[rel]
-#			    (rel  Fols (rel o sub i o client))"
+#       "network_rel == \<Inter>i \<in> lessThan Nclients.
+#                           Increasing (rel o sub i o client)
+#                           guarantees[rel]
+#                           (rel  Fols (rel o sub i o client))"
 
 #    {*spec: preserves part*}
-#	network_preserves :: "'a systemState program set
-#	"network_preserves == preserves giv  Int
-#			      (\<Inter>i \<in> lessThan Nclients.
-#			       preserves (funPair rel ask o sub i o client))"
+#       network_preserves :: "'a systemState program set
+#       "network_preserves == preserves giv  Int
+#                             (\<Inter>i \<in> lessThan Nclients.
+#                              preserves (funPair rel ask o sub i o client))"
 
 #    network_spec :: "'a systemState program set
-#	"network_spec == network_ask Int network_giv Int
-#			 network_rel Int network_preserves"
+#       "network_spec == network_ask Int network_giv Int
+#                        network_rel Int network_preserves"
 
 
 #  {** State mappings **}
 #    sysOfAlloc :: "((nat => merge) * 'a) allocState_d => 'a systemState"
-#	"sysOfAlloc == %s. let (cl,xtr) = allocState_d.dummy s
-#			   in (| giv = giv s,
-#				 ask = ask s,
-#				 rel = rel s,
-#				 client   = cl,
-#				 dummy    = xtr|)"
+#       "sysOfAlloc == %s. let (cl,xtr) = allocState_d.dummy s
+#                          in (| giv = giv s,
+#                                ask = ask s,
+#                                rel = rel s,
+#                                client   = cl,
+#                                dummy    = xtr|)"
 
 
 #    sysOfClient :: "(nat => merge) * 'a allocState_d => 'a systemState"
-#	"sysOfClient == %(cl,al). (| giv = giv al,
-#				     ask = ask al,
-#				     rel = rel al,
-#				     client   = cl,
-#				     systemState.dummy = allocState_d.dummy al|)"
+#       "sysOfClient == %(cl,al). (| giv = giv al,
+#                                    ask = ask al,
+#                                    rel = rel al,
+#                                    client   = cl,
+#                                    systemState.dummy = allocState_d.dummy al|)"
 ****)
 
 
--- a/src/HOL/UNITY/Comp/Client.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Comp/Client.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      HOL/UNITY/Client.thy
-    ID:         $Id$
+(*  Title:      HOL/UNITY/Comp/Client.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 *)
@@ -9,13 +8,13 @@
 theory Client imports Rename AllocBase begin
 
 types
-  tokbag = nat	   --{*tokbags could be multisets...or any ordered type?*}
+  tokbag = nat     --{*tokbags could be multisets...or any ordered type?*}
 
 record state =
   giv :: "tokbag list" --{*input history: tokens granted*}
   ask :: "tokbag list" --{*output history: tokens requested*}
   rel :: "tokbag list" --{*output history: tokens released*}
-  tok :: tokbag	       --{*current token request*}
+  tok :: tokbag        --{*current token request*}
 
 record 'a state_d =
   state +  
@@ -30,10 +29,10 @@
   
   rel_act :: "('a state_d * 'a state_d) set"
     "rel_act == {(s,s').
-		  \<exists>nrel. nrel = size (rel s) &
-		         s' = s (| rel := rel s @ [giv s!nrel] |) &
-		         nrel < size (giv s) &
-		         ask s!nrel \<le> giv s!nrel}"
+                  \<exists>nrel. nrel = size (rel s) &
+                         s' = s (| rel := rel s @ [giv s!nrel] |) &
+                         nrel < size (giv s) &
+                         ask s!nrel \<le> giv s!nrel}"
 
   (** Choose a new token requirement **)
 
@@ -45,16 +44,16 @@
   
   ask_act :: "('a state_d * 'a state_d) set"
     "ask_act == {(s,s'). s'=s |
-		         (s' = s (|ask := ask s @ [tok s]|))}"
+                         (s' = s (|ask := ask s @ [tok s]|))}"
 
   Client :: "'a state_d program"
     "Client ==
        mk_total_program
             ({s. tok s \<in> atMost NbT &
-		 giv s = [] & ask s = [] & rel s = []},
-	     {rel_act, tok_act, ask_act},
-	     \<Union>G \<in> preserves rel Int preserves ask Int preserves tok.
-		   Acts G)"
+                 giv s = [] & ask s = [] & rel s = []},
+             {rel_act, tok_act, ask_act},
+             \<Union>G \<in> preserves rel Int preserves ask Int preserves tok.
+                   Acts G)"
 
   (*Maybe want a special theory section to declare such maps*)
   non_dummy :: "'a state_d => state"
--- a/src/HOL/UNITY/Comp/Counter.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Comp/Counter.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,12 +1,11 @@
-(*  Title:      HOL/UNITY/Counter
-    ID:         $Id$
+(*  Title:      HOL/UNITY/Comp/Counter.thy
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2001  University of Cambridge
 
 From Charpentier and Chandy,
 Examples of Program Composition Illustrating the Use of Universal Properties
    In J. Rolim (editor), Parallel and Distributed Processing,
-   Spriner LNCS 1586 (1999), pages 1215-1227.
+   Springer LNCS 1586 (1999), pages 1215-1227.
 *)
 
 header{*A Family of Similar Counters: Original Version*}
@@ -38,7 +37,7 @@
   Component :: "nat => state program"
   "Component i ==
     mk_total_program({s. s C = 0 & s (c i) = 0}, {a i},
-	             \<Union>G \<in> preserves (%s. s (c i)). Acts G)"
+                     \<Union>G \<in> preserves (%s. s (c i)). Acts G)"
 
 
 
--- a/src/HOL/UNITY/Comp/Counterc.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Comp/Counterc.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,9 @@
-(*  Title:      HOL/UNITY/Counterc
-    ID:         $Id$
+(*  Title:      HOL/UNITY/Comp/Counterc.thy
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2001  University of Cambridge
 
-A family of similar counters, version with a full use of "compatibility "
+A family of similar counters, version with a full use of
+"compatibility ".
 
 From Charpentier and Chandy,
 Examples of Program Composition Illustrating the Use of Universal Properties
@@ -41,8 +41,8 @@
  
   Component :: "nat => state program"
   "Component i == mk_total_program({s. C s = 0 & (c s) i = 0},
-				   {a i},
- 	                           \<Union>G \<in> preserves (%s. (c s) i). Acts G)"
+                                   {a i},
+                                   \<Union>G \<in> preserves (%s. (c s) i). Acts G)"
 
 
 declare Component_def [THEN def_prg_Init, simp]
@@ -101,7 +101,7 @@
 lemma p2_p3_lemma1: 
      "[| OK {i. i<I} Component; i<I |] ==>  
       \<forall>k. Component i \<in> stable ({s. C s = c s i + sumj I i k} Int  
-	 	                {s. \<forall>j\<in>{i. i<I}. j\<noteq>i --> c s j = c k j})"
+                                {s. \<forall>j\<in>{i. i<I}. j\<noteq>i --> c s j = c k j})"
 by (blast intro: stable_Int [OF p2 p3])
 
 lemma p2_p3_lemma2:
--- a/src/HOL/UNITY/Constrains.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Constrains.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      HOL/UNITY/Constrains
-    ID:         $Id$
+(*  Title:      HOL/UNITY/Constrains.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -21,7 +20,7 @@
     Init:  "s \<in> init ==> (s,[]) \<in> traces init acts"
 
   | Acts:  "[| act: acts;  (s,evs) \<in> traces init acts;  (s,s'): act |]
-	    ==> (s', s#evs) \<in> traces init acts"
+            ==> (s', s#evs) \<in> traces init acts"
 
 
 inductive_set
@@ -31,7 +30,7 @@
     Init:  "s \<in> Init F ==> s \<in> reachable F"
 
   | Acts:  "[| act: Acts F;  s \<in> reachable F;  (s,s'): act |]
-	    ==> s' \<in> reachable F"
+            ==> s' \<in> reachable F"
 
 constdefs
   Constrains :: "['a set, 'a set] => 'a program set"  (infixl "Co" 60)
--- a/src/HOL/UNITY/ELT.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/ELT.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      HOL/UNITY/ELT
-    ID:         $Id$
+(*  Title:      HOL/UNITY/ELT.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1999  University of Cambridge
 
@@ -15,7 +14,7 @@
     Weaken:  "A <= B ==> A : elt CC F B"
 
     ETrans:  "[| F : A ensures A';  A-A' : CC;  A' : elt CC F B |]
-	      ==> A : elt CC F B"
+              ==> A : elt CC F B"
 
     Union:  "{A. A: S} : Pow (elt CC F B) ==> (Union S) : elt CC F B"
 
--- a/src/HOL/UNITY/Extend.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Extend.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/Extend.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -36,16 +35,16 @@
 
   extend :: "['a*'b => 'c, 'a program] => 'c program"
     "extend h F == mk_program (extend_set h (Init F),
-			       extend_act h ` Acts F,
-			       project_act h -` AllowedActs F)"
+                               extend_act h ` Acts F,
+                               project_act h -` AllowedActs F)"
 
   (*Argument C allows weak safety laws to be projected*)
   project :: "['a*'b => 'c, 'c set, 'c program] => 'a program"
     "project h C F ==
        mk_program (project_set h (Init F),
-		   project_act h ` Restrict C ` Acts F,
-		   {act. Restrict (project_set h C) act :
-		         project_act h ` Restrict C ` AllowedActs F})"
+                   project_act h ` Restrict C ` Acts F,
+                   {act. Restrict (project_set h C) act :
+                         project_act h ` Restrict C ` AllowedActs F})"
 
 locale Extend =
   fixes f     :: "'c => 'a"
@@ -104,7 +103,7 @@
 (*Possibly easier than reasoning about "inv h"*)
 lemma good_mapI: 
      assumes surj_h: "surj h"
-	 and prem:   "!! x x' y y'. h(x,y) = h(x',y') ==> x=x'"
+         and prem:   "!! x x' y y'. h(x,y) = h(x',y') ==> x=x'"
      shows "good_map h"
 apply (simp add: good_map_def) 
 apply (safe intro!: surj_h)
@@ -120,7 +119,7 @@
 (*A convenient way of finding a closed form for inv h*)
 lemma fst_inv_equalityI: 
      assumes surj_h: "surj h"
-	 and prem:   "!! x y. g (h(x,y)) = x"
+         and prem:   "!! x y. g (h(x,y)) = x"
      shows "fst (inv h z) = g z"
 apply (unfold inv_def)
 apply (rule_tac y1 = z in surj_h [THEN surjD, THEN exE])
--- a/src/HOL/UNITY/Follows.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Follows.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,4 +1,4 @@
-(*  Title:      HOL/UNITY/Follows
+(*  Title:      HOL/UNITY/Follows.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 *)
@@ -39,8 +39,8 @@
 
 lemma mono_Follows_o: "mono h ==> f Fols g \<subseteq> (h o f) Fols (h o g)"
 by (auto simp add: Follows_def mono_Increasing_o [THEN [2] rev_subsetD]
-		   mono_Always_o [THEN [2] rev_subsetD]
-		   mono_LeadsTo_o [THEN [2] rev_subsetD, THEN INT_D])
+                   mono_Always_o [THEN [2] rev_subsetD]
+                   mono_LeadsTo_o [THEN [2] rev_subsetD, THEN INT_D])
 
 lemma mono_Follows_apply:
      "mono h ==> f Fols g \<subseteq> (%x. h (f x)) Fols (%x. h (g x))"
--- a/src/HOL/UNITY/Guar.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Guar.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,19 +1,13 @@
 (*  Title:      HOL/UNITY/Guar.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1999  University of Cambridge
+    Author:     Sidi Ehmety
 
 From Chandy and Sanders, "Reasoning About Program Composition",
 Technical Report 2000-003, University of Florida, 2000.
 
-Revised by Sidi Ehmety on January 2001
-
-Added: Compatibility, weakest guarantees, etc.
-
-and Weakest existential property,
-from Charpentier and Chandy "Theorems about Composition",
+Compatibility, weakest guarantees, etc.  and Weakest existential
+property, from Charpentier and Chandy "Theorems about Composition",
 Fifth International Conference on Mathematics of Program, 2000.
-
 *)
 
 header{*Guarantees Specifications*}
@@ -66,7 +60,7 @@
   "welldef == {F. Init F \<noteq> {}}"
   
   refines :: "['a program, 'a program, 'a program set] => bool"
-			("(3_ refines _ wrt _)" [10,10,10] 10)
+                        ("(3_ refines _ wrt _)" [10,10,10] 10)
   "G refines F wrt X ==
      \<forall>H. (F ok H & G ok H & F\<squnion>H \<in> welldef \<inter> X) --> 
          (G\<squnion>H \<in> welldef \<inter> X)"
--- a/src/HOL/UNITY/ListOrder.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/ListOrder.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      HOL/UNITY/ListOrder
-    ID:         $Id$
+(*  Title:      HOL/UNITY/ListOrder.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -9,8 +8,6 @@
      and corresponding elements of xs, xs' are pairwise related by r
 
 Also overloads <= and < for lists!
-
-Based on Lex/Prefix
 *)
 
 header {*The Prefix Ordering on Lists*}
@@ -26,7 +23,7 @@
    Nil:     "([],[]) : genPrefix(r)"
 
  | prepend: "[| (xs,ys) : genPrefix(r);  (x,y) : r |] ==>
-	     (x#xs, y#ys) : genPrefix(r)"
+             (x#xs, y#ys) : genPrefix(r)"
 
  | append:  "(xs,ys) : genPrefix(r) ==> (xs, ys@zs) : genPrefix(r)"
 
--- a/src/HOL/UNITY/PPROD.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/PPROD.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/PPROD.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -274,7 +273,7 @@
     done
 
     lemma const_PLam_Increasing:
-	 "[| i \<in> I;  finite I |]
+         "[| i \<in> I;  finite I |]
           ==> ((plam x \<in> I. F) \<in> Increasing (f o sub i)) = (F \<in> Increasing f)"
     apply (unfold Increasing_def)
     apply (subgoal_tac "\<forall>z. {s. z \<subseteq> (f o sub i) s} = lift_set i {s. z \<subseteq> f s}")
--- a/src/HOL/UNITY/ProgressSets.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/ProgressSets.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,4 +1,4 @@
-(*  Title:      HOL/UNITY/ProgressSets
+(*  Title:      HOL/UNITY/ProgressSets.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   2003  University of Cambridge
 
@@ -23,7 +23,7 @@
   lattice :: "'a set set => bool"
    --{*Meier calls them closure sets, but they are just complete lattices*}
    "lattice L ==
-	 (\<forall>M. M \<subseteq> L --> \<Inter>M \<in> L) & (\<forall>M. M \<subseteq> L --> \<Union>M \<in> L)"
+         (\<forall>M. M \<subseteq> L --> \<Inter>M \<in> L) & (\<forall>M. M \<subseteq> L --> \<Union>M \<in> L)"
 
   cl :: "['a set set, 'a set] => 'a set"
    --{*short for ``closure''*}
@@ -197,14 +197,14 @@
               [OF _ lattice_awp_lemma [OF TXC BsubX latt TC BC clos, of r]
                     latt])
 apply (subgoal_tac
-	 "T \<inter> (B \<union> wp act X) \<inter> (T \<inter> (B \<union> awp F (X \<union> cl C (T\<inter>r)))) = 
-	  T \<inter> (B \<union> wp act X \<inter> awp F (X \<union> cl C (T\<inter>r)))") 
+         "T \<inter> (B \<union> wp act X) \<inter> (T \<inter> (B \<union> awp F (X \<union> cl C (T\<inter>r)))) = 
+          T \<inter> (B \<union> wp act X \<inter> awp F (X \<union> cl C (T\<inter>r)))") 
  prefer 2 apply blast 
 apply simp  
 apply (drule Un_in_lattice [OF _ TXC latt])  
 apply (subgoal_tac
-	 "T \<inter> (B \<union> wp act X \<inter> awp F (X \<union> cl C (T\<inter>r))) \<union> T\<inter>X = 
-	  T \<inter> (wp act X \<inter> awp F (X \<union> cl C (T\<inter>r)) \<union> X)")
+         "T \<inter> (B \<union> wp act X \<inter> awp F (X \<union> cl C (T\<inter>r))) \<union> T\<inter>X = 
+          T \<inter> (wp act X \<inter> awp F (X \<union> cl C (T\<inter>r)) \<union> X)")
  apply simp 
 apply (blast intro: BsubX [THEN subsetD]) 
 done
@@ -408,7 +408,7 @@
 constdefs
   decoupled :: "['a program, 'a program] => bool"
    "decoupled F G ==
-	\<forall>act \<in> Acts F. \<forall>B. G \<in> stable B --> G \<in> stable (wp act B)"
+        \<forall>act \<in> Acts F. \<forall>B. G \<in> stable B --> G \<in> stable (wp act B)"
 
 
 text{*Rao's Decoupling Theorem*}
--- a/src/HOL/UNITY/Project.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Project.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,11 +1,10 @@
 (*  Title:      HOL/UNITY/Project.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1999  University of Cambridge
 
-Projections of state sets (also of actions and programs)
+Projections of state sets (also of actions and programs).
 
-Inheritance of GUARANTEES properties under extension
+Inheritance of GUARANTEES properties under extension.
 *)
 
 header{*Projections of State Sets*}
@@ -14,15 +13,15 @@
 
 constdefs
   projecting :: "['c program => 'c set, 'a*'b => 'c, 
-		  'a program, 'c program set, 'a program set] => bool"
+                  'a program, 'c program set, 'a program set] => bool"
     "projecting C h F X' X ==
        \<forall>G. extend h F\<squnion>G \<in> X' --> F\<squnion>project h (C G) G \<in> X"
 
   extending :: "['c program => 'c set, 'a*'b => 'c, 'a program, 
-		 'c program set, 'a program set] => bool"
+                 'c program set, 'a program set] => bool"
     "extending C h F Y' Y ==
        \<forall>G. extend h F  ok G --> F\<squnion>project h (C G) G \<in> Y
-	      --> extend h F\<squnion>G \<in> Y'"
+              --> extend h F\<squnion>G \<in> Y'"
 
   subset_closed :: "'a set set => bool"
     "subset_closed U == \<forall>A \<in> U. Pow A \<subseteq> U"  
@@ -566,8 +565,8 @@
      "project_act h (Restrict C act) \<subseteq> project_act h act"
 apply (auto simp add: project_act_def)
 done
-					   
-							   
+                                           
+                                                           
 lemma (in Extend) subset_closed_ok_extend_imp_ok_project:
      "[| extend h F ok G; subset_closed (AllowedActs F) |]  
       ==> F ok project h C G"
--- a/src/HOL/UNITY/Simple/Deadlock.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Simple/Deadlock.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      HOL/UNITY/Deadlock
-    ID:         $Id$
+(*  Title:      HOL/UNITY/Deadlock.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -63,7 +62,7 @@
 lemma
   assumes zeroprem: "F \<in> (A 0 \<inter> A (Suc n)) co (A 0)"
       and allprem:
-	    "!!i. i \<in> atMost n ==> F \<in> (A(Suc i) \<inter> A i) co (-A i \<union> A(Suc i))"
+            "!!i. i \<in> atMost n ==> F \<in> (A(Suc i) \<inter> A i) co (-A i \<union> A(Suc i))"
   shows "F \<in> stable (\<Inter>i \<in> atMost (Suc n). A i)"
 apply (unfold stable_def) 
 apply (rule constrains_Int [THEN constrains_weaken])
--- a/src/HOL/UNITY/Simple/Lift.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Simple/Lift.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,8 @@
 (*  Title:      HOL/UNITY/Lift.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
-The Lift-Control Example
+The Lift-Control Example.
 *)
 
 theory Lift
@@ -12,12 +11,12 @@
 begin
 
 record state =
-  floor :: "int"	    --{*current position of the lift*}
-  "open" :: "bool"	    --{*whether the door is opened at floor*}
-  stop  :: "bool"	    --{*whether the lift is stopped at floor*}
-  req   :: "int set"	    --{*for each floor, whether the lift is requested*}
-  up    :: "bool"	    --{*current direction of movement*}
-  move  :: "bool"	    --{*whether moving takes precedence over opening*}
+  floor :: "int"            --{*current position of the lift*}
+  "open" :: "bool"          --{*whether the door is opened at floor*}
+  stop  :: "bool"           --{*whether the lift is stopped at floor*}
+  req   :: "int set"        --{*for each floor, whether the lift is requested*}
+  up    :: "bool"           --{*current direction of movement*}
+  move  :: "bool"           --{*whether moving takes precedence over opening*}
 
 consts
   Min :: "int"       --{*least and greatest floors*}
@@ -74,15 +73,15 @@
   
   request_act :: "(state*state) set"
     "request_act == {(s,s'). s' = s (|stop:=True, move:=False|)
-		                  & ~ stop s & floor s \<in> req s}"
+                                  & ~ stop s & floor s \<in> req s}"
 
   open_act :: "(state*state) set"
     "open_act ==
          {(s,s'). s' = s (|open :=True,
-			   req  := req s - {floor s},
-			   move := True|)
-		       & stop s & ~ open s & floor s \<in> req s
-	               & ~(move s & s \<in> queueing)}"
+                           req  := req s - {floor s},
+                           move := True|)
+                       & stop s & ~ open s & floor s \<in> req s
+                       & ~(move s & s \<in> queueing)}"
 
   close_act :: "(state*state) set"
     "close_act == {(s,s'). s' = s (|open := False|) & open s}"
@@ -90,47 +89,47 @@
   req_up :: "(state*state) set"
     "req_up ==
          {(s,s'). s' = s (|stop  :=False,
-			   floor := floor s + 1,
-			   up    := True|)
-		       & s \<in> (ready \<inter> goingup)}"
+                           floor := floor s + 1,
+                           up    := True|)
+                       & s \<in> (ready \<inter> goingup)}"
 
   req_down :: "(state*state) set"
     "req_down ==
          {(s,s'). s' = s (|stop  :=False,
-			   floor := floor s - 1,
-			   up    := False|)
-		       & s \<in> (ready \<inter> goingdown)}"
+                           floor := floor s - 1,
+                           up    := False|)
+                       & s \<in> (ready \<inter> goingdown)}"
 
   move_up :: "(state*state) set"
     "move_up ==
          {(s,s'). s' = s (|floor := floor s + 1|)
-		       & ~ stop s & up s & floor s \<notin> req s}"
+                       & ~ stop s & up s & floor s \<notin> req s}"
 
   move_down :: "(state*state) set"
     "move_down ==
          {(s,s'). s' = s (|floor := floor s - 1|)
-		       & ~ stop s & ~ up s & floor s \<notin> req s}"
+                       & ~ stop s & ~ up s & floor s \<notin> req s}"
 
   button_press  :: "(state*state) set"
       --{*This action is omitted from prior treatments, which therefore are
-	unrealistic: nobody asks the lift to do anything!  But adding this
-	action invalidates many of the existing progress arguments: various
-	"ensures" properties fail. Maybe it should be constrained to only
+        unrealistic: nobody asks the lift to do anything!  But adding this
+        action invalidates many of the existing progress arguments: various
+        "ensures" properties fail. Maybe it should be constrained to only
         allow button presses in the current direction of travel, like in a
         real lift.*}
     "button_press ==
          {(s,s'). \<exists>n. s' = s (|req := insert n (req s)|)
-		        & Min \<le> n & n \<le> Max}"
+                        & Min \<le> n & n \<le> Max}"
 
 
   Lift :: "state program"
     --{*for the moment, we OMIT @{text button_press}*}
     "Lift == mk_total_program
                 ({s. floor s = Min & ~ up s & move s & stop s &
-		          ~ open s & req s = {}},
-		 {request_act, open_act, close_act,
- 		  req_up, req_down, move_up, move_down},
-		 UNIV)"
+                          ~ open s & req s = {}},
+                 {request_act, open_act, close_act,
+                  req_up, req_down, move_up, move_down},
+                 UNIV)"
 
 
   --{*Invariants*}
@@ -158,10 +157,10 @@
   metric :: "[int,state] => int"
     "metric ==
        %n s. if floor s < n then (if up s then n - floor s
-			          else (floor s - Min) + (n-Min))
+                                  else (floor s - Min) + (n-Min))
              else
              if n < floor s then (if up s then (Max - floor s) + (Max-n)
-		                  else floor s - n)
+                                  else floor s - n)
              else 0"
 
 locale Floor =
--- a/src/HOL/UNITY/Simple/Mutex.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Simple/Mutex.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,8 @@
 (*  Title:      HOL/UNITY/Mutex.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
-Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra
+Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra.
 *)
 
 theory Mutex imports "../UNITY_Main" begin
@@ -56,7 +55,7 @@
   Mutex :: "state program"
     "Mutex == mk_total_program
                  ({s. ~ u s & ~ v s & m s = 0 & n s = 0},
-		  {U0, U1, U2, U3, U4, V0, V1, V2, V3, V4},
+                  {U0, U1, U2, U3, U4, V0, V1, V2, V3, V4},
                   UNIV)"
 
 
@@ -72,7 +71,7 @@
 
   bad_IU :: "state set"
     "bad_IU == {s. (u s = (1 \<le> m s & m s \<le> 3)) &
-	           (3 \<le> m s & m s \<le> 4 --> ~ p s)}"
+                   (3 \<le> m s & m s \<le> 4 --> ~ p s)}"
 
 
 declare Mutex_def [THEN def_prg_Init, simp]
--- a/src/HOL/UNITY/Simple/NSP_Bad.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Simple/NSP_Bad.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      HOL/Auth/NSP_Bad
-    ID:         $Id$
+(*  Title:      HOL/Auth/NSP_Bad.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1996  University of Cambridge
 
@@ -25,33 +24,33 @@
     all similar protocols.*)
   Fake :: "(state*state) set"
     "Fake == {(s,s').
-	      \<exists>B X. s' = Says Spy B X # s
-		    & X \<in> synth (analz (spies s))}"
+              \<exists>B X. s' = Says Spy B X # s
+                    & X \<in> synth (analz (spies s))}"
 
   (*The numeric suffixes on A identify the rule*)
 
   (*Alice initiates a protocol run, sending a nonce to Bob*)
   NS1 :: "(state*state) set"
     "NS1 == {(s1,s').
-	     \<exists>A1 B NA.
-	         s' = Says A1 B (Crypt (pubK B) {|Nonce NA, Agent A1|}) # s1
-	       & Nonce NA \<notin> used s1}"
+             \<exists>A1 B NA.
+                 s' = Says A1 B (Crypt (pubK B) {|Nonce NA, Agent A1|}) # s1
+               & Nonce NA \<notin> used s1}"
 
   (*Bob responds to Alice's message with a further nonce*)
   NS2 :: "(state*state) set"
     "NS2 == {(s2,s').
-	     \<exists>A' A2 B NA NB.
-	         s' = Says B A2 (Crypt (pubK A2) {|Nonce NA, Nonce NB|}) # s2
+             \<exists>A' A2 B NA NB.
+                 s' = Says B A2 (Crypt (pubK A2) {|Nonce NA, Nonce NB|}) # s2
                & Says A' B (Crypt (pubK B) {|Nonce NA, Agent A2|}) \<in> set s2
-	       & Nonce NB \<notin> used s2}"
+               & Nonce NB \<notin> used s2}"
 
   (*Alice proves her existence by sending NB back to Bob.*)
   NS3 :: "(state*state) set"
     "NS3 == {(s3,s').
-	     \<exists>A3 B' B NA NB.
-	         s' = Says A3 B (Crypt (pubK B) (Nonce NB)) # s3
+             \<exists>A3 B' B NA NB.
+                 s' = Says A3 B (Crypt (pubK B) (Nonce NB)) # s3
                & Says A3  B (Crypt (pubK B) {|Nonce NA, Agent A3|}) \<in> set s3
-	       & Says B' A3 (Crypt (pubK A3) {|Nonce NA, Nonce NB|}) \<in> set s3}"
+               & Says B' A3 (Crypt (pubK A3) {|Nonce NA, Nonce NB|}) \<in> set s3}"
 
 
 constdefs
@@ -104,14 +103,14 @@
 fun ns_constrains_tac(cs,ss) i =
    SELECT_GOAL
       (EVERY [REPEAT (etac @{thm Always_ConstrainsI} 1),
-	      REPEAT (resolve_tac [@{thm StableI}, @{thm stableI},
-				   @{thm constrains_imp_Constrains}] 1),
-	      rtac @{thm ns_constrainsI} 1,
-	      full_simp_tac ss 1,
-	      REPEAT (FIRSTGOAL (etac disjE)),
-	      ALLGOALS (clarify_tac (cs delrules [impI, @{thm impCE}])),
-	      REPEAT (FIRSTGOAL analz_mono_contra_tac),
-	      ALLGOALS (asm_simp_tac ss)]) i;
+              REPEAT (resolve_tac [@{thm StableI}, @{thm stableI},
+                                   @{thm constrains_imp_Constrains}] 1),
+              rtac @{thm ns_constrainsI} 1,
+              full_simp_tac ss 1,
+              REPEAT (FIRSTGOAL (etac disjE)),
+              ALLGOALS (clarify_tac (cs delrules [impI, @{thm impCE}])),
+              REPEAT (FIRSTGOAL analz_mono_contra_tac),
+              ALLGOALS (asm_simp_tac ss)]) i;
 
 (*Tactic for proving secrecy theorems*)
 fun ns_induct_tac(cs,ss) =
--- a/src/HOL/UNITY/Simple/Network.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Simple/Network.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,11 +1,10 @@
-(*  Title:      HOL/UNITY/Network
-    ID:         $Id$
+(*  Title:      HOL/UNITY/Network.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
-The Communication Network
+The Communication Network.
 
-From Misra, "A Logic for Concurrent Programming" (1994), section 5.7
+From Misra, "A Logic for Concurrent Programming" (1994), section 5.7.
 *)
 
 theory Network imports UNITY begin
@@ -52,9 +51,9 @@
 
 lemma (in F_props)
   shows "F \<in> stable {s. s(Aproc,Idle) = Suc 0 & s(Bproc,Idle) = Suc 0 &  
-			s(Aproc,Sent) = s(Bproc,Rcvd) &  
-			s(Bproc,Sent) = s(Aproc,Rcvd) &  
-			s(Aproc,Rcvd) = m & s(Bproc,Rcvd) = n}"
+                        s(Aproc,Sent) = s(Bproc,Rcvd) &  
+                        s(Bproc,Sent) = s(Aproc,Rcvd) &  
+                        s(Aproc,Rcvd) = m & s(Bproc,Rcvd) = n}"
 apply (unfold stable_def) 
 apply (rule constrainsI)
 apply (drule constrains_Int [OF rs_AB [unfolded stable_def] nondec_idle, 
--- a/src/HOL/UNITY/Simple/Reach.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Simple/Reach.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,10 +1,9 @@
 (*  Title:      HOL/UNITY/Reach.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
 Reachability in Directed Graphs.  From Chandy and Misra, section 6.4.
-	[this example took only four days!]
+[this example took only four days!]
 *)
 
 theory Reach imports "../UNITY_Main" begin
--- a/src/HOL/UNITY/Simple/Reachability.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Simple/Reachability.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,11 +1,11 @@
-(*  Title:      HOL/UNITY/Reachability
-    ID:         $Id$
+(*  Title:      HOL/UNITY/Reachability.thy
     Author:     Tanja Vos, Cambridge University Computer Laboratory
     Copyright   2000  University of Cambridge
 
-Reachability in Graphs
+Reachability in Graphs.
 
-From Chandy and Misra, "Parallel Program Design" (1989), sections 6.2 and 11.3
+From Chandy and Misra, "Parallel Program Design" (1989), sections 6.2
+and 11.3.
 *)
 
 theory Reachability imports "../Detects" Reach begin
@@ -300,7 +300,7 @@
              ((reachable v <==> {s. (root,v): REACHABLE}) \<inter> nmsg_eq 0 (v,w))"
 apply (rule LeadsTo_Reachability [THEN LeadsTo_Trans], blast)
 apply (subgoal_tac
-	 "F \<in> (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> 
+         "F \<in> (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> 
               UNIV LeadsTo (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> 
               nmsg_eq 0 (v,w) ")
 apply simp
--- a/src/HOL/UNITY/UNITY.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/UNITY.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,11 +1,11 @@
-(*  Title:      HOL/UNITY/UNITY
-    ID:         $Id$
+(*  Title:      HOL/UNITY/UNITY.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
-The basic UNITY theory (revised version, based upon the "co" operator)
+The basic UNITY theory (revised version, based upon the "co"
+operator).
 
-From Misra, "A Logic for Concurrent Programming", 1994
+From Misra, "A Logic for Concurrent Programming", 1994.
 *)
 
 header {*The Basic UNITY Theory*}
@@ -14,7 +14,7 @@
 
 typedef (Program)
   'a program = "{(init:: 'a set, acts :: ('a * 'a)set set,
-		   allowed :: ('a * 'a)set set). Id \<in> acts & Id: allowed}" 
+                   allowed :: ('a * 'a)set set). Id \<in> acts & Id: allowed}" 
   by blast
 
 constdefs
@@ -28,7 +28,7 @@
     "A unless B == (A-B) co (A \<union> B)"
 
   mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set)
-		   => 'a program"
+                   => 'a program"
     "mk_program == %(init, acts, allowed).
                       Abs_Program (init, insert Id acts, insert Id allowed)"
 
@@ -363,11 +363,11 @@
 
   totalize :: "'a program => 'a program"
     "totalize F == mk_program (Init F,
-			       totalize_act ` Acts F,
-			       AllowedActs F)"
+                               totalize_act ` Acts F,
+                               AllowedActs F)"
 
   mk_total_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set)
-		   => 'a program"
+                   => 'a program"
     "mk_total_program args == totalize (mk_program args)"
 
   all_total :: "'a program => bool"
--- a/src/HOL/UNITY/Union.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/UNITY/Union.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,8 @@
 (*  Title:      HOL/UNITY/Union.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
-Partly from Misra's Chapter 5: Asynchronous Compositions of Programs
+Partly from Misra's Chapter 5: Asynchronous Compositions of Programs.
 *)
 
 header{*Unions of Programs*}
@@ -23,11 +22,11 @@
 
   JOIN  :: "['a set, 'a => 'b program] => 'b program"
     "JOIN I F == mk_program (\<Inter>i \<in> I. Init (F i), \<Union>i \<in> I. Acts (F i),
-			     \<Inter>i \<in> I. AllowedActs (F i))"
+                             \<Inter>i \<in> I. AllowedActs (F i))"
 
   Join :: "['a program, 'a program] => 'a program"      (infixl "Join" 65)
     "F Join G == mk_program (Init F \<inter> Init G, Acts F \<union> Acts G,
-			     AllowedActs F \<inter> AllowedActs G)"
+                             AllowedActs F \<inter> AllowedActs G)"
 
   SKIP :: "'a program"
     "SKIP == mk_program (UNIV, {}, UNIV)"
@@ -406,7 +405,7 @@
 lemma safety_prop_INTER1 [simp]:
      "(!!i. safety_prop (X i)) ==> safety_prop (\<Inter>i. X i)"
 by (auto simp add: safety_prop_def, blast)
-							       
+
 lemma safety_prop_INTER [simp]:
      "(!!i. i \<in> I ==> safety_prop (X i)) ==> safety_prop (\<Inter>i \<in> I. X i)"
 by (auto simp add: safety_prop_def, blast)
--- a/src/HOL/Wellfounded.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/Wellfounded.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,7 +1,8 @@
-(*  Author:     Tobias Nipkow
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Author:     Konrad Slind, Alexander Krauss
-    Copyright   1992-2008  University of Cambridge and TU Muenchen
+(*  Title:      HOL/Wellfounded.thy
+    Author:     Tobias Nipkow
+    Author:     Lawrence C Paulson
+    Author:     Konrad Slind
+    Author:     Alexander Krauss
 *)
 
 header {*Well-founded Recursion*}
@@ -94,21 +95,21 @@
       fix y assume "(y, x) : r^+"
       with `wf r` show "P y"
       proof (induct x arbitrary: y)
-	case (less x)
-	note hyp = `\<And>x' y'. (x', x) : r ==> (y', x') : r^+ ==> P y'`
-	from `(y, x) : r^+` show "P y"
-	proof cases
-	  case base
-	  show "P y"
-	  proof (rule induct_step)
-	    fix y' assume "(y', y) : r^+"
-	    with `(y, x) : r` show "P y'" by (rule hyp [of y y'])
-	  qed
-	next
-	  case step
-	  then obtain x' where "(x', x) : r" and "(y, x') : r^+" by simp
-	  then show "P y" by (rule hyp [of x' y])
-	qed
+        case (less x)
+        note hyp = `\<And>x' y'. (x', x) : r ==> (y', x') : r^+ ==> P y'`
+        from `(y, x) : r^+` show "P y"
+        proof cases
+          case base
+          show "P y"
+          proof (rule induct_step)
+            fix y' assume "(y', y) : r^+"
+            with `(y, x) : r` show "P y'" by (rule hyp [of y y'])
+          qed
+        next
+          case step
+          then obtain x' where "(x', x) : r" and "(y, x') : r^+" by simp
+          then show "P y" by (rule hyp [of x' y])
+        qed
       qed
     qed
   } then show ?thesis unfolding wf_def by blast
--- a/src/HOL/ZF/Games.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ZF/Games.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,8 @@
 (*  Title:      HOL/ZF/Games.thy
-    ID:         $Id$
     Author:     Steven Obua
 
-    An application of HOLZF: Partizan Games.
-    See "Partizan Games in Isabelle/HOLZF", available from http://www4.in.tum.de/~obua/partizan
+An application of HOLZF: Partizan Games.  See "Partizan Games in
+Isabelle/HOLZF", available from http://www4.in.tum.de/~obua/partizan
 *)
 
 theory Games 
@@ -421,23 +420,23 @@
       from goal1 have "(y, g) \<in> option_of" by (auto)
       with 1 have "ge_game (y, y)" by auto
       with goal1 have "\<not> ge_game (g, y)" 
-	by (subst ge_game_eq, auto)
+        by (subst ge_game_eq, auto)
       with goal1 show ?case by auto}
     note right = this
     {case (goal2 y)
       from goal2 have "(y, g) \<in> option_of" by (auto)
       with 1 have "ge_game (y, y)" by auto
       with goal2 have "\<not> ge_game (y, g)" 
-	by (subst ge_game_eq, auto)
+        by (subst ge_game_eq, auto)
       with goal2 show ?case by auto}
     note left = this
     {case goal3
       from left right show ?case
-	by (subst ge_game_eq, auto)
+        by (subst ge_game_eq, auto)
     }
   qed
 qed
-	
+        
 constdefs
   eq_game :: "game \<Rightarrow> game \<Rightarrow> bool"
   "eq_game G H \<equiv> ge_game (G, H) \<and> ge_game (H, G)" 
@@ -462,36 +461,36 @@
       case (1 a)
       show ?case
       proof (rule allI | rule impI)+
-	case (goal1 x y z)
-	show ?case
-	proof -
-	  { fix xr
+        case (goal1 x y z)
+        show ?case
+        proof -
+          { fix xr
             assume xr:"zin xr (right_options x)"
-	    assume "ge_game (z, xr)"
-	    have "ge_game (y, xr)"
-	      apply (rule 1[rule_format, where y="[y,z,xr]"])
-	      apply (auto intro: xr lprod_3_1 simp add: prems)
-	      done
-	    moreover from xr have "\<not> ge_game (y, xr)"
-	      by (simp add: goal1(2)[simplified ge_game_eq[of x y], rule_format, of xr, simplified xr])
-	    ultimately have "False" by auto      
-	  }
-	  note xr = this
-	  { fix zl
-	    assume zl:"zin zl (left_options z)"
-	    assume "ge_game (zl, x)"
-	    have "ge_game (zl, y)"
-	      apply (rule 1[rule_format, where y="[zl,x,y]"])
-	      apply (auto intro: zl lprod_3_2 simp add: prems)
-	      done
-	    moreover from zl have "\<not> ge_game (zl, y)"
-	      by (simp add: goal1(3)[simplified ge_game_eq[of y z], rule_format, of zl, simplified zl])
-	    ultimately have "False" by auto
-	  }
-	  note zl = this
-	  show ?thesis
-	    by (auto simp add: ge_game_eq[of x z] intro: xr zl)
-	qed
+            assume "ge_game (z, xr)"
+            have "ge_game (y, xr)"
+              apply (rule 1[rule_format, where y="[y,z,xr]"])
+              apply (auto intro: xr lprod_3_1 simp add: prems)
+              done
+            moreover from xr have "\<not> ge_game (y, xr)"
+              by (simp add: goal1(2)[simplified ge_game_eq[of x y], rule_format, of xr, simplified xr])
+            ultimately have "False" by auto      
+          }
+          note xr = this
+          { fix zl
+            assume zl:"zin zl (left_options z)"
+            assume "ge_game (zl, x)"
+            have "ge_game (zl, y)"
+              apply (rule 1[rule_format, where y="[zl,x,y]"])
+              apply (auto intro: zl lprod_3_2 simp add: prems)
+              done
+            moreover from zl have "\<not> ge_game (zl, y)"
+              by (simp add: goal1(3)[simplified ge_game_eq[of y z], rule_format, of zl, simplified zl])
+            ultimately have "False" by auto
+          }
+          note zl = this
+          show ?thesis
+            by (auto simp add: ge_game_eq[of x z] intro: xr zl)
+        qed
       qed
     qed
   } 
@@ -550,12 +549,12 @@
       case (goal1 G H)
       note induct_hyp = prems[simplified goal1, simplified] and prems
       show ?case
-	apply (simp only: plus_game.simps[where G=G and H=H])
-	apply (simp add: game_ext_eq prems)
-	apply (auto simp add: 
-	  zimage_cong[where f = "\<lambda> g. plus_game (g, zero_game)" and g = "id"] 
-	  induct_hyp)
-	done
+        apply (simp only: plus_game.simps[where G=G and H=H])
+        apply (simp add: game_ext_eq prems)
+        apply (auto simp add: 
+          zimage_cong[where f = "\<lambda> g. plus_game (g, zero_game)" and g = "id"] 
+          induct_hyp)
+        done
     qed
   }
   then show ?thesis by auto
@@ -578,7 +577,7 @@
   "right_options (plus_game (u, v)) =  zunion (zimage (\<lambda>g. plus_game (g, v)) (right_options u)) (zimage (\<lambda>h. plus_game (u, h)) (right_options v))"
   by (subst plus_game.simps, simp)
 
-lemma left_options_neg: "left_options (neg_game u) = zimage neg_game (right_options u)"	 
+lemma left_options_neg: "left_options (neg_game u) = zimage neg_game (right_options u)"  
   by (subst neg_game.simps, simp)
 
 lemma right_options_neg: "right_options (neg_game u) = zimage neg_game (left_options u)"
@@ -595,25 +594,25 @@
       let ?R = "plus_game (F, plus_game (G, H))"
       note options_plus = left_options_plus right_options_plus
       {
-	fix opt
-	note hyp = goal1(1)[simplified goal1(2), rule_format] 
-	have F: "zin opt (options F)  \<Longrightarrow> plus_game (plus_game (opt, G), H) = plus_game (opt, plus_game (G, H))"
-	  by (blast intro: hyp lprod_3_3)
-	have G: "zin opt (options G) \<Longrightarrow> plus_game (plus_game (F, opt), H) = plus_game (F, plus_game (opt, H))"
-	  by (blast intro: hyp lprod_3_4)
-	have H: "zin opt (options H) \<Longrightarrow> plus_game (plus_game (F, G), opt) = plus_game (F, plus_game (G, opt))" 
-	  by (blast intro: hyp lprod_3_5)
-	note F and G and H
+        fix opt
+        note hyp = goal1(1)[simplified goal1(2), rule_format] 
+        have F: "zin opt (options F)  \<Longrightarrow> plus_game (plus_game (opt, G), H) = plus_game (opt, plus_game (G, H))"
+          by (blast intro: hyp lprod_3_3)
+        have G: "zin opt (options G) \<Longrightarrow> plus_game (plus_game (F, opt), H) = plus_game (F, plus_game (opt, H))"
+          by (blast intro: hyp lprod_3_4)
+        have H: "zin opt (options H) \<Longrightarrow> plus_game (plus_game (F, G), opt) = plus_game (F, plus_game (G, opt))" 
+          by (blast intro: hyp lprod_3_5)
+        note F and G and H
       }
       note induct_hyp = this
       have "left_options ?L = left_options ?R \<and> right_options ?L = right_options ?R"
-	by (auto simp add: 
-	  plus_game.simps[where G="plus_game (F,G)" and H=H]
-	  plus_game.simps[where G="F" and H="plus_game (G,H)"] 
-	  zet_ext_eq zunion zimage_iff options_plus
-	  induct_hyp left_imp_options right_imp_options)
+        by (auto simp add: 
+          plus_game.simps[where G="plus_game (F,G)" and H=H]
+          plus_game.simps[where G="F" and H="plus_game (G,H)"] 
+          zet_ext_eq zunion zimage_iff options_plus
+          induct_hyp left_imp_options right_imp_options)
       then show ?case
-	by (simp add: game_ext_eq)
+        by (simp add: game_ext_eq)
     qed
   }
   then show ?thesis by auto
@@ -702,99 +701,99 @@
       case (goal1 a x y z)
       note induct_hyp = goal1(1)[rule_format, simplified goal1(2)]
       { 
-	assume hyp: "ge_game(plus_game (x, y), plus_game (x, z))"
-	have "ge_game (y, z)"
-	proof -
-	  { fix yr
-	    assume yr: "zin yr (right_options y)"
-	    from hyp have "\<not> (ge_game (plus_game (x, z), plus_game (x, yr)))"
-	      by (auto simp add: ge_game_eq[of "plus_game (x,y)" "plus_game(x,z)"]
-		right_options_plus zunion zimage_iff intro: yr)
-	    then have "\<not> (ge_game (z, yr))"
-	      apply (subst induct_hyp[where y="[x, z, yr]", of "x" "z" "yr"])
-	      apply (simp_all add: yr lprod_3_6)
-	      done
-	  }
-	  note yr = this
-	  { fix zl
-	    assume zl: "zin zl (left_options z)"
-	    from hyp have "\<not> (ge_game (plus_game (x, zl), plus_game (x, y)))"
-	      by (auto simp add: ge_game_eq[of "plus_game (x,y)" "plus_game(x,z)"]
-		left_options_plus zunion zimage_iff intro: zl)
-	    then have "\<not> (ge_game (zl, y))"
-	      apply (subst goal1(1)[rule_format, where y="[x, zl, y]", of "x" "zl" "y"])
-	      apply (simp_all add: goal1(2) zl lprod_3_7)
-	      done
-	  }	
-	  note zl = this
-	  show "ge_game (y, z)"
-	    apply (subst ge_game_eq)
-	    apply (auto simp add: yr zl)
-	    done
-	qed      
+        assume hyp: "ge_game(plus_game (x, y), plus_game (x, z))"
+        have "ge_game (y, z)"
+        proof -
+          { fix yr
+            assume yr: "zin yr (right_options y)"
+            from hyp have "\<not> (ge_game (plus_game (x, z), plus_game (x, yr)))"
+              by (auto simp add: ge_game_eq[of "plus_game (x,y)" "plus_game(x,z)"]
+                right_options_plus zunion zimage_iff intro: yr)
+            then have "\<not> (ge_game (z, yr))"
+              apply (subst induct_hyp[where y="[x, z, yr]", of "x" "z" "yr"])
+              apply (simp_all add: yr lprod_3_6)
+              done
+          }
+          note yr = this
+          { fix zl
+            assume zl: "zin zl (left_options z)"
+            from hyp have "\<not> (ge_game (plus_game (x, zl), plus_game (x, y)))"
+              by (auto simp add: ge_game_eq[of "plus_game (x,y)" "plus_game(x,z)"]
+                left_options_plus zunion zimage_iff intro: zl)
+            then have "\<not> (ge_game (zl, y))"
+              apply (subst goal1(1)[rule_format, where y="[x, zl, y]", of "x" "zl" "y"])
+              apply (simp_all add: goal1(2) zl lprod_3_7)
+              done
+          }     
+          note zl = this
+          show "ge_game (y, z)"
+            apply (subst ge_game_eq)
+            apply (auto simp add: yr zl)
+            done
+        qed      
       }
       note right_imp_left = this
       {
-	assume yz: "ge_game (y, z)"
-	{
-	  fix x'
-	  assume x': "zin x' (right_options x)"
-	  assume hyp: "ge_game (plus_game (x, z), plus_game (x', y))"
-	  then have n: "\<not> (ge_game (plus_game (x', y), plus_game (x', z)))"
-	    by (auto simp add: ge_game_eq[of "plus_game (x,z)" "plus_game (x', y)"] 
-	      right_options_plus zunion zimage_iff intro: x')
-	  have t: "ge_game (plus_game (x', y), plus_game (x', z))"
-	    apply (subst induct_hyp[symmetric])
-	    apply (auto intro: lprod_3_3 x' yz)
-	    done
-	  from n t have "False" by blast
-	}    
-	note case1 = this
-	{
-	  fix x'
-	  assume x': "zin x' (left_options x)"
-	  assume hyp: "ge_game (plus_game (x', z), plus_game (x, y))"
-	  then have n: "\<not> (ge_game (plus_game (x', y), plus_game (x', z)))"
-	    by (auto simp add: ge_game_eq[of "plus_game (x',z)" "plus_game (x, y)"] 
-	      left_options_plus zunion zimage_iff intro: x')
-	  have t: "ge_game (plus_game (x', y), plus_game (x', z))"
-	    apply (subst induct_hyp[symmetric])
-	    apply (auto intro: lprod_3_3 x' yz)
-	    done
-	  from n t have "False" by blast
-	}
-	note case3 = this
-	{
-	  fix y'
-	  assume y': "zin y' (right_options y)"
-	  assume hyp: "ge_game (plus_game(x, z), plus_game (x, y'))"
-	  then have "ge_game(z, y')"
-	    apply (subst induct_hyp[of "[x, z, y']" "x" "z" "y'"])
-	    apply (auto simp add: hyp lprod_3_6 y')
-	    done
-	  with yz have "ge_game (y, y')"
-	    by (blast intro: ge_game_trans)      
-	  with y' have "False" by (auto simp add: ge_game_leftright_refl)
-	}
-	note case2 = this
-	{
-	  fix z'
-	  assume z': "zin z' (left_options z)"
-	  assume hyp: "ge_game (plus_game(x, z'), plus_game (x, y))"
-	  then have "ge_game(z', y)"
-	    apply (subst induct_hyp[of "[x, z', y]" "x" "z'" "y"])
-	    apply (auto simp add: hyp lprod_3_7 z')
-	    done    
-	  with yz have "ge_game (z', z)"
-	    by (blast intro: ge_game_trans)      
-	  with z' have "False" by (auto simp add: ge_game_leftright_refl)
-	}
-	note case4 = this   
-	have "ge_game(plus_game (x, y), plus_game (x, z))"
-	  apply (subst ge_game_eq)
-	  apply (auto simp add: right_options_plus left_options_plus zunion zimage_iff)
-	  apply (auto intro: case1 case2 case3 case4)
-	  done
+        assume yz: "ge_game (y, z)"
+        {
+          fix x'
+          assume x': "zin x' (right_options x)"
+          assume hyp: "ge_game (plus_game (x, z), plus_game (x', y))"
+          then have n: "\<not> (ge_game (plus_game (x', y), plus_game (x', z)))"
+            by (auto simp add: ge_game_eq[of "plus_game (x,z)" "plus_game (x', y)"] 
+              right_options_plus zunion zimage_iff intro: x')
+          have t: "ge_game (plus_game (x', y), plus_game (x', z))"
+            apply (subst induct_hyp[symmetric])
+            apply (auto intro: lprod_3_3 x' yz)
+            done
+          from n t have "False" by blast
+        }    
+        note case1 = this
+        {
+          fix x'
+          assume x': "zin x' (left_options x)"
+          assume hyp: "ge_game (plus_game (x', z), plus_game (x, y))"
+          then have n: "\<not> (ge_game (plus_game (x', y), plus_game (x', z)))"
+            by (auto simp add: ge_game_eq[of "plus_game (x',z)" "plus_game (x, y)"] 
+              left_options_plus zunion zimage_iff intro: x')
+          have t: "ge_game (plus_game (x', y), plus_game (x', z))"
+            apply (subst induct_hyp[symmetric])
+            apply (auto intro: lprod_3_3 x' yz)
+            done
+          from n t have "False" by blast
+        }
+        note case3 = this
+        {
+          fix y'
+          assume y': "zin y' (right_options y)"
+          assume hyp: "ge_game (plus_game(x, z), plus_game (x, y'))"
+          then have "ge_game(z, y')"
+            apply (subst induct_hyp[of "[x, z, y']" "x" "z" "y'"])
+            apply (auto simp add: hyp lprod_3_6 y')
+            done
+          with yz have "ge_game (y, y')"
+            by (blast intro: ge_game_trans)      
+          with y' have "False" by (auto simp add: ge_game_leftright_refl)
+        }
+        note case2 = this
+        {
+          fix z'
+          assume z': "zin z' (left_options z)"
+          assume hyp: "ge_game (plus_game(x, z'), plus_game (x, y))"
+          then have "ge_game(z', y)"
+            apply (subst induct_hyp[of "[x, z', y]" "x" "z'" "y"])
+            apply (auto simp add: hyp lprod_3_7 z')
+            done    
+          with yz have "ge_game (z', z)"
+            by (blast intro: ge_game_trans)      
+          with z' have "False" by (auto simp add: ge_game_leftright_refl)
+        }
+        note case4 = this   
+        have "ge_game(plus_game (x, y), plus_game (x, z))"
+          apply (subst ge_game_eq)
+          apply (auto simp add: right_options_plus left_options_plus zunion zimage_iff)
+          apply (auto intro: case1 case2 case3 case4)
+          done
       }
       note left_imp_right = this
       show ?case by (auto intro: right_imp_left left_imp_right)
@@ -815,24 +814,24 @@
       case (goal1 a x y)
       note ihyp = goal1(1)[rule_format, simplified goal1(2)]
       { fix xl
-	assume xl: "zin xl (left_options x)"
-	have "ge_game (neg_game y, neg_game xl) = ge_game (xl, y)"
-	  apply (subst ihyp)
-	  apply (auto simp add: lprod_2_1 xl)
-	  done
+        assume xl: "zin xl (left_options x)"
+        have "ge_game (neg_game y, neg_game xl) = ge_game (xl, y)"
+          apply (subst ihyp)
+          apply (auto simp add: lprod_2_1 xl)
+          done
       }
       note xl = this
       { fix yr
-	assume yr: "zin yr (right_options y)"
-	have "ge_game (neg_game yr, neg_game x) = ge_game (x, yr)"
-	  apply (subst ihyp)
-	  apply (auto simp add: lprod_2_2 yr)
-	  done
+        assume yr: "zin yr (right_options y)"
+        have "ge_game (neg_game yr, neg_game x) = ge_game (x, yr)"
+          apply (subst ihyp)
+          apply (auto simp add: lprod_2_2 yr)
+          done
       }
       note yr = this
       show ?case
-	by (auto simp add: ge_game_eq[of "neg_game x" "neg_game y"] ge_game_eq[of "y" "x"]
-	  right_options_neg left_options_neg zimage_iff  xl yr)
+        by (auto simp add: ge_game_eq[of "neg_game x" "neg_game y"] ge_game_eq[of "y" "x"]
+          right_options_neg left_options_neg zimage_iff  xl yr)
     qed
   }
   note a = this[of "[x,y]"]
--- a/src/HOL/ZF/HOLZF.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ZF/HOLZF.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,8 @@
 (*  Title:      HOL/ZF/HOLZF.thy
-    ID:         $Id$
     Author:     Steven Obua
 
-    Axiomatizes the ZFC universe as an HOL type.
-    See "Partizan Games in Isabelle/HOLZF", available from http://www4.in.tum.de/~obua/partizan
+Axiomatizes the ZFC universe as an HOL type.  See "Partizan Games in
+Isabelle/HOLZF", available from http://www4.in.tum.de/~obua/partizan
 *)
 
 theory HOLZF 
@@ -360,25 +359,25 @@
     assume a_in_U: "Elem a U"
     have "P a"
       proof -
-	term "P"
-	term Sep
-	let ?Z = "Sep U (Not o P)"
-	have "?Z = Empty \<longrightarrow> P a" by (simp add: Ext Sep Empty a_in_U)	
-	moreover have "?Z \<noteq> Empty \<longrightarrow> False"
-	  proof 
-	    assume not_empty: "?Z \<noteq> Empty" 
-	    note thereis_x = Regularity[where A="?Z", simplified not_empty, simplified]
-	    then obtain x where x_def: "Elem x ?Z & (! y. Elem y x \<longrightarrow> Not (Elem y ?Z))" ..
+        term "P"
+        term Sep
+        let ?Z = "Sep U (Not o P)"
+        have "?Z = Empty \<longrightarrow> P a" by (simp add: Ext Sep Empty a_in_U)     
+        moreover have "?Z \<noteq> Empty \<longrightarrow> False"
+          proof 
+            assume not_empty: "?Z \<noteq> Empty" 
+            note thereis_x = Regularity[where A="?Z", simplified not_empty, simplified]
+            then obtain x where x_def: "Elem x ?Z & (! y. Elem y x \<longrightarrow> Not (Elem y ?Z))" ..
             then have x_induct:"! y. Elem y x \<longrightarrow> Elem y U \<longrightarrow> P y" by (simp add: Sep)
-	    have "Elem x U \<longrightarrow> P x" 
-	      by (rule impE[OF spec[OF P_induct, where x=x], OF x_induct], assumption)
-	    moreover have "Elem x U & Not(P x)"
-	      apply (insert x_def)
-	      apply (simp add: Sep)
-	      done
-	    ultimately show "False" by auto
-	  qed
-	ultimately show "P a" by auto
+            have "Elem x U \<longrightarrow> P x" 
+              by (rule impE[OF spec[OF P_induct, where x=x], OF x_induct], assumption)
+            moreover have "Elem x U & Not(P x)"
+              apply (insert x_def)
+              apply (simp add: Sep)
+              done
+            ultimately show "False" by auto
+          qed
+        ultimately show "P a" by auto
       qed
   }
   with hyps show ?thesis by blast
@@ -563,18 +562,18 @@
       assume n_less_u: "n < u"
       let ?y = "nat2Nat (u - 1)"
       have "Elem ?y (nat2Nat u)"
-	apply (rule increasing_nat2Nat)
-	apply (insert n_less_u)
-	apply arith
-	done
+        apply (rule increasing_nat2Nat)
+        apply (insert n_less_u)
+        apply arith
+        done
       with u have "Elem ?y x" by auto
       with x have "Not (Elem ?y ?Z)" by auto
       moreover have "Elem ?y ?Z" 
-	apply (insert n_Elem_NatInterval[where q = "u - n - 1" and n=n and m=m])
-	apply (insert n_less_u)
-	apply (insert u)
-	apply auto
-	done
+        apply (insert n_Elem_NatInterval[where q = "u - n - 1" and n=n and m=m])
+        apply (insert n_less_u)
+        apply (insert u)
+        apply auto
+        done
       ultimately show False by auto
     qed
     moreover have "u = n \<longrightarrow> False"
@@ -584,17 +583,17 @@
       then have nm_eq_x: "nat2Nat (n+m) = x" by (simp add: nm)
       let ?y = "nat2Nat (n+m - 1)"
       have "Elem ?y (nat2Nat (n+m))"
-	apply (rule increasing_nat2Nat)
-	apply (insert mg0)
-	apply arith
-	done
+        apply (rule increasing_nat2Nat)
+        apply (insert mg0)
+        apply arith
+        done
       with nm_eq_x have "Elem ?y x" by auto
       with x have "Not (Elem ?y ?Z)" by auto
       moreover have "Elem ?y ?Z" 
-	apply (insert n_Elem_NatInterval[where q = "m - 1" and n=n and m=m])
-	apply (insert mg0)
-	apply auto
-	done
+        apply (insert n_Elem_NatInterval[where q = "m - 1" and n=n and m=m])
+        apply (insert mg0)
+        apply auto
+        done
       ultimately show False by auto      
     qed
     ultimately have "False" using u by arith
@@ -822,12 +821,12 @@
     proof (rule ccontr)
       assume C: "?C \<noteq> {}"
       from
-	wfzf_minimal[OF wfzf C]
+        wfzf_minimal[OF wfzf C]
       obtain x where x: "x \<in> ?C \<and> (\<forall>y. (y, x) \<in> R \<longrightarrow> y \<notin> ?C)" ..
       then have "P x"
-	apply (rule_tac induct[rule_format])
-	apply auto
-	done
+        apply (rule_tac induct[rule_format])
+        apply auto
+        done
       with x show "False" by auto
     qed
     then have "! x. P x" by auto
--- a/src/HOL/ex/Antiquote.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/Antiquote.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -13,10 +13,10 @@
 *}
 
 syntax
-  "_Expr" :: "'a => 'a"				("EXPR _" [1000] 999)
+  "_Expr" :: "'a => 'a"                         ("EXPR _" [1000] 999)
 
 constdefs
-  var :: "'a => ('a => nat) => nat"		("VAR _" [1000] 999)
+  var :: "'a => ('a => nat) => nat"             ("VAR _" [1000] 999)
   "var x env == env x"
 
   Expr :: "(('a => nat) => nat) => ('a => nat) => nat"
@@ -29,9 +29,9 @@
 term "EXPR (a + b + c + VAR x + VAR y + 1)"
 term "EXPR (VAR (f w) + VAR x)"
 
-term "Expr (\<lambda>env. env x)"				(*improper*)
+term "Expr (\<lambda>env. env x)"    -- {* improper *}
 term "Expr (\<lambda>env. f env)"
-term "Expr (\<lambda>env. f env + env x)"			(*improper*)
+term "Expr (\<lambda>env. f env + env x)"    -- {* improper *}
 term "Expr (\<lambda>env. f env y z)"
 term "Expr (\<lambda>env. f env + g y env)"
 term "Expr (\<lambda>env. f env + g env y + h a env z)"
--- a/src/HOL/ex/CTL.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/CTL.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -251,21 +251,21 @@
       show "?lhs \<subseteq> p" ..
       show "?lhs \<subseteq> \<AX> ?lhs"
       proof -
-	{
-	  have "\<AG> (p \<rightarrow> \<AX> p) \<subseteq> p \<rightarrow> \<AX> p" by (rule AG_fp_1)
+        {
+          have "\<AG> (p \<rightarrow> \<AX> p) \<subseteq> p \<rightarrow> \<AX> p" by (rule AG_fp_1)
           moreover have "p \<inter> p \<rightarrow> \<AX> p \<subseteq> \<AX> p" ..
           ultimately have "?lhs \<subseteq> \<AX> p" by auto
-	}  
-	moreover
-	{
-	  have "p \<inter> \<AG> (p \<rightarrow> \<AX> p) \<subseteq> \<AG> (p \<rightarrow> \<AX> p)" ..
+        }  
+        moreover
+        {
+          have "p \<inter> \<AG> (p \<rightarrow> \<AX> p) \<subseteq> \<AG> (p \<rightarrow> \<AX> p)" ..
           also have "\<dots> \<subseteq> \<AX> \<dots>" by (rule AG_fp_2)
           finally have "?lhs \<subseteq> \<AX> \<AG> (p \<rightarrow> \<AX> p)" .
-	}  
-	ultimately have "?lhs \<subseteq> \<AX> p \<inter> \<AX> \<AG> (p \<rightarrow> \<AX> p)"
-	  by (rule Int_greatest)
-	also have "\<dots> = \<AX> ?lhs" by (simp only: AX_int)
-	finally show ?thesis .
+        }  
+        ultimately have "?lhs \<subseteq> \<AX> p \<inter> \<AX> \<AG> (p \<rightarrow> \<AX> p)"
+          by (rule Int_greatest)
+        also have "\<dots> = \<AX> ?lhs" by (simp only: AX_int)
+        finally show ?thesis .
       qed
     qed
   qed
--- a/src/HOL/ex/Classical.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/Classical.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -713,7 +713,7 @@
        (\<forall>x. grain x \<longrightarrow> plant x) & (\<exists>x. grain x) &
        (\<forall>x. animal x \<longrightarrow>
              ((\<forall>y. plant y \<longrightarrow> eats x y)  \<or> 
-	      (\<forall>y. animal y & smaller_than y x &
+              (\<forall>y. animal y & smaller_than y x &
                     (\<exists>z. plant z & eats y z) \<longrightarrow> eats x y))) &
        (\<forall>x y. bird y & (snail x \<or> caterpillar x) \<longrightarrow> smaller_than x y) &
        (\<forall>x y. bird x & fox y \<longrightarrow> smaller_than x y) &
--- a/src/HOL/ex/Hilbert_Classical.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/Hilbert_Classical.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -47,45 +47,45 @@
       assume Q: "Eps ?Q = True"
       have neq: "?P \<noteq> ?Q"
       proof
-	assume "?P = ?Q"
-	hence "Eps ?P = Eps ?Q" by (rule arg_cong)
-	also note P
-	also note Q
-	finally show False by (rule False_neq_True)
+        assume "?P = ?Q"
+        hence "Eps ?P = Eps ?Q" by (rule arg_cong)
+        also note P
+        also note Q
+        finally show False by (rule False_neq_True)
       qed
       have "\<not> A"
       proof
-	assume a: A
-	have "?P = ?Q"
-	proof (rule ext)
-	  fix x show "?P x = ?Q x"
-	  proof
-	    assume "?P x"
-	    thus "?Q x"
-	    proof
-	      assume "x = False"
-	      from this and a have "x = False \<and> A" ..
-	      thus "?Q x" ..
-	    next
-	      assume "x = True \<and> A"
-	      hence "x = True" ..
-	      thus "?Q x" ..
-	    qed
-	  next
-	    assume "?Q x"
-	    thus "?P x"
-	    proof
-	      assume "x = False \<and> A"
-	      hence "x = False" ..
-	      thus "?P x" ..
-	    next
-	      assume "x = True"
-	      from this and a have "x = True \<and> A" ..
-	      thus "?P x" ..
-	    qed
-	  qed
-	qed
-	with neq show False by contradiction
+        assume a: A
+        have "?P = ?Q"
+        proof (rule ext)
+          fix x show "?P x = ?Q x"
+          proof
+            assume "?P x"
+            thus "?Q x"
+            proof
+              assume "x = False"
+              from this and a have "x = False \<and> A" ..
+              thus "?Q x" ..
+            next
+              assume "x = True \<and> A"
+              hence "x = True" ..
+              thus "?Q x" ..
+            qed
+          next
+            assume "?Q x"
+            thus "?P x"
+            proof
+              assume "x = False \<and> A"
+              hence "x = False" ..
+              thus "?P x" ..
+            next
+              assume "x = True"
+              from this and a have "x = True \<and> A" ..
+              thus "?P x" ..
+            qed
+          qed
+        qed
+        with neq show False by contradiction
       qed
       thus ?thesis ..
     qed
--- a/src/HOL/ex/InductiveInvariant.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/InductiveInvariant.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,5 @@
-(*  ID:         $Id$
-    Author:	Sava Krsti\'{c} and John Matthews
+(*
+    Author:     Sava Krsti\'{c} and John Matthews
 *)
 
 header {* Some of the results in Inductive Invariants for Nested Recursion *}
--- a/src/HOL/ex/InductiveInvariant_examples.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/InductiveInvariant_examples.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,5 @@
-(*  ID:         $Id$
-    Author:	Sava Krsti\'{c} and John Matthews
+(*
+    Author:     Sava Krsti\'{c} and John Matthews
 *)
 
 header {* Example use if an inductive invariant to solve termination conditions *}
--- a/src/HOL/ex/LocaleTest2.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/LocaleTest2.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -302,8 +302,8 @@
     proof
       show "y \<sqsubseteq> w"
       proof -
-	have "y \<sqsubseteq> x \<squnion> y" ..
-	also have "... \<sqsubseteq> w" by fact
+        have "y \<sqsubseteq> x \<squnion> y" ..
+        also have "... \<sqsubseteq> w" by fact
         finally show ?thesis .
       qed
       show "z \<sqsubseteq> w" by fact
@@ -451,15 +451,15 @@
   proof -
     { assume c: "y \<sqsubseteq> x" "z \<sqsubseteq> x"
       from c have "?l = y \<squnion> z"
-	by (metis c (*join_commute*) join_connection2 join_related2 (*meet_commute*) meet_connection meet_related2 total)
+        by (metis c (*join_commute*) join_connection2 join_related2 (*meet_commute*) meet_connection meet_related2 total)
       also from c have "... = ?r" by (metis (*c*) (*join_commute*) meet_related2)
       finally have "?l = ?r" . }
     moreover
     { assume c: "x \<sqsubseteq> y | x \<sqsubseteq> z"
       from c have "?l = x"
-	by (metis (*anti_sym*) (*c*) (*circular*) (*join_assoc*)(* join_commute *) join_connection2 (*join_left*) join_related2 meet_connection(* meet_related2*) total trans)
+        by (metis (*anti_sym*) (*c*) (*circular*) (*join_assoc*)(* join_commute *) join_connection2 (*join_left*) join_related2 meet_connection(* meet_related2*) total trans)
       also from c have "... = ?r"
-	by (metis join_commute join_related2 meet_connection meet_related2 total)
+        by (metis join_commute join_related2 meet_connection meet_related2 total)
       finally have "?l = ?r" . }
     moreover note total
     ultimately show ?thesis by blast
@@ -855,15 +855,15 @@
       proof (rule inj_onI)
         fix x y
         assume "f x = f y"
-	then have "(g o f) x = (g o f) y" by simp
-	with id2 show "x = y" by simp
+        then have "(g o f) x = (g o f) y" by simp
+        with id2 show "x = y" by simp
       qed
     next
       show "surj f"
       proof (rule surjI)
-	fix x
+        fix x
         from id1 have "(f o g) x = x" by simp
-	then show "f (g x) = x" by simp
+        then show "f (g x) = x" by simp
       qed
     qed
   next
--- a/src/HOL/ex/MergeSort.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/MergeSort.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -42,7 +42,7 @@
   "msort [] = []"
 | "msort [x] = [x]"
 | "msort xs = merge (msort (take (size xs div 2) xs))
-	                  (msort (drop (size xs div 2) xs))"
+                    (msort (drop (size xs div 2) xs))"
 
 theorem sorted_msort: "sorted (op \<le>) (msort xs)"
 by (induct xs rule: msort.induct) simp_all
--- a/src/HOL/ex/Multiquote.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/Multiquote.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -13,7 +13,7 @@
 *}
 
 syntax
-  "_quote" :: "'b => ('a => 'b)"	     ("\<guillemotleft>_\<guillemotright>" [0] 1000)
+  "_quote" :: "'b => ('a => 'b)"             ("\<guillemotleft>_\<guillemotright>" [0] 1000)
   "_antiquote" :: "('a => 'b) => 'b"         ("\<acute>_" [1000] 1000)
 
 parse_translation {*
--- a/src/HOL/ex/ReflectionEx.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/ReflectionEx.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -382,7 +382,7 @@
 oops
 
 
-  (* An example for equations containing type variables *)	
+(* An example for equations containing type variables *)
 datatype prod = Zero | One | Var nat | Mul prod prod 
   | Pw prod nat | PNM nat nat prod
 consts Iprod :: " prod \<Rightarrow> ('a::{ordered_idom}) list \<Rightarrow>'a" 
--- a/src/HOL/ex/Unification.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOL/ex/Unification.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -234,8 +234,7 @@
       by auto
     
     have "\<sigma>' =\<^sub>s \<sigma> \<bullet> \<rho>" unfolding \<sigma>
-      by (rule eqv_intro, auto simp:eqv_dest[OF eqv] 
-	eqv_dest[OF eqv2])
+      by (rule eqv_intro, auto simp:eqv_dest[OF eqv] eqv_dest[OF eqv2])
     thus "\<exists>\<gamma>. \<sigma>' =\<^sub>s \<sigma> \<bullet> \<gamma>" ..
   qed
 qed (auto split:split_if_asm) -- "Solve the remaining cases automatically"
@@ -388,22 +387,22 @@
     
     show "v \<in> vars_of (M \<cdot> N) \<union> vars_of (M' \<cdot> N') \<union> vars_of t"
     proof (cases "v \<notin> vars_of M \<and> v \<notin> vars_of M'
-	    \<and> v \<notin> vars_of N \<and> v \<notin> vars_of N'")
+        \<and> v \<notin> vars_of N \<and> v \<notin> vars_of N'")
       case True
       with ih1 have l:"\<And>t. v \<in> vars_of (t \<triangleleft> \<theta>1) \<Longrightarrow> v \<in> vars_of t"
-	    by auto
+        by auto
       
       from a and ih2[where t="t \<triangleleft> \<theta>1"]
       have "v \<in> vars_of (N \<triangleleft> \<theta>1) \<union> vars_of (N' \<triangleleft> \<theta>1) 
         \<or> v \<in> vars_of (t \<triangleleft> \<theta>1)" unfolding \<sigma>
-	    by auto
+        by auto
       hence "v \<in> vars_of t"
       proof
-	    assume "v \<in> vars_of (N \<triangleleft> \<theta>1) \<union> vars_of (N' \<triangleleft> \<theta>1)"
-	    with True show ?thesis by (auto dest:l)
+        assume "v \<in> vars_of (N \<triangleleft> \<theta>1) \<union> vars_of (N' \<triangleleft> \<theta>1)"
+        with True show ?thesis by (auto dest:l)
       next
-	    assume "v \<in> vars_of (t \<triangleleft> \<theta>1)" 
-	    thus ?thesis by (rule l)
+        assume "v \<in> vars_of (t \<triangleleft> \<theta>1)" 
+        thus ?thesis by (rule l)
       qed
       
       thus ?thesis by auto
@@ -515,12 +514,12 @@
     -- {* Either a variable is eliminated \ldots *}
     assume "(\<exists>v\<in>vars_of M \<union> vars_of M'. elim \<theta> v)"
     then obtain v 
-	  where "elim \<theta> v" 
-	  and "v\<in>vars_of M \<union> vars_of M'" by auto
+      where "elim \<theta> v" 
+      and "v\<in>vars_of M \<union> vars_of M'" by auto
     with unify_vars[OF inner]
     have "vars_of (N\<triangleleft>\<theta>) \<union> vars_of (N'\<triangleleft>\<theta>)
-	  \<subset> vars_of (M\<cdot>N) \<union> vars_of (M'\<cdot>N')"
-	  by auto
+      \<subset> vars_of (M\<cdot>N) \<union> vars_of (M'\<cdot>N')"
+      by auto
     
     thus ?thesis
       by (auto intro!: measures_less intro: psubset_card_mono)
@@ -528,7 +527,7 @@
     -- {* Or the substitution is empty *}
     assume "\<theta> =\<^sub>s []"
     hence "N \<triangleleft> \<theta> = N" 
-	  and "N' \<triangleleft> \<theta> = N'" by auto
+      and "N' \<triangleleft> \<theta> = N'" by auto
     thus ?thesis 
        by (auto intro!: measures_less intro: psubset_card_mono)
   qed
--- a/src/HOLCF/FOCUS/Fstream.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/FOCUS/Fstream.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,9 @@
 (*  Title:      HOLCF/FOCUS/Fstream.thy
-    ID:         $Id$
     Author:     David von Oheimb, TU Muenchen
 
-FOCUS streams (with lifted elements)
-###TODO: integrate Fstreams.thy
+FOCUS streams (with lifted elements).
+
+TODO: integrate Fstreams.thy
 *)
 
 header {* FOCUS flat streams *}
--- a/src/HOLCF/FOCUS/Fstreams.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/FOCUS/Fstreams.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,9 @@
-(*  Title: 	HOLCF/FOCUS/Fstreams.thy
-    ID:         $Id$
-    Author: 	Borislav Gajanovic
+(*  Title:      HOLCF/FOCUS/Fstreams.thy
+    Author:     Borislav Gajanovic
 
-FOCUS flat streams (with lifted elements)
-###TODO: integrate this with Fstream.*
+FOCUS flat streams (with lifted elements).
+
+TODO: integrate this with Fstream.
 *)
 
 theory Fstreams imports "../ex/Stream" begin
@@ -21,7 +21,7 @@
   "fsfilter A = sfilter\<cdot>(flift2 (\<lambda>x. x\<in>A))"
 
 definition
-  fsmap		:: "('a => 'b) => 'a fstream -> 'b fstream" where
+  fsmap         :: "('a => 'b) => 'a fstream -> 'b fstream" where
   "fsmap f = smap$(flift2 f)"
 
 definition
@@ -42,7 +42,7 @@
   "<> == \<bottom>"
 
 abbreviation
-  fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream"	("(_'(C')_)" [64,63] 63) where
+  fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream"       ("(_'(C')_)" [64,63] 63) where
   "A(C)s == fsfilter A\<cdot>s"
 
 notation (xsymbols)
--- a/src/HOLCF/FOCUS/ROOT.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/FOCUS/ROOT.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,1 @@
-(*  Title:      HOLCF/FOCUS/ROOT.ML
-    ID:         $Id$
-    Author: 	David von Oheimb, TU Muenchen
-
-ROOT file for the FOCUS extension of HOLCF.
-See README.html for further information.
-*)
-
 use_thys ["Fstreams", "FOCUS", "Buffer_adm"];
--- a/src/HOLCF/IOA/Modelcheck/MuIOA.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/IOA/Modelcheck/MuIOA.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,3 @@
-
-(* $Id$ *)
-
 theory MuIOA
 imports IOA "../../../HOL/Modelcheck/MuckeSyn"
 begin
@@ -20,8 +17,8 @@
 ML {*
 
 (* MuIOA.ML declares function mk_sim_oracle used by oracle "Sim" (see MuIOAOracle.thy).
-	There, implementation relations for I/O automatons are proved using
-	the model checker mucke (invoking cal_mucke_tac defined in MCSyn.ML). *)
+        There, implementation relations for I/O automatons are proved using
+        the model checker mucke (invoking cal_mucke_tac defined in MCSyn.ML). *)
 
 exception SimFailureExn of string;
 
@@ -161,13 +158,13 @@
 
 fun double_tupel_of _ _ _ _ [] = "" |
 double_tupel_of sign s t n [a] = s ^ (Int.toString(n)) ^ "," ^
-				 t ^ (Int.toString(n)) |
+                                 t ^ (Int.toString(n)) |
 double_tupel_of sign s t n (a::r) = s ^ (Int.toString(n)) ^ "," ^
-		t ^ (Int.toString(n)) ^ "," ^ (double_tupel_of sign s t (n+1) r);
+                t ^ (Int.toString(n)) ^ "," ^ (double_tupel_of sign s t (n+1) r);
 
 fun eq_string _ _ _ [] = "" |
 eq_string s t n [a] = s ^ (Int.toString(n)) ^ " = " ^
-			 t ^ (Int.toString(n)) |
+                         t ^ (Int.toString(n)) |
 eq_string s t n (a::r) =
  s ^ (Int.toString(n)) ^ " = " ^
  t ^ (Int.toString(n)) ^ " & " ^ (eq_string s t (n+1) r);
@@ -188,60 +185,60 @@
     val concl = Logic.strip_imp_concl subgoal;
     val ic_str = delete_ul_string(Syntax.string_of_term_global sign (IntC sign concl));
     val ia_str = delete_ul_string(Syntax.string_of_term_global sign (IntA sign concl));
-	val sc_str = delete_ul_string(Syntax.string_of_term_global sign (StartC sign concl));	
-	val sa_str = delete_ul_string(Syntax.string_of_term_global sign (StartA sign concl));
-	val tc_str = delete_ul_string(Syntax.string_of_term_global sign (TransC sign concl));
-	val ta_str = delete_ul_string(Syntax.string_of_term_global sign (TransA sign concl));
-	
-	val action_type_str =
-	Syntax.string_of_typ_global sign (element_type(#T (rep_cterm(cterm_of sign (IntA sign concl)))));
+        val sc_str = delete_ul_string(Syntax.string_of_term_global sign (StartC sign concl));   
+        val sa_str = delete_ul_string(Syntax.string_of_term_global sign (StartA sign concl));
+        val tc_str = delete_ul_string(Syntax.string_of_term_global sign (TransC sign concl));
+        val ta_str = delete_ul_string(Syntax.string_of_term_global sign (TransA sign concl));
+        
+        val action_type_str =
+        Syntax.string_of_typ_global sign (element_type(#T (rep_cterm(cterm_of sign (IntA sign concl)))));
 
-	val abs_state_type_list =
-	type_list_of sign (element_type(#T (rep_cterm(cterm_of sign (StartA sign concl)))));
-	val con_state_type_list =
-	type_list_of sign (element_type(#T (rep_cterm(cterm_of sign (StartC sign concl)))));
+        val abs_state_type_list =
+        type_list_of sign (element_type(#T (rep_cterm(cterm_of sign (StartA sign concl)))));
+        val con_state_type_list =
+        type_list_of sign (element_type(#T (rep_cterm(cterm_of sign (StartC sign concl)))));
 
-	val tupel_eq = eq_string "s" "t" 0 abs_state_type_list;
+        val tupel_eq = eq_string "s" "t" 0 abs_state_type_list;
 
-	val abs_pre_tupel_typed = tupel_typed_of sign "s" 0 abs_state_type_list;
-	val abs_s_t_tupel_typed = double_tupel_typed_of sign "s" "t" 0 abs_state_type_list;
-	val con_pre_tupel_typed = tupel_typed_of sign "cs" 0 con_state_type_list;
-	val con_s_t_tupel_typed = double_tupel_typed_of sign "cs" "ct" 0 con_state_type_list;
-	
-	val abs_pre_tupel = tupel_of sign "s" 0 abs_state_type_list;
-	val abs_post_tupel = tupel_of sign "t" 0 abs_state_type_list;
-	val abs_s_t_tupel = double_tupel_of sign "s" "t" 0 abs_state_type_list;
-	val abs_s_u_tupel = double_tupel_of sign "s" "u" 0 abs_state_type_list;
-	val abs_u_t_tupel = double_tupel_of sign "u" "t" 0 abs_state_type_list;
-	val abs_u_v_tupel = double_tupel_of sign "u" "v" 0 abs_state_type_list;
-	val abs_v_t_tupel = double_tupel_of sign "v" "t" 0 abs_state_type_list;
-	val con_pre_tupel = tupel_of sign "cs" 0 con_state_type_list;
+        val abs_pre_tupel_typed = tupel_typed_of sign "s" 0 abs_state_type_list;
+        val abs_s_t_tupel_typed = double_tupel_typed_of sign "s" "t" 0 abs_state_type_list;
+        val con_pre_tupel_typed = tupel_typed_of sign "cs" 0 con_state_type_list;
+        val con_s_t_tupel_typed = double_tupel_typed_of sign "cs" "ct" 0 con_state_type_list;
+        
+        val abs_pre_tupel = tupel_of sign "s" 0 abs_state_type_list;
+        val abs_post_tupel = tupel_of sign "t" 0 abs_state_type_list;
+        val abs_s_t_tupel = double_tupel_of sign "s" "t" 0 abs_state_type_list;
+        val abs_s_u_tupel = double_tupel_of sign "s" "u" 0 abs_state_type_list;
+        val abs_u_t_tupel = double_tupel_of sign "u" "t" 0 abs_state_type_list;
+        val abs_u_v_tupel = double_tupel_of sign "u" "v" 0 abs_state_type_list;
+        val abs_v_t_tupel = double_tupel_of sign "v" "t" 0 abs_state_type_list;
+        val con_pre_tupel = tupel_of sign "cs" 0 con_state_type_list;
         val con_post_tupel = tupel_of sign "ct" 0 con_state_type_list;
-	val con_s_t_tupel = double_tupel_of sign "cs" "ct" 0 con_state_type_list;
+        val con_s_t_tupel = double_tupel_of sign "cs" "ct" 0 con_state_type_list;
 
         val abs_pre_varlist = varlist_of 0 "s" abs_state_type_list;
         val abs_post_varlist = varlist_of 0 "t" abs_state_type_list;
         val abs_u_varlist = varlist_of 0 "u" abs_state_type_list;
         val abs_v_varlist = varlist_of 0 "v" abs_state_type_list;
-	val con_pre_varlist = varlist_of 0 "cs" con_state_type_list;
+        val con_pre_varlist = varlist_of 0 "cs" con_state_type_list;
         val con_post_varlist = varlist_of 0 "ct" con_state_type_list;
 
         val abs_post_str = string_of abs_post_varlist;
-	val abs_u_str = string_of abs_u_varlist;
-	val con_post_str = string_of con_post_varlist;
-	
-	val abs_pre_str_typed = merge_var_and_type abs_pre_varlist abs_state_type_list;
-	val abs_u_str_typed = merge_var_and_type abs_u_varlist abs_state_type_list;
+        val abs_u_str = string_of abs_u_varlist;
+        val con_post_str = string_of con_post_varlist;
+        
+        val abs_pre_str_typed = merge_var_and_type abs_pre_varlist abs_state_type_list;
+        val abs_u_str_typed = merge_var_and_type abs_u_varlist abs_state_type_list;
         val abs_v_str_typed = merge_var_and_type abs_v_varlist abs_state_type_list;
-	val con_pre_str_typed = merge_var_and_type con_pre_varlist con_state_type_list;
+        val con_pre_str_typed = merge_var_and_type con_pre_varlist con_state_type_list;
 
-	val abs_pre_tupel_struct = snd(
+        val abs_pre_tupel_struct = snd(
 structured_tuple abs_pre_varlist (element_type(#T(rep_cterm(cterm_of sign (StartA sign concl))))));
-	val abs_post_tupel_struct = snd(
+        val abs_post_tupel_struct = snd(
 structured_tuple abs_post_varlist (element_type(#T(rep_cterm(cterm_of sign (StartA sign concl))))));
-	val con_pre_tupel_struct = snd(
+        val con_pre_tupel_struct = snd(
 structured_tuple con_pre_varlist (element_type(#T(rep_cterm(cterm_of sign (StartC sign concl))))));
-	val con_post_tupel_struct = snd(
+        val con_post_tupel_struct = snd(
 structured_tuple con_post_varlist (element_type(#T(rep_cterm(cterm_of sign (StartC sign concl))))));
 in
 (
@@ -250,37 +247,37 @@
 ( "(Internal_of_A = (% a::(" ^ action_type_str ^ "). a:(" ^ ia_str^
   "))) --> (Internal_of_C = (% a::(" ^ action_type_str ^ "). a:(" ^ ic_str ^ 
   "))) --> (Start_of_A = (% (" ^ abs_pre_tupel_typed ^
-	"). (" ^ abs_pre_tupel_struct ^ "):(" ^ sa_str ^
+        "). (" ^ abs_pre_tupel_struct ^ "):(" ^ sa_str ^
   "))) --> (Start_of_C = (% (" ^ con_pre_tupel_typed ^
-	"). (" ^ con_pre_tupel_struct ^ "):(" ^ sc_str ^
+        "). (" ^ con_pre_tupel_struct ^ "):(" ^ sc_str ^
   "))) --> (Trans_of_A = (% (" ^ abs_s_t_tupel_typed ^ ") (a::(" ^ action_type_str ^ 
-	")). ((" ^ abs_pre_tupel_struct ^ "),a,(" ^ abs_post_tupel_struct ^ ")):(" ^ ta_str ^
+        ")). ((" ^ abs_pre_tupel_struct ^ "),a,(" ^ abs_post_tupel_struct ^ ")):(" ^ ta_str ^
   "))) --> (Trans_of_C = (% (" ^ con_s_t_tupel_typed ^ ") (a::(" ^ action_type_str ^ 
-	")). ((" ^ con_pre_tupel_struct ^ "),a,(" ^ con_post_tupel_struct ^ ")):(" ^ tc_str ^
+        ")). ((" ^ con_pre_tupel_struct ^ "),a,(" ^ con_post_tupel_struct ^ ")):(" ^ tc_str ^
   "))) --> (IntStep_of_A = (% (" ^ abs_s_t_tupel_typed ^ 
-	"). ? (a::(" ^ action_type_str ^
-	")). Internal_of_A a & Trans_of_A (" ^ abs_s_t_tupel ^ ") a" ^
+        "). ? (a::(" ^ action_type_str ^
+        ")). Internal_of_A a & Trans_of_A (" ^ abs_s_t_tupel ^ ") a" ^
   ")) --> (IntStepStar_of_A =  mu (% P (" ^ abs_s_t_tupel_typed ^
-	 "). (" ^ tupel_eq ^ ") | (? " ^ abs_u_str ^ ". IntStep_of_A (" ^ abs_s_u_tupel ^
-	 ") & P(" ^ abs_u_t_tupel ^ ")))" ^
+         "). (" ^ tupel_eq ^ ") | (? " ^ abs_u_str ^ ". IntStep_of_A (" ^ abs_s_u_tupel ^
+         ") & P(" ^ abs_u_t_tupel ^ ")))" ^
   ") --> (Move_of_A = (% (" ^ abs_s_t_tupel_typed ^ ") (a::(" ^ action_type_str ^ 
-	")). (Internal_of_C a & IntStepStar_of_A(" ^ abs_s_t_tupel ^ 
-	")) | (? " ^ abs_u_str_typed ^ ". ? " ^ abs_v_str_typed ^
-	". IntStepStar_of_A(" ^ abs_s_u_tupel ^ ") & " ^
-	"Trans_of_A (" ^ abs_u_v_tupel ^ ") a  & IntStepStar_of_A(" ^ abs_v_t_tupel ^ "))" ^
+        ")). (Internal_of_C a & IntStepStar_of_A(" ^ abs_s_t_tupel ^ 
+        ")) | (? " ^ abs_u_str_typed ^ ". ? " ^ abs_v_str_typed ^
+        ". IntStepStar_of_A(" ^ abs_s_u_tupel ^ ") & " ^
+        "Trans_of_A (" ^ abs_u_v_tupel ^ ") a  & IntStepStar_of_A(" ^ abs_v_t_tupel ^ "))" ^
   ")) --> (isSimCA = nu (% P (" ^ con_pre_tupel_typed ^ "," ^ abs_pre_tupel_typed ^ 
-	"). (! (a::(" ^ action_type_str ^ ")) " ^ con_post_str ^
-	". Trans_of_C (" ^ con_s_t_tupel ^ ") a  -->" ^ " (? " ^ abs_post_str ^
+        "). (! (a::(" ^ action_type_str ^ ")) " ^ con_post_str ^
+        ". Trans_of_C (" ^ con_s_t_tupel ^ ") a  -->" ^ " (? " ^ abs_post_str ^
   ". Move_of_A (" ^ abs_s_t_tupel ^ ") a  & P(" ^ con_post_tupel ^ "," ^ abs_post_tupel ^ "))))" ^
  ") --> (! " ^ con_pre_str_typed ^ ". ! " ^ abs_pre_str_typed ^
-	". (Start_of_C (" ^ con_pre_tupel ^ ") & Start_of_A (" ^ abs_pre_tupel ^
-	")) --> isSimCA(" ^ con_pre_tupel ^ "," ^ abs_pre_tupel ^ "))");
+        ". (Start_of_C (" ^ con_pre_tupel ^ ") & Start_of_A (" ^ abs_pre_tupel ^
+        ")) --> isSimCA(" ^ con_pre_tupel ^ "," ^ abs_pre_tupel ^ "))");
 OldGoals.by (REPEAT (rtac impI 1));
 OldGoals.by (REPEAT (dtac eq_reflection 1));
 (* Bis hierher wird im Kapitel 4 erl"autert, ab hier im Kapitel 5 *)
 OldGoals.by (full_simp_tac ((global_simpset_of sign delcongs (if_weak_cong :: weak_case_congs)
                               delsimps [not_iff,split_part])
-			      addsimps (thl @ comp_simps @ restrict_simps @ hide_simps @
+                              addsimps (thl @ comp_simps @ restrict_simps @ hide_simps @
                                         rename_simps @ ioa_simps @ asig_simps)) 1);
 OldGoals.by (full_simp_tac
   (Mucke_ss delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) 1);
--- a/src/HOLCF/IOA/Modelcheck/MuIOAOracle.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/IOA/Modelcheck/MuIOAOracle.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,3 @@
-
-(* $Id$ *)
-
 theory MuIOAOracle
 imports MuIOA
 begin
@@ -17,21 +14,21 @@
 val ioa_implements_def = thm "ioa_implements_def";
 
 (* is_sim_tac makes additionally to call_sim_tac some simplifications,
-	which are suitable for implementation realtion formulas *)
+        which are suitable for implementation realtion formulas *)
 fun is_sim_tac ss thm_list = SUBGOAL (fn (goal, i) =>
   EVERY [REPEAT (etac thin_rl i),
-	 simp_tac (ss addsimps [ioa_implements_def]) i,
+         simp_tac (ss addsimps [ioa_implements_def]) i,
          rtac conjI i,
          rtac conjI (i+1),
-	 TRY(call_sim_tac thm_list (i+2)),
-	 TRY(atac (i+2)), 
+         TRY(call_sim_tac thm_list (i+2)),
+         TRY(atac (i+2)), 
          REPEAT(rtac refl (i+2)),
-	 simp_tac (ss addsimps (thm_list @
-				       comp_simps @ restrict_simps @ hide_simps @
-				       rename_simps @  ioa_simps @ asig_simps)) (i+1),
-	 simp_tac (ss addsimps (thm_list @
-				       comp_simps @ restrict_simps @ hide_simps @
-				       rename_simps @ ioa_simps @ asig_simps)) (i)]);
+         simp_tac (ss addsimps (thm_list @
+                                       comp_simps @ restrict_simps @ hide_simps @
+                                       rename_simps @  ioa_simps @ asig_simps)) (i+1),
+         simp_tac (ss addsimps (thm_list @
+                                       comp_simps @ restrict_simps @ hide_simps @
+                                       rename_simps @ ioa_simps @ asig_simps)) (i)]);
 
 *}
 
--- a/src/HOLCF/IOA/Modelcheck/Ring3.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/IOA/Modelcheck/Ring3.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,3 @@
-
-(* $Id$ *)
-
 theory Ring3
 imports MuIOAOracle
 begin
@@ -112,7 +109,7 @@
   aut2_asig_def aut2_trans_def aut2_initial_def aut2_def
 
 (* property to prove: at most one (of 3) members of the ring will
-	declare itself to leader *)
+   declare itself to leader *)
 lemma at_most_one_leader: "ring =<| one_leader"
 apply (tactic {* is_sim_tac @{simpset} (thms "aut_simps") 1 *})
 apply auto
--- a/src/HOLCF/IOA/meta_theory/automaton.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/IOA/meta_theory/automaton.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -200,14 +200,14 @@
 (* the case with transrel *)
  else 
  (trm,actstr,"(" ^ r ^ ") & " ^ (poststring statetupel ((a,b)::c)),
-	"(" ^ r ^ ") & " ^ (pseudo_poststring ((a,b)::c)),flag))
+  "(" ^ r ^ ") & " ^ (pseudo_poststring ((a,b)::c)),flag))
 else
 error("transition " ^ actstr ^ " is not a pure constructor term")
 end;
 (* used by make_alt_string *) 
 fun extended_list _ _ _ [] = [] |
 extended_list thy atyp statetupel (a::r) =
-	 (extend thy atyp statetupel a) :: (extended_list thy atyp statetupel r);
+   (extend thy atyp statetupel a) :: (extended_list thy atyp statetupel r);
 
 (* used by write_alts *)
 fun write_alt thy (chead,tr) inp out int [] =
@@ -230,10 +230,10 @@
 (if ((chead mem inp) andalso e) then (
 error("Input action " ^ b ^ " has a precondition")
 ) else (if (chead mem (inp@out@int)) then 
-		(if (occurs_again chead r) then (
+    (if (occurs_again chead r) then (
 error("Two specifications for action: " ^ b)
-		) else (b ^ " => " ^ c,b ^ " => " ^ d))
-	else (
+    ) else (b ^ " => " ^ c,b ^ " => " ^ d))
+  else (
 error("Action " ^ b ^ " is not in automaton signature")
 ))) else (write_alt thy (chead,ctrm) inp out int r)
 handle malformed =>
@@ -344,7 +344,7 @@
 val int_head_list = constructor_head_list thy action_type int;
 val overlap_flag = ((overlap inp out) andalso (overlap inp int) andalso (overlap out int)); 
 val alt_string = make_alt_string thy inp_head_list out_head_list int_head_list 
-							atyp statetupel trans;
+              atyp statetupel trans;
 val thy2 = (thy
 |> Sign.add_consts
 [(Binding.name (automaton_name ^ "_initial"), "(" ^ state_type_string ^ ")set" ,NoSyn),
--- a/src/HOLCF/Sprod.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/Sprod.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -23,9 +23,9 @@
 by (rule typedef_chfin [OF type_definition_Sprod below_Sprod_def])
 
 syntax (xsymbols)
-  "**"		:: "[type, type] => type"	 ("(_ \<otimes>/ _)" [21,20] 20)
+  "**"          :: "[type, type] => type"        ("(_ \<otimes>/ _)" [21,20] 20)
 syntax (HTML output)
-  "**"		:: "[type, type] => type"	 ("(_ \<otimes>/ _)" [21,20] 20)
+  "**"          :: "[type, type] => type"        ("(_ \<otimes>/ _)" [21,20] 20)
 
 lemma spair_lemma:
   "(strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a) \<in> Sprod"
--- a/src/HOLCF/Ssum.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/Ssum.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -25,9 +25,9 @@
 by (rule typedef_chfin [OF type_definition_Ssum below_Ssum_def])
 
 syntax (xsymbols)
-  "++"		:: "[type, type] => type"	("(_ \<oplus>/ _)" [21, 20] 20)
+  "++"          :: "[type, type] => type"       ("(_ \<oplus>/ _)" [21, 20] 20)
 syntax (HTML output)
-  "++"		:: "[type, type] => type"	("(_ \<oplus>/ _)" [21, 20] 20)
+  "++"          :: "[type, type] => type"       ("(_ \<oplus>/ _)" [21, 20] 20)
 
 subsection {* Definitions of constructors *}
 
--- a/src/HOLCF/Tools/Domain/domain_theorems.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/HOLCF/Tools/Domain/domain_theorems.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -168,7 +168,7 @@
 val rep_strict = ax_abs_iso RS (allI RS retraction_strict);
 val abs_defin' = iso_locale RS iso_abs_defin';
 val rep_defin' = iso_locale RS iso_rep_defin';
-val iso_rews = map standard [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict];
+val iso_rews = map Drule.standard [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict];
 
 (* ----- generating beta reduction rules from definitions-------------------- *)
 
@@ -251,7 +251,7 @@
   val exhaust = pg con_appls (mk_trp exh) (K tacs);
   val _ = trace " Proving casedist...";
   val casedist =
-    standard (rewrite_rule exh_casedists (exhaust RS exh_casedist0));
+    Drule.standard (rewrite_rule exh_casedists (exhaust RS exh_casedist0));
 end;
 
 local 
@@ -542,7 +542,7 @@
         flat
           (ListPair.map (distinct c) ((map #1 cs),leqs)) @
         distincts cs;
-  in map standard (distincts (cons ~~ distincts_le)) end;
+  in map Drule.standard (distincts (cons ~~ distincts_le)) end;
 
 local 
   fun pgterm rel con args =
@@ -738,7 +738,7 @@
         maps (eq_tacs ctxt) eqs;
     in pg copy_take_defs goal tacs end;
 in
-  val take_rews = map standard
+  val take_rews = map Drule.standard
     (atomize global_ctxt take_stricts @ take_0s @ atomize global_ctxt take_apps);
 end; (* local *)
 
--- a/src/Provers/classical.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Provers/classical.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -363,7 +363,7 @@
     (warning ("Ignoring duplicate safe elimination (elim!)\n" ^
         Display.string_of_thm_without_context th); cs)
   else if has_fewer_prems 1 th then
-    	error ("Ill-formed elimination rule\n" ^ Display.string_of_thm_without_context th)
+        error ("Ill-formed elimination rule\n" ^ Display.string_of_thm_without_context th)
   else
   let
       val th' = classical_rule (flat_rule th)
@@ -390,7 +390,7 @@
 
 fun make_elim th =
     if has_fewer_prems 1 th then
-    	  error ("Ill-formed destruction rule\n" ^ Display.string_of_thm_without_context th)
+          error ("Ill-formed destruction rule\n" ^ Display.string_of_thm_without_context th)
     else Tactic.make_elim th;
 
 fun cs addSDs ths = cs addSEs (map make_elim ths);
@@ -431,7 +431,7 @@
     (warning ("Ignoring duplicate elimination (elim)\n" ^
         Display.string_of_thm_without_context th); cs)
   else if has_fewer_prems 1 th then
-    	error("Ill-formed elimination rule\n" ^ Display.string_of_thm_without_context th)
+        error("Ill-formed elimination rule\n" ^ Display.string_of_thm_without_context th)
   else
   let
       val th' = classical_rule (flat_rule th)
--- a/src/Provers/hypsubst.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Provers/hypsubst.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -165,7 +165,7 @@
 
 end;
 
-val ssubst = standard (Data.sym RS Data.subst);
+val ssubst = Drule.standard (Data.sym RS Data.subst);
 
 fun inst_subst_tac b rl = CSUBGOAL (fn (cBi, i) =>
   case try (Logic.strip_assums_hyp #> hd #>
--- a/src/Provers/typedsimp.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Provers/typedsimp.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,5 @@
-(*  Title: 	typedsimp
-    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
+(*  Title:      Provers/typedsimp.ML
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 
 Functor for constructing simplifiers.  Suitable for Constructive Type
@@ -9,11 +9,11 @@
 
 signature TSIMP_DATA =
   sig
-  val refl: thm		(*Reflexive law*)
-  val sym: thm		(*Symmetric law*)
-  val trans: thm	(*Transitive law*)
-  val refl_red: thm	(* reduce(a,a) *)
-  val trans_red: thm	(* [|a=b; reduce(b,c) |] ==> a=c *)
+  val refl: thm         (*Reflexive law*)
+  val sym: thm          (*Symmetric law*)
+  val trans: thm        (*Transitive law*)
+  val refl_red: thm     (* reduce(a,a) *)
+  val trans_red: thm    (* [|a=b; reduce(b,c) |] ==> a=c *)
   val red_if_equal: thm (* a=b ==> reduce(a,b) *)
   (*Built-in rewrite rules*)
   val default_rls: thm list
@@ -43,11 +43,11 @@
 (*For simplifying both sides of an equation:
       [| a=c; b=c |] ==> b=a
   Can use resolve_tac [split_eqn] to prepare an equation for simplification. *)
-val split_eqn = standard (sym RSN (2,trans) RS sym);
+val split_eqn = Drule.standard (sym RSN (2,trans) RS sym);
 
 
 (*    [| a=b; b=c |] ==> reduce(a,c)  *)
-val red_trans = standard (trans RS red_if_equal);
+val red_trans = Drule.standard (trans RS red_if_equal);
 
 (*For REWRITE rule: Make a reduction rule for simplification, e.g.
   [| a: C(0); ... ; a=c: C(0) |] ==> rec(0,a,b) = c: C(0) *)
@@ -76,7 +76,7 @@
   case process_rules rls of
       (simp_rls,[])  =>  simp_rls
     | (_,others) => raise THM 
-	("process_rewrites: Ill-formed rewrite", 0, others);
+        ("process_rewrites: Ill-formed rewrite", 0, others);
 
 (*Process the default rewrite rules*)
 val simp_rls = process_rewrites default_rls;
@@ -97,8 +97,8 @@
 (*Resolve with asms, whether rewrites or not*)
 fun asm_res_tac asms =
     let val (xsimp_rls,xother_rls) = process_rules asms
-    in 	routine_tac xother_rls  ORELSE'  
-	filt_resolve_tac xsimp_rls 2
+    in  routine_tac xother_rls  ORELSE'  
+        filt_resolve_tac xsimp_rls 2
     end;
 
 (*Single step for simple rewriting*)
--- a/src/Pure/drule.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Pure/drule.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -29,8 +29,6 @@
   val zero_var_indexes_list: thm list -> thm list
   val zero_var_indexes: thm -> thm
   val implies_intr_hyps: thm -> thm
-  val standard: thm -> thm
-  val standard': thm -> thm
   val rotate_prems: int -> thm -> thm
   val rearrange_prems: int list -> thm -> thm
   val RSN: thm * (int * thm) -> thm
@@ -77,6 +75,8 @@
   val beta_conv: cterm -> cterm -> cterm
   val types_sorts: thm -> (indexname-> typ option) * (indexname-> sort option)
   val flexflex_unique: thm -> thm
+  val standard: thm -> thm
+  val standard': thm -> thm
   val get_def: theory -> xstring -> thm
   val store_thm: bstring -> thm -> thm
   val store_standard_thm: bstring -> thm -> thm
--- a/src/Pure/old_goals.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Pure/old_goals.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -305,7 +305,7 @@
             val th = Goal.conclude ath
             val thy' = Thm.theory_of_thm th
             val {hyps, prop, ...} = Thm.rep_thm th
-            val final_th = standard th
+            val final_th = Drule.standard th
         in  if not check then final_th
             else if not (Theory.eq_thy(thy,thy')) then !result_error_fn state
                 ("Theory of proof state has changed!" ^
@@ -512,7 +512,7 @@
             0 => thm
           | i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
   in
-    standard (implies_intr_list As
+    Drule.standard (implies_intr_list As
       (check (Seq.pull (EVERY (f asms) (trivial B)))))
   end;
 
--- a/src/Sequents/modal.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Sequents/modal.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -86,7 +86,7 @@
 fun step_tac n = 
     COND (has_fewer_prems 1) all_tac 
          (DETERM(fres_safe_tac n) ORELSE 
-	  (fres_unsafe_tac n APPEND fres_bound_tac n));
+          (fres_unsafe_tac n APPEND fres_bound_tac n));
 
 end;
 end;
--- a/src/Sequents/prover.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Sequents/prover.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -32,21 +32,21 @@
 
 fun (Pack(safes,unsafes)) add_safes ths   = 
     let val dups = warn_duplicates (gen_inter Thm.eq_thm_prop (ths,safes))
-	val ths' = subtract Thm.eq_thm_prop dups ths
+        val ths' = subtract Thm.eq_thm_prop dups ths
     in
         Pack(sort (make_ord less) (ths'@safes), unsafes)
     end;
 
 fun (Pack(safes,unsafes)) add_unsafes ths = 
     let val dups = warn_duplicates (gen_inter Thm.eq_thm_prop (ths,unsafes))
-	val ths' = subtract Thm.eq_thm_prop dups ths
+        val ths' = subtract Thm.eq_thm_prop dups ths
     in
-	Pack(safes, sort (make_ord less) (ths'@unsafes))
+        Pack(safes, sort (make_ord less) (ths'@unsafes))
     end;
 
 fun merge_pack (Pack(safes,unsafes), Pack(safes',unsafes')) =
         Pack(sort (make_ord less) (safes@safes'), 
-	     sort (make_ord less) (unsafes@unsafes'));
+             sort (make_ord less) (unsafes@unsafes'));
 
 
 fun print_pack (Pack(safes,unsafes)) =
@@ -75,10 +75,10 @@
   case (prem,conc) of
       (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp),
        _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) =>
-	  could_res (leftp,leftc) andalso could_res (rightp,rightc)
+          could_res (leftp,leftc) andalso could_res (rightp,rightc)
     | (_ $ Abs(_,_,leftp) $ rightp,
        _ $ Abs(_,_,leftc) $ rightc) =>
-	  could_res (leftp,leftc)  andalso  Term.could_unify (rightp,rightc)
+          could_res (leftp,leftc)  andalso  Term.could_unify (rightp,rightc)
     | _ => false;
 
 
@@ -88,10 +88,10 @@
 fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) =>
   let val rls = filter_thms could_resolve_seq (maxr+1, prem, rules)
   in  if length rls > maxr  then  no_tac
-	  else (*((rtac derelict 1 THEN rtac impl 1
-		 THEN (rtac identity 2 ORELSE rtac ll_mp 2)
-		 THEN rtac context1 1)
-		 ORELSE *) resolve_tac rls i
+          else (*((rtac derelict 1 THEN rtac impl 1
+                 THEN (rtac identity 2 ORELSE rtac ll_mp 2)
+                 THEN rtac context1 1)
+                 ORELSE *) resolve_tac rls i
   end);
 
 
@@ -112,12 +112,12 @@
 fun RESOLVE_THEN rules =
   let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules;
       fun tac nextac i state = state |>
-	     (filseq_resolve_tac rls0 9999 i 
-	      ORELSE
-	      (DETERM(filseq_resolve_tac rls1 9999 i) THEN  TRY(nextac i))
-	      ORELSE
-	      (DETERM(filseq_resolve_tac rls2 9999 i) THEN  TRY(nextac(i+1))
-					    THEN  TRY(nextac i)))
+             (filseq_resolve_tac rls0 9999 i 
+              ORELSE
+              (DETERM(filseq_resolve_tac rls1 9999 i) THEN  TRY(nextac i))
+              ORELSE
+              (DETERM(filseq_resolve_tac rls2 9999 i) THEN  TRY(nextac(i+1))
+                                            THEN  TRY(nextac i)))
   in  tac  end;
 
 
@@ -133,9 +133,9 @@
   let val restac  =    RESOLVE_THEN safes
       and lastrestac = RESOLVE_THEN unsafes;
       fun gtac i = restac gtac i  
-	           ORELSE  (if !trace then
-				(print_tac "" THEN lastrestac gtac i)
-			    else lastrestac gtac i)
+                   ORELSE  (if !trace then
+                                (print_tac "" THEN lastrestac gtac i)
+                            else lastrestac gtac i)
   in  gtac  end; 
 
 
@@ -164,7 +164,7 @@
 
 fun best_tac pack  = 
   SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) 
-	       (step_tac pack 1));
+               (step_tac pack 1));
 
 end;
 
--- a/src/Sequents/simpdata.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Sequents/simpdata.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -49,7 +49,7 @@
 
 (*Congruence rules for = or <-> (instead of ==)*)
 fun mk_meta_cong rl =
-  standard(mk_meta_eq (mk_meta_prems rl))
+  Drule.standard(mk_meta_eq (mk_meta_prems rl))
   handle THM _ =>
   error("Premises and conclusion of congruence rules must use =-equality or <->");
 
--- a/src/Tools/Compute_Oracle/am.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Tools/Compute_Oracle/am.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -58,7 +58,7 @@
 fun forall_consts pred (Const c) = pred c
   | forall_consts pred (Var x) = true
   | forall_consts pred (App (u,v)) = forall_consts pred u 
-				     andalso forall_consts pred v
+                                     andalso forall_consts pred v
   | forall_consts pred (Abs m) = forall_consts pred m
   | forall_consts pred (Computed t) = forall_consts pred t
 
--- a/src/Tools/Compute_Oracle/am_compiler.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Tools/Compute_Oracle/am_compiler.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -29,45 +29,45 @@
       List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps
 
 fun print_rule (p, t) = 
-    let	
-	fun str x = Int.toString x		    
-	fun print_pattern n PVar = (n+1, "x"^(str n))
-	  | print_pattern n (PConst (c, [])) = (n, "c"^(str c))
-	  | print_pattern n (PConst (c, args)) = 
-	    let
-		val h = print_pattern n (PConst (c,[]))
-	    in
-		print_pattern_list h args
-	    end
-	and print_pattern_list r [] = r
-	  | print_pattern_list (n, p) (t::ts) = 
-	    let
-		val (n, t) = print_pattern n t
-	    in
-		print_pattern_list (n, "App ("^p^", "^t^")") ts
-	    end
+    let
+        fun str x = Int.toString x
+        fun print_pattern n PVar = (n+1, "x"^(str n))
+          | print_pattern n (PConst (c, [])) = (n, "c"^(str c))
+          | print_pattern n (PConst (c, args)) = 
+            let
+                val h = print_pattern n (PConst (c,[]))
+            in
+                print_pattern_list h args
+            end
+        and print_pattern_list r [] = r
+          | print_pattern_list (n, p) (t::ts) = 
+            let
+                val (n, t) = print_pattern n t
+            in
+                print_pattern_list (n, "App ("^p^", "^t^")") ts
+            end
 
-	val (n, pattern) = print_pattern 0 p
-	val pattern =
+        val (n, pattern) = print_pattern 0 p
+        val pattern =
             if exists_string Symbol.is_ascii_blank pattern then "(" ^ pattern ^")"
             else pattern
-	
-	fun print_term d (Var x) = (*if x < d then "Var "^(str x) else "x"^(str (n-(x-d)-1))*)
+        
+        fun print_term d (Var x) = (*if x < d then "Var "^(str x) else "x"^(str (n-(x-d)-1))*)
               "Var " ^ str x
-	  | print_term d (Const c) = "c" ^ str c
-	  | print_term d (App (a,b)) = "App (" ^ print_term d a ^ ", " ^ print_term d b ^ ")"
-	  | print_term d (Abs c) = "Abs (" ^ print_term (d + 1) c ^ ")"
+          | print_term d (Const c) = "c" ^ str c
+          | print_term d (App (a,b)) = "App (" ^ print_term d a ^ ", " ^ print_term d b ^ ")"
+          | print_term d (Abs c) = "Abs (" ^ print_term (d + 1) c ^ ")"
           | print_term d (Computed c) = print_term d c
 
-	fun listvars n = if n = 0 then "x0" else "x"^(str n)^", "^(listvars (n-1))
+        fun listvars n = if n = 0 then "x0" else "x"^(str n)^", "^(listvars (n-1))
 
-	val term = print_term 0 t
-	val term =
+        val term = print_term 0 t
+        val term =
             if n > 0 then "Closure (["^(listvars (n-1))^"], "^term^")"
             else "Closure ([], "^term^")"
-			   
+                           
     in
-	"  | weak_reduce (false, stack, "^pattern^") = Continue (false, stack, "^term^")"
+        "  | weak_reduce (false, stack, "^pattern^") = Continue (false, stack, "^term^")"
     end
 
 fun constants_of PVar = []
@@ -81,131 +81,131 @@
     
 fun load_rules sname name prog = 
     let
-	val buffer = Unsynchronized.ref ""
-	fun write s = (buffer := (!buffer)^s)
-	fun writeln s = (write s; write "\n")
-	fun writelist [] = ()
-	  | writelist (s::ss) = (writeln s; writelist ss)
-	fun str i = Int.toString i
-	val _ = writelist [
-		"structure "^name^" = struct",
-		"",
-		"datatype term = Dummy | App of term * term | Abs of term | Var of int | Const of int | Closure of term list * term"]
-	val constants = distinct (op =) (maps (fn (p, r) => ((constants_of p)@(constants_of_term r))) prog)
-	val _ = map (fn x => write (" | c"^(str x))) constants
-	val _ = writelist [
-		"",
-		"datatype stack = SEmpty | SAppL of term * stack | SAppR of term * stack | SAbs of stack",
-		"",
-		"type state = bool * stack * term",
-		"",
-		"datatype loopstate = Continue of state | Stop of stack * term",
-		"",
-		"fun proj_C (Continue s) = s",
-		"  | proj_C _ = raise Match",
-		"",
-		"fun proj_S (Stop s) = s",
-		"  | proj_S _ = raise Match",
-		"",
-		"fun cont (Continue _) = true",
-		"  | cont _ = false",
-		"",
-		"fun do_reduction reduce p =",
-		"    let",
-		"       val s = Unsynchronized.ref (Continue p)",
-		"       val _ = while cont (!s) do (s := reduce (proj_C (!s)))",
-		"   in",
-		"       proj_S (!s)",
-		"   end",
-		""]
+        val buffer = Unsynchronized.ref ""
+        fun write s = (buffer := (!buffer)^s)
+        fun writeln s = (write s; write "\n")
+        fun writelist [] = ()
+          | writelist (s::ss) = (writeln s; writelist ss)
+        fun str i = Int.toString i
+        val _ = writelist [
+                "structure "^name^" = struct",
+                "",
+                "datatype term = Dummy | App of term * term | Abs of term | Var of int | Const of int | Closure of term list * term"]
+        val constants = distinct (op =) (maps (fn (p, r) => ((constants_of p)@(constants_of_term r))) prog)
+        val _ = map (fn x => write (" | c"^(str x))) constants
+        val _ = writelist [
+                "",
+                "datatype stack = SEmpty | SAppL of term * stack | SAppR of term * stack | SAbs of stack",
+                "",
+                "type state = bool * stack * term",
+                "",
+                "datatype loopstate = Continue of state | Stop of stack * term",
+                "",
+                "fun proj_C (Continue s) = s",
+                "  | proj_C _ = raise Match",
+                "",
+                "fun proj_S (Stop s) = s",
+                "  | proj_S _ = raise Match",
+                "",
+                "fun cont (Continue _) = true",
+                "  | cont _ = false",
+                "",
+                "fun do_reduction reduce p =",
+                "    let",
+                "       val s = Unsynchronized.ref (Continue p)",
+                "       val _ = while cont (!s) do (s := reduce (proj_C (!s)))",
+                "   in",
+                "       proj_S (!s)",
+                "   end",
+                ""]
 
-	val _ = writelist [
-		"fun weak_reduce (false, stack, Closure (e, App (a, b))) = Continue (false, SAppL (Closure (e, b), stack), Closure (e, a))",
-		"  | weak_reduce (false, SAppL (b, stack), Closure (e, Abs m)) = Continue (false, stack, Closure (b::e, m))",
-		"  | weak_reduce (false, stack, c as Closure (e, Abs m)) = Continue (true, stack, c)",
-		"  | weak_reduce (false, stack, Closure (e, Var n)) = Continue (false, stack, case "^sname^".list_nth (e, n) of Dummy => Var n | r => r)",
-		"  | weak_reduce (false, stack, Closure (e, c)) = Continue (false, stack, c)"]
-	val _ = writelist (map print_rule prog)
+        val _ = writelist [
+                "fun weak_reduce (false, stack, Closure (e, App (a, b))) = Continue (false, SAppL (Closure (e, b), stack), Closure (e, a))",
+                "  | weak_reduce (false, SAppL (b, stack), Closure (e, Abs m)) = Continue (false, stack, Closure (b::e, m))",
+                "  | weak_reduce (false, stack, c as Closure (e, Abs m)) = Continue (true, stack, c)",
+                "  | weak_reduce (false, stack, Closure (e, Var n)) = Continue (false, stack, case "^sname^".list_nth (e, n) of Dummy => Var n | r => r)",
+                "  | weak_reduce (false, stack, Closure (e, c)) = Continue (false, stack, c)"]
+        val _ = writelist (map print_rule prog)
         val _ = writelist [
-		"  | weak_reduce (false, stack, clos) = Continue (true, stack, clos)",
-		"  | weak_reduce (true, SAppR (a, stack), b) = Continue (false, stack, App (a,b))",
-		"  | weak_reduce (true, s as (SAppL (b, stack)), a) = Continue (false, SAppR (a, stack), b)",
-		"  | weak_reduce (true, stack, c) = Stop (stack, c)",
-		"",
-		"fun strong_reduce (false, stack, Closure (e, Abs m)) =",
-		"    let",
-		"        val (stack', wnf) = do_reduction weak_reduce (false, SEmpty, Closure (Dummy::e, m))",
-		"    in",
-		"        case stack' of",
-		"            SEmpty => Continue (false, SAbs stack, wnf)",
-		"          | _ => raise ("^sname^".Run \"internal error in strong: weak failed\")",
-		"    end",		
-		"  | strong_reduce (false, stack, clos as (App (u, v))) = Continue (false, SAppL (v, stack), u)",
-		"  | strong_reduce (false, stack, clos) = Continue (true, stack, clos)",
-		"  | strong_reduce (true, SAbs stack, m) = Continue (false, stack, Abs m)",
-		"  | strong_reduce (true, SAppL (b, stack), a) = Continue (false, SAppR (a, stack), b)",
-		"  | strong_reduce (true, SAppR (a, stack), b) = Continue (true, stack, App (a, b))",
-		"  | strong_reduce (true, stack, clos) = Stop (stack, clos)",
-		""]
-	
-	val ic = "(case c of "^(implode (map (fn c => (str c)^" => c"^(str c)^" | ") constants))^" _ => Const c)"						  	
-	val _ = writelist [
-		"fun importTerm ("^sname^".Var x) = Var x",
-		"  | importTerm ("^sname^".Const c) =  "^ic,
-		"  | importTerm ("^sname^".App (a, b)) = App (importTerm a, importTerm b)",
-		"  | importTerm ("^sname^".Abs m) = Abs (importTerm m)",
-		""]
+                "  | weak_reduce (false, stack, clos) = Continue (true, stack, clos)",
+                "  | weak_reduce (true, SAppR (a, stack), b) = Continue (false, stack, App (a,b))",
+                "  | weak_reduce (true, s as (SAppL (b, stack)), a) = Continue (false, SAppR (a, stack), b)",
+                "  | weak_reduce (true, stack, c) = Stop (stack, c)",
+                "",
+                "fun strong_reduce (false, stack, Closure (e, Abs m)) =",
+                "    let",
+                "        val (stack', wnf) = do_reduction weak_reduce (false, SEmpty, Closure (Dummy::e, m))",
+                "    in",
+                "        case stack' of",
+                "            SEmpty => Continue (false, SAbs stack, wnf)",
+                "          | _ => raise ("^sname^".Run \"internal error in strong: weak failed\")",
+                "    end",              
+                "  | strong_reduce (false, stack, clos as (App (u, v))) = Continue (false, SAppL (v, stack), u)",
+                "  | strong_reduce (false, stack, clos) = Continue (true, stack, clos)",
+                "  | strong_reduce (true, SAbs stack, m) = Continue (false, stack, Abs m)",
+                "  | strong_reduce (true, SAppL (b, stack), a) = Continue (false, SAppR (a, stack), b)",
+                "  | strong_reduce (true, SAppR (a, stack), b) = Continue (true, stack, App (a, b))",
+                "  | strong_reduce (true, stack, clos) = Stop (stack, clos)",
+                ""]
+        
+        val ic = "(case c of "^(implode (map (fn c => (str c)^" => c"^(str c)^" | ") constants))^" _ => Const c)"                                                       
+        val _ = writelist [
+                "fun importTerm ("^sname^".Var x) = Var x",
+                "  | importTerm ("^sname^".Const c) =  "^ic,
+                "  | importTerm ("^sname^".App (a, b)) = App (importTerm a, importTerm b)",
+                "  | importTerm ("^sname^".Abs m) = Abs (importTerm m)",
+                ""]
 
-	fun ec c = "  | exportTerm c"^(str c)^" = "^sname^".Const "^(str c)
-	val _ = writelist [
-		"fun exportTerm (Var x) = "^sname^".Var x",
-		"  | exportTerm (Const c) = "^sname^".Const c",
-		"  | exportTerm (App (a,b)) = "^sname^".App (exportTerm a, exportTerm b)",
-		"  | exportTerm (Abs m) = "^sname^".Abs (exportTerm m)",
-		"  | exportTerm (Closure (closlist, clos)) = raise ("^sname^".Run \"internal error, cannot export Closure\")",
-		"  | exportTerm Dummy = raise ("^sname^".Run \"internal error, cannot export Dummy\")"]
-	val _ = writelist (map ec constants)
-		
-	val _ = writelist [
-		"",
-		"fun rewrite t = ",
-		"    let",
-		"      val (stack, wnf) = do_reduction weak_reduce (false, SEmpty, Closure ([], importTerm t))",
-		"    in",
-		"      case stack of ",
-		"           SEmpty => (case do_reduction strong_reduce (false, SEmpty, wnf) of",
-		"                          (SEmpty, snf) => exportTerm snf",
-		"                        | _ => raise ("^sname^".Run \"internal error in rewrite: strong failed\"))",
-		"         | _ => (raise ("^sname^".Run \"internal error in rewrite: weak failed\"))",
-		"    end",
-		"",
-		"val _ = "^sname^".set_compiled_rewriter rewrite",
-		"",
-		"end;"]
+        fun ec c = "  | exportTerm c"^(str c)^" = "^sname^".Const "^(str c)
+        val _ = writelist [
+                "fun exportTerm (Var x) = "^sname^".Var x",
+                "  | exportTerm (Const c) = "^sname^".Const c",
+                "  | exportTerm (App (a,b)) = "^sname^".App (exportTerm a, exportTerm b)",
+                "  | exportTerm (Abs m) = "^sname^".Abs (exportTerm m)",
+                "  | exportTerm (Closure (closlist, clos)) = raise ("^sname^".Run \"internal error, cannot export Closure\")",
+                "  | exportTerm Dummy = raise ("^sname^".Run \"internal error, cannot export Dummy\")"]
+        val _ = writelist (map ec constants)
+                
+        val _ = writelist [
+                "",
+                "fun rewrite t = ",
+                "    let",
+                "      val (stack, wnf) = do_reduction weak_reduce (false, SEmpty, Closure ([], importTerm t))",
+                "    in",
+                "      case stack of ",
+                "           SEmpty => (case do_reduction strong_reduce (false, SEmpty, wnf) of",
+                "                          (SEmpty, snf) => exportTerm snf",
+                "                        | _ => raise ("^sname^".Run \"internal error in rewrite: strong failed\"))",
+                "         | _ => (raise ("^sname^".Run \"internal error in rewrite: weak failed\"))",
+                "    end",
+                "",
+                "val _ = "^sname^".set_compiled_rewriter rewrite",
+                "",
+                "end;"]
 
     in
-	compiled_rewriter := NONE;	
-	use_text ML_Env.local_context (1, "") false (!buffer);
-	case !compiled_rewriter of 
-	    NONE => raise (Compile "cannot communicate with compiled function")
-	  | SOME r => (compiled_rewriter := NONE; r)
-    end	
+        compiled_rewriter := NONE;      
+        use_text ML_Env.local_context (1, "") false (!buffer);
+        case !compiled_rewriter of 
+            NONE => raise (Compile "cannot communicate with compiled function")
+          | SOME r => (compiled_rewriter := NONE; r)
+    end 
 
 fun compile cache_patterns const_arity eqs = 
     let
-	val _ = if exists (fn (a,b,c) => not (null a)) eqs then raise Compile ("cannot deal with guards") else ()
-	val eqs = map (fn (a,b,c) => (b,c)) eqs
-	fun check (p, r) = if check_freevars (count_patternvars p) r then () else raise Compile ("unbound variables in rule") 
-	val _ = map (fn (p, r) => 
+        val _ = if exists (fn (a,b,c) => not (null a)) eqs then raise Compile ("cannot deal with guards") else ()
+        val eqs = map (fn (a,b,c) => (b,c)) eqs
+        fun check (p, r) = if check_freevars (count_patternvars p) r then () else raise Compile ("unbound variables in rule") 
+        val _ = map (fn (p, r) => 
                   (check (p, r); 
                    case p of PVar => raise (Compile "pattern is just a variable") | _ => ())) eqs
     in
-	load_rules "AM_Compiler" "AM_compiled_code" eqs
-    end	
+        load_rules "AM_Compiler" "AM_compiled_code" eqs
+    end 
 
 fun run prog t = (prog t)
 
 fun discard p = ()
-			 	  
+                                  
 end
 
--- a/src/Tools/Compute_Oracle/am_ghc.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Tools/Compute_Oracle/am_ghc.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -14,7 +14,7 @@
 
 fun update_arity arity code a = 
     (case Inttab.lookup arity code of
-	 NONE => Inttab.update_new (code, a) arity
+         NONE => Inttab.update_new (code, a) arity
        | SOME (a': int) => if a > a' then Inttab.update (code, a) arity else arity)
 
 (* We have to find out the maximal arity of each constant *)
@@ -39,65 +39,65 @@
 
 fun adjust_rules rules =
     let
-	val arity = fold (fn (p, t) => fn arity => collect_term_arity t (collect_pattern_arity p arity)) rules Inttab.empty
-	fun arity_of c = the (Inttab.lookup arity c)
-	fun adjust_pattern PVar = PVar
-	  | adjust_pattern (C as PConst (c, args)) = if (length args <> arity_of c) then raise Compile ("Constant inside pattern must have maximal arity") else C
-	fun adjust_rule (PVar, t) = raise Compile ("pattern may not be a variable")
-	  | adjust_rule (rule as (p as PConst (c, args),t)) = 
-	    let
-		val _ = if not (check_freevars (count_patternvars p) t) then raise Compile ("unbound variables on right hand side") else () 
-		val args = map adjust_pattern args	        
-		val len = length args
-		val arity = arity_of c
-		fun lift level n (Var m) = if m < level then Var m else Var (m+n) 
-		  | lift level n (Const c) = Const c
-		  | lift level n (App (a,b)) = App (lift level n a, lift level n b)
-		  | lift level n (Abs b) = Abs (lift (level+1) n b)
-		val lift = lift 0
-		fun adjust_term n t = if n=0 then t else adjust_term (n-1) (App (t, Var (n-1))) 
-	    in
-		if len = arity then
-		    rule
-		else if arity >= len then  
-		    (PConst (c, args @ (rep (arity-len) PVar)), adjust_term (arity-len) (lift (arity-len) t))
-		else (raise Compile "internal error in adjust_rule")
-	    end
+        val arity = fold (fn (p, t) => fn arity => collect_term_arity t (collect_pattern_arity p arity)) rules Inttab.empty
+        fun arity_of c = the (Inttab.lookup arity c)
+        fun adjust_pattern PVar = PVar
+          | adjust_pattern (C as PConst (c, args)) = if (length args <> arity_of c) then raise Compile ("Constant inside pattern must have maximal arity") else C
+        fun adjust_rule (PVar, t) = raise Compile ("pattern may not be a variable")
+          | adjust_rule (rule as (p as PConst (c, args),t)) = 
+            let
+                val _ = if not (check_freevars (count_patternvars p) t) then raise Compile ("unbound variables on right hand side") else () 
+                val args = map adjust_pattern args              
+                val len = length args
+                val arity = arity_of c
+                fun lift level n (Var m) = if m < level then Var m else Var (m+n) 
+                  | lift level n (Const c) = Const c
+                  | lift level n (App (a,b)) = App (lift level n a, lift level n b)
+                  | lift level n (Abs b) = Abs (lift (level+1) n b)
+                val lift = lift 0
+                fun adjust_term n t = if n=0 then t else adjust_term (n-1) (App (t, Var (n-1))) 
+            in
+                if len = arity then
+                    rule
+                else if arity >= len then  
+                    (PConst (c, args @ (rep (arity-len) PVar)), adjust_term (arity-len) (lift (arity-len) t))
+                else (raise Compile "internal error in adjust_rule")
+            end
     in
-	(arity, map adjust_rule rules)
-    end		    
+        (arity, map adjust_rule rules)
+    end             
 
 fun print_term arity_of n =
 let
     fun str x = string_of_int x
     fun protect_blank s = if exists_string Symbol.is_ascii_blank s then "(" ^ s ^")" else s
-											  
+                                                                                          
     fun print_apps d f [] = f
       | print_apps d f (a::args) = print_apps d ("app "^(protect_blank f)^" "^(protect_blank (print_term d a))) args
     and print_call d (App (a, b)) args = print_call d a (b::args) 
       | print_call d (Const c) args = 
-	(case arity_of c of 
-	     NONE => print_apps d ("Const "^(str c)) args 
-	   | SOME a =>
-	     let
-		 val len = length args
-	     in
-		 if a <= len then 
-		     let
-			 val s = "c"^(str c)^(implode (map (fn t => " "^(protect_blank (print_term d t))) (List.take (args, a))))
-		     in
-			 print_apps d s (List.drop (args, a))
-		     end
-		 else 
-		     let
-			 fun mk_apps n t = if n = 0 then t else mk_apps (n-1) (App (t, Var (n-1)))
-			 fun mk_lambdas n t = if n = 0 then t else mk_lambdas (n-1) (Abs t)
-			 fun append_args [] t = t
-			   | append_args (c::cs) t = append_args cs (App (t, c))
-		     in
-			 print_term d (mk_lambdas (a-len) (mk_apps (a-len) (nlift 0 (a-len) (append_args args (Const c)))))
-		     end
-	     end)
+        (case arity_of c of 
+             NONE => print_apps d ("Const "^(str c)) args 
+           | SOME a =>
+             let
+                 val len = length args
+             in
+                 if a <= len then 
+                     let
+                         val s = "c"^(str c)^(implode (map (fn t => " "^(protect_blank (print_term d t))) (List.take (args, a))))
+                     in
+                         print_apps d s (List.drop (args, a))
+                     end
+                 else 
+                     let
+                         fun mk_apps n t = if n = 0 then t else mk_apps (n-1) (App (t, Var (n-1)))
+                         fun mk_lambdas n t = if n = 0 then t else mk_lambdas (n-1) (Abs t)
+                         fun append_args [] t = t
+                           | append_args (c::cs) t = append_args cs (App (t, c))
+                     in
+                         print_term d (mk_lambdas (a-len) (mk_apps (a-len) (nlift 0 (a-len) (append_args args (Const c)))))
+                     end
+             end)
       | print_call d t args = print_apps d (print_term d t) args
     and print_term d (Var x) = if x < d then "b"^(str (d-x-1)) else "x"^(str (n-(x-d)-1))
       | print_term d (Abs c) = "Abs (\\b"^(str d)^" -> "^(print_term (d + 1) c)^")"
@@ -105,108 +105,108 @@
 in
     print_term 0 
 end
-			 			
+                                                
 fun print_rule arity_of (p, t) = 
-    let	
-	fun str x = Int.toString x		    
-	fun print_pattern top n PVar = (n+1, "x"^(str n))
-	  | print_pattern top n (PConst (c, [])) = (n, (if top then "c" else "C")^(str c))
-	  | print_pattern top n (PConst (c, args)) = 
-	    let
-		val (n,s) = print_pattern_list (n, (if top then "c" else "C")^(str c)) args
-	    in
-		(n, if top then s else "("^s^")")
-	    end
-	and print_pattern_list r [] = r
-	  | print_pattern_list (n, p) (t::ts) = 
-	    let
-		val (n, t) = print_pattern false n t
-	    in
-		print_pattern_list (n, p^" "^t) ts
-	    end
-	val (n, pattern) = print_pattern true 0 p
+    let 
+        fun str x = Int.toString x                  
+        fun print_pattern top n PVar = (n+1, "x"^(str n))
+          | print_pattern top n (PConst (c, [])) = (n, (if top then "c" else "C")^(str c))
+          | print_pattern top n (PConst (c, args)) = 
+            let
+                val (n,s) = print_pattern_list (n, (if top then "c" else "C")^(str c)) args
+            in
+                (n, if top then s else "("^s^")")
+            end
+        and print_pattern_list r [] = r
+          | print_pattern_list (n, p) (t::ts) = 
+            let
+                val (n, t) = print_pattern false n t
+            in
+                print_pattern_list (n, p^" "^t) ts
+            end
+        val (n, pattern) = print_pattern true 0 p
     in
-	pattern^" = "^(print_term arity_of n t)	
+        pattern^" = "^(print_term arity_of n t) 
     end
 
 fun group_rules rules =
     let
-	fun add_rule (r as (PConst (c,_), _)) groups =
-	    let
-		val rs = (case Inttab.lookup groups c of NONE => [] | SOME rs => rs)
-	    in
-		Inttab.update (c, r::rs) groups
-	    end
-	  | add_rule _ _ = raise Compile "internal error group_rules"
+        fun add_rule (r as (PConst (c,_), _)) groups =
+            let
+                val rs = (case Inttab.lookup groups c of NONE => [] | SOME rs => rs)
+            in
+                Inttab.update (c, r::rs) groups
+            end
+          | add_rule _ _ = raise Compile "internal error group_rules"
     in
-	fold_rev add_rule rules Inttab.empty
+        fold_rev add_rule rules Inttab.empty
     end
 
 fun haskell_prog name rules = 
     let
-	val buffer = Unsynchronized.ref ""
-	fun write s = (buffer := (!buffer)^s)
-	fun writeln s = (write s; write "\n")
-	fun writelist [] = ()
-	  | writelist (s::ss) = (writeln s; writelist ss)
-	fun str i = Int.toString i
-	val (arity, rules) = adjust_rules rules
-	val rules = group_rules rules
-	val constants = Inttab.keys arity
-	fun arity_of c = Inttab.lookup arity c
-	fun rep_str s n = implode (rep n s)
-	fun indexed s n = s^(str n)
-	fun section n = if n = 0 then [] else (section (n-1))@[n-1]
-	fun make_show c = 
-	    let
-		val args = section (the (arity_of c))
-	    in
-		"  show ("^(indexed "C" c)^(implode (map (indexed " a") args))^") = "
-		^"\""^(indexed "C" c)^"\""^(implode (map (fn a => "++(show "^(indexed "a" a)^")") args))
-	    end
-	fun default_case c = 
-	    let
-		val args = implode (map (indexed " x") (section (the (arity_of c))))
-	    in
-		(indexed "c" c)^args^" = "^(indexed "C" c)^args
-	    end
-	val _ = writelist [        
-		"module "^name^" where",
-		"",
-		"data Term = Const Integer | App Term Term | Abs (Term -> Term)",
-		"         "^(implode (map (fn c => " | C"^(str c)^(rep_str " Term" (the (arity_of c)))) constants)),
-		"",
-		"instance Show Term where"]
-	val _ = writelist (map make_show constants)
-	val _ = writelist [
-		"  show (Const c) = \"c\"++(show c)",
-		"  show (App a b) = \"A\"++(show a)++(show b)",
-		"  show (Abs _) = \"L\"",
-		""]
-	val _ = writelist [
-		"app (Abs a) b = a b",
-		"app a b = App a b",
-		"",
-		"calc s c = writeFile s (show c)",
-		""]
-	fun list_group c = (writelist (case Inttab.lookup rules c of 
-					   NONE => [default_case c, ""] 
-					 | SOME (rs as ((PConst (_, []), _)::rs')) => 
-					   if not (null rs') then raise Compile "multiple declaration of constant"
-					   else (map (print_rule arity_of) rs) @ [""]
-					 | SOME rs => (map (print_rule arity_of) rs) @ [default_case c, ""]))
-	val _ = map list_group constants
+        val buffer = Unsynchronized.ref ""
+        fun write s = (buffer := (!buffer)^s)
+        fun writeln s = (write s; write "\n")
+        fun writelist [] = ()
+          | writelist (s::ss) = (writeln s; writelist ss)
+        fun str i = Int.toString i
+        val (arity, rules) = adjust_rules rules
+        val rules = group_rules rules
+        val constants = Inttab.keys arity
+        fun arity_of c = Inttab.lookup arity c
+        fun rep_str s n = implode (rep n s)
+        fun indexed s n = s^(str n)
+        fun section n = if n = 0 then [] else (section (n-1))@[n-1]
+        fun make_show c = 
+            let
+                val args = section (the (arity_of c))
+            in
+                "  show ("^(indexed "C" c)^(implode (map (indexed " a") args))^") = "
+                ^"\""^(indexed "C" c)^"\""^(implode (map (fn a => "++(show "^(indexed "a" a)^")") args))
+            end
+        fun default_case c = 
+            let
+                val args = implode (map (indexed " x") (section (the (arity_of c))))
+            in
+                (indexed "c" c)^args^" = "^(indexed "C" c)^args
+            end
+        val _ = writelist [        
+                "module "^name^" where",
+                "",
+                "data Term = Const Integer | App Term Term | Abs (Term -> Term)",
+                "         "^(implode (map (fn c => " | C"^(str c)^(rep_str " Term" (the (arity_of c)))) constants)),
+                "",
+                "instance Show Term where"]
+        val _ = writelist (map make_show constants)
+        val _ = writelist [
+                "  show (Const c) = \"c\"++(show c)",
+                "  show (App a b) = \"A\"++(show a)++(show b)",
+                "  show (Abs _) = \"L\"",
+                ""]
+        val _ = writelist [
+                "app (Abs a) b = a b",
+                "app a b = App a b",
+                "",
+                "calc s c = writeFile s (show c)",
+                ""]
+        fun list_group c = (writelist (case Inttab.lookup rules c of 
+                                           NONE => [default_case c, ""] 
+                                         | SOME (rs as ((PConst (_, []), _)::rs')) => 
+                                           if not (null rs') then raise Compile "multiple declaration of constant"
+                                           else (map (print_rule arity_of) rs) @ [""]
+                                         | SOME rs => (map (print_rule arity_of) rs) @ [default_case c, ""]))
+        val _ = map list_group constants
     in
-	(arity, !buffer)
+        (arity, !buffer)
     end
 
 val guid_counter = Unsynchronized.ref 0
 fun get_guid () = 
     let
-	val c = !guid_counter
-	val _ = guid_counter := !guid_counter + 1
+        val c = !guid_counter
+        val _ = guid_counter := !guid_counter + 1
     in
-	(LargeInt.toString (Time.toMicroseconds (Time.now ()))) ^ (string_of_int c)
+        (LargeInt.toString (Time.toMicroseconds (Time.now ()))) ^ (string_of_int c)
     end
 
 fun tmp_file s = Path.implode (Path.expand (File.tmp_path (Path.make [s])));
@@ -220,106 +220,106 @@
 
 fun compile cache_patterns const_arity eqs = 
     let
-	val _ = if exists (fn (a,b,c) => not (null a)) eqs then raise Compile ("cannot deal with guards") else ()
-	val eqs = map (fn (a,b,c) => (b,c)) eqs
-	val guid = get_guid ()
-	val module = "AMGHC_Prog_"^guid
-	val (arity, source) = haskell_prog module eqs
-	val module_file = tmp_file (module^".hs")
-	val object_file = tmp_file (module^".o")
-	val _ = writeTextFile module_file source
-	val _ = system ((!ghc)^" -c "^module_file)
-	val _ = if not (fileExists object_file) then raise Compile ("Failure compiling haskell code (GHC_PATH = '"^(!ghc)^"')") else ()
+        val _ = if exists (fn (a,b,c) => not (null a)) eqs then raise Compile ("cannot deal with guards") else ()
+        val eqs = map (fn (a,b,c) => (b,c)) eqs
+        val guid = get_guid ()
+        val module = "AMGHC_Prog_"^guid
+        val (arity, source) = haskell_prog module eqs
+        val module_file = tmp_file (module^".hs")
+        val object_file = tmp_file (module^".o")
+        val _ = writeTextFile module_file source
+        val _ = system ((!ghc)^" -c "^module_file)
+        val _ = if not (fileExists object_file) then raise Compile ("Failure compiling haskell code (GHC_PATH = '"^(!ghc)^"')") else ()
     in
-	(guid, module_file, arity)	
+        (guid, module_file, arity)      
     end
 
 fun readResultFile name = File.read (Path.explode name) 
-			  			  			      			    
+                                                                                                    
 fun parse_result arity_of result =
     let
-	val result = String.explode result
-	fun shift NONE x = SOME x
-	  | shift (SOME y) x = SOME (y*10 + x)
-	fun parse_int' x (#"0"::rest) = parse_int' (shift x 0) rest
-	  | parse_int' x (#"1"::rest) = parse_int' (shift x 1) rest
-	  | parse_int' x (#"2"::rest) = parse_int' (shift x 2) rest
-	  | parse_int' x (#"3"::rest) = parse_int' (shift x 3) rest
-	  | parse_int' x (#"4"::rest) = parse_int' (shift x 4) rest
-	  | parse_int' x (#"5"::rest) = parse_int' (shift x 5) rest
-	  | parse_int' x (#"6"::rest) = parse_int' (shift x 6) rest
-	  | parse_int' x (#"7"::rest) = parse_int' (shift x 7) rest
-	  | parse_int' x (#"8"::rest) = parse_int' (shift x 8) rest
-	  | parse_int' x (#"9"::rest) = parse_int' (shift x 9) rest
-	  | parse_int' x rest = (x, rest)
-	fun parse_int rest = parse_int' NONE rest
+        val result = String.explode result
+        fun shift NONE x = SOME x
+          | shift (SOME y) x = SOME (y*10 + x)
+        fun parse_int' x (#"0"::rest) = parse_int' (shift x 0) rest
+          | parse_int' x (#"1"::rest) = parse_int' (shift x 1) rest
+          | parse_int' x (#"2"::rest) = parse_int' (shift x 2) rest
+          | parse_int' x (#"3"::rest) = parse_int' (shift x 3) rest
+          | parse_int' x (#"4"::rest) = parse_int' (shift x 4) rest
+          | parse_int' x (#"5"::rest) = parse_int' (shift x 5) rest
+          | parse_int' x (#"6"::rest) = parse_int' (shift x 6) rest
+          | parse_int' x (#"7"::rest) = parse_int' (shift x 7) rest
+          | parse_int' x (#"8"::rest) = parse_int' (shift x 8) rest
+          | parse_int' x (#"9"::rest) = parse_int' (shift x 9) rest
+          | parse_int' x rest = (x, rest)
+        fun parse_int rest = parse_int' NONE rest
 
-	fun parse (#"C"::rest) = 
-	    (case parse_int rest of 
-		 (SOME c, rest) => 
-		 let
-		     val (args, rest) = parse_list (the (arity_of c)) rest
-		     fun app_args [] t = t
-		       | app_args (x::xs) t = app_args xs (App (t, x))
-		 in
-		     (app_args args (Const c), rest)
-		 end		     
-	       | (NONE, rest) => raise Run "parse C")
-	  | parse (#"c"::rest) = 
-	    (case parse_int rest of
-		 (SOME c, rest) => (Const c, rest)
-	       | _ => raise Run "parse c")
-	  | parse (#"A"::rest) = 
-	    let
-		val (a, rest) = parse rest
-		val (b, rest) = parse rest
-	    in
-		(App (a,b), rest)
-	    end
-	  | parse (#"L"::rest) = raise Run "there may be no abstraction in the result"
-	  | parse _ = raise Run "invalid result"
-	and parse_list n rest = 
-	    if n = 0 then 
-		([], rest) 
-	    else 
-		let 
-		    val (x, rest) = parse rest
-		    val (xs, rest) = parse_list (n-1) rest
-		in
-		    (x::xs, rest)
-		end
-	val (parsed, rest) = parse result
-	fun is_blank (#" "::rest) = is_blank rest
-	  | is_blank (#"\n"::rest) = is_blank rest
-	  | is_blank [] = true
-	  | is_blank _ = false
+        fun parse (#"C"::rest) = 
+            (case parse_int rest of 
+                 (SOME c, rest) => 
+                 let
+                     val (args, rest) = parse_list (the (arity_of c)) rest
+                     fun app_args [] t = t
+                       | app_args (x::xs) t = app_args xs (App (t, x))
+                 in
+                     (app_args args (Const c), rest)
+                 end                 
+               | (NONE, rest) => raise Run "parse C")
+          | parse (#"c"::rest) = 
+            (case parse_int rest of
+                 (SOME c, rest) => (Const c, rest)
+               | _ => raise Run "parse c")
+          | parse (#"A"::rest) = 
+            let
+                val (a, rest) = parse rest
+                val (b, rest) = parse rest
+            in
+                (App (a,b), rest)
+            end
+          | parse (#"L"::rest) = raise Run "there may be no abstraction in the result"
+          | parse _ = raise Run "invalid result"
+        and parse_list n rest = 
+            if n = 0 then 
+                ([], rest) 
+            else 
+                let 
+                    val (x, rest) = parse rest
+                    val (xs, rest) = parse_list (n-1) rest
+                in
+                    (x::xs, rest)
+                end
+        val (parsed, rest) = parse result
+        fun is_blank (#" "::rest) = is_blank rest
+          | is_blank (#"\n"::rest) = is_blank rest
+          | is_blank [] = true
+          | is_blank _ = false
     in
-	if is_blank rest then parsed else raise Run "non-blank suffix in result file"	
+        if is_blank rest then parsed else raise Run "non-blank suffix in result file"   
     end
 
 fun run (guid, module_file, arity) t = 
     let
-	val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms")
-	fun arity_of c = Inttab.lookup arity c			 
-	val callguid = get_guid()
-	val module = "AMGHC_Prog_"^guid
-	val call = module^"_Call_"^callguid
-	val result_file = tmp_file (module^"_Result_"^callguid^".txt")
-	val call_file = tmp_file (call^".hs")
-	val term = print_term arity_of 0 t
-	val call_source = "module "^call^" where\n\nimport "^module^"\n\ncall = "^module^".calc \""^result_file^"\" ("^term^")"
-	val _ = writeTextFile call_file call_source
-	val _ = system ((!ghc)^" -e \""^call^".call\" "^module_file^" "^call_file)
-	val result = readResultFile result_file handle IO.Io _ => raise Run ("Failure running haskell compiler (GHC_PATH = '"^(!ghc)^"')")
-	val t' = parse_result arity_of result
-	val _ = OS.FileSys.remove call_file
-	val _ = OS.FileSys.remove result_file
+        val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms")
+        fun arity_of c = Inttab.lookup arity c                   
+        val callguid = get_guid()
+        val module = "AMGHC_Prog_"^guid
+        val call = module^"_Call_"^callguid
+        val result_file = tmp_file (module^"_Result_"^callguid^".txt")
+        val call_file = tmp_file (call^".hs")
+        val term = print_term arity_of 0 t
+        val call_source = "module "^call^" where\n\nimport "^module^"\n\ncall = "^module^".calc \""^result_file^"\" ("^term^")"
+        val _ = writeTextFile call_file call_source
+        val _ = system ((!ghc)^" -e \""^call^".call\" "^module_file^" "^call_file)
+        val result = readResultFile result_file handle IO.Io _ => raise Run ("Failure running haskell compiler (GHC_PATH = '"^(!ghc)^"')")
+        val t' = parse_result arity_of result
+        val _ = OS.FileSys.remove call_file
+        val _ = OS.FileSys.remove result_file
     in
-	t'
+        t'
     end
 
-	
+        
 fun discard _ = ()
-		 	  
+                          
 end
 
--- a/src/Tools/Compute_Oracle/am_interpreter.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Tools/Compute_Oracle/am_interpreter.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -51,11 +51,11 @@
 
 fun resolve (stack, closure) = 
     let
-	val _ = writeln "start resolving"
-	val t =	resolve_stack (resolve_closure' closure) stack
-	val _ = writeln "finished resolving"
+        val _ = writeln "start resolving"
+        val t = resolve_stack (resolve_closure' closure) stack
+        val _ = writeln "finished resolving"
     in
-	t
+        t
     end
 
 fun strip_closure args (CApp (a,b)) = strip_closure (b::args) a
@@ -102,13 +102,13 @@
 
 fun compile cache_patterns const_arity eqs =
     let
-	fun check p r = if check_freevars p r then () else raise Compile ("unbound variables in rule") 
-	fun check_guard p (Guard (a,b)) = (check p a; check p b) 
+        fun check p r = if check_freevars p r then () else raise Compile ("unbound variables in rule") 
+        fun check_guard p (Guard (a,b)) = (check p a; check p b) 
         fun clos_of_guard (Guard (a,b)) = (clos_of_term a, clos_of_term b)
         val eqs = map (fn (guards, p, r) => let val pcount = count_patternvars p val _ = map (check_guard pcount) (guards) val _ = check pcount r in 
                                               (pattern_key p, (p, clos_of_term r, map clos_of_guard guards)) end) eqs
-	fun merge (k, a) table = prog_struct.update (k, case prog_struct.lookup table k of NONE => [a] | SOME l => a::l) table
-	val p = fold merge eqs prog_struct.empty 
+        fun merge (k, a) table = prog_struct.update (k, case prog_struct.lookup table k of NONE => [a] | SOME l => a::l) table
+        val p = fold merge eqs prog_struct.empty 
     in
         Program p
     end
@@ -133,15 +133,15 @@
 
 fun do_reduction reduce p =
     let
-	val s = Unsynchronized.ref (Continue p)
-	val counter = Unsynchronized.ref 0
-	val _ = case !max_reductions of 
-		    NONE => while cont (!s) do (s := reduce (proj_C (!s)))
-		  | SOME m => while cont (!s) andalso (!counter < m) do (s := reduce (proj_C (!s)); counter := (!counter) + 1)
+        val s = Unsynchronized.ref (Continue p)
+        val counter = Unsynchronized.ref 0
+        val _ = case !max_reductions of 
+                    NONE => while cont (!s) do (s := reduce (proj_C (!s)))
+                  | SOME m => while cont (!s) andalso (!counter < m) do (s := reduce (proj_C (!s)); counter := (!counter) + 1)
     in
-	case !max_reductions of
-	    SOME m => if !counter >= m then raise InterruptedExecution (proj_S (!s)) else proj_S (!s)
-	  | NONE => proj_S (!s)
+        case !max_reductions of
+            SOME m => if !counter >= m then raise InterruptedExecution (proj_S (!s)) else proj_S (!s)
+          | NONE => proj_S (!s)
     end
 
 fun match_rules prog n [] clos = NONE
--- a/src/Tools/Compute_Oracle/am_sml.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Tools/Compute_Oracle/am_sml.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,16 +1,16 @@
 (*  Title:      Tools/Compute_Oracle/am_sml.ML
     Author:     Steven Obua
 
-    ToDO: "parameterless rewrite cannot be used in pattern": In a lot of cases it CAN be used, and these cases should be handled properly; 
-          right now, all cases throw an exception.
- 
+TODO: "parameterless rewrite cannot be used in pattern": In a lot of
+cases it CAN be used, and these cases should be handled
+properly; right now, all cases raise an exception. 
 *)
 
 signature AM_SML = 
 sig
   include ABSTRACT_MACHINE
   val save_result : (string * term) -> unit
-  val set_compiled_rewriter : (term -> term) -> unit				       
+  val set_compiled_rewriter : (term -> term) -> unit                                   
   val list_nth : 'a list * int -> 'a
   val dump_output : (string option) Unsynchronized.ref 
 end
@@ -42,7 +42,7 @@
 
 fun update_arity arity code a = 
     (case Inttab.lookup arity code of
-	 NONE => Inttab.update_new (code, a) arity
+         NONE => Inttab.update_new (code, a) arity
        | SOME (a': int) => if a > a' then Inttab.update (code, a) arity else arity)
 
 (* We have to find out the maximal arity of each constant *)
@@ -72,7 +72,7 @@
   | beta (App (Abs m, b)) = beta (unlift 0 (subst 0 m (lift 0 b)))
   | beta (App (a, b)) = 
     (case beta a of
-	 Abs m => beta (App (Abs m, b))
+         Abs m => beta (App (Abs m, b))
        | a => App (a, beta b))
   | beta (Abs m) = Abs (beta m)
   | beta (Computed t) = Computed t
@@ -102,48 +102,48 @@
 (* Remove all rules that are just parameterless rewrites. This is necessary because SML does not allow functions with no parameters. *)
 fun inline_rules rules =
     let
-	fun term_contains_const c (App (a, b)) = term_contains_const c a orelse term_contains_const c b
-	  | term_contains_const c (Abs m) = term_contains_const c m
-	  | term_contains_const c (Var i) = false
-	  | term_contains_const c (Const c') = (c = c')
-	fun find_rewrite [] = NONE
-	  | find_rewrite ((prems, PConst (c, []), r) :: _) = 
-	    if check_freevars 0 r then 
-		if term_contains_const c r then 
-		    raise Compile "parameterless rewrite is caught in cycle"
-		else if not (null prems) then
-		    raise Compile "parameterless rewrite may not be guarded"
-		else
-		    SOME (c, r) 
-	    else raise Compile "unbound variable on right hand side or guards of rule"
-	  | find_rewrite (_ :: rules) = find_rewrite rules
-	fun remove_rewrite (c,r) [] = []
-	  | remove_rewrite (cr as (c,r)) ((rule as (prems', PConst (c', args), r'))::rules) = 
-	    (if c = c' then 
-		 if null args andalso r = r' andalso null (prems') then 
-		     remove_rewrite cr rules 
-		 else raise Compile "incompatible parameterless rewrites found"
-	     else
-		 rule :: (remove_rewrite cr rules))
-	  | remove_rewrite cr (r::rs) = r::(remove_rewrite cr rs)
-	fun pattern_contains_const c (PConst (c', args)) = (c = c' orelse exists (pattern_contains_const c) args)
-	  | pattern_contains_const c (PVar) = false
-	fun inline_rewrite (ct as (c, _)) (prems, p, r) = 
-	    if pattern_contains_const c p then 
-		raise Compile "parameterless rewrite cannot be used in pattern"
-	    else (map (fn (Guard (a,b)) => Guard (subst_const ct a, subst_const ct b)) prems, p, subst_const ct r)
-	fun inline inlined rules =
-	    (case find_rewrite rules of 
-		 NONE => (Inttab.make inlined, rules)
-	       | SOME ct => 
-		 let
-		     val rules = map (inline_rewrite ct) (remove_rewrite ct rules)
-		     val inlined =  ct :: (map (fn (c', r) => (c', subst_const ct r)) inlined)
-		 in
-		     inline inlined rules
-		 end)		
+        fun term_contains_const c (App (a, b)) = term_contains_const c a orelse term_contains_const c b
+          | term_contains_const c (Abs m) = term_contains_const c m
+          | term_contains_const c (Var i) = false
+          | term_contains_const c (Const c') = (c = c')
+        fun find_rewrite [] = NONE
+          | find_rewrite ((prems, PConst (c, []), r) :: _) = 
+            if check_freevars 0 r then 
+                if term_contains_const c r then 
+                    raise Compile "parameterless rewrite is caught in cycle"
+                else if not (null prems) then
+                    raise Compile "parameterless rewrite may not be guarded"
+                else
+                    SOME (c, r) 
+            else raise Compile "unbound variable on right hand side or guards of rule"
+          | find_rewrite (_ :: rules) = find_rewrite rules
+        fun remove_rewrite (c,r) [] = []
+          | remove_rewrite (cr as (c,r)) ((rule as (prems', PConst (c', args), r'))::rules) = 
+            (if c = c' then 
+                 if null args andalso r = r' andalso null (prems') then 
+                     remove_rewrite cr rules 
+                 else raise Compile "incompatible parameterless rewrites found"
+             else
+                 rule :: (remove_rewrite cr rules))
+          | remove_rewrite cr (r::rs) = r::(remove_rewrite cr rs)
+        fun pattern_contains_const c (PConst (c', args)) = (c = c' orelse exists (pattern_contains_const c) args)
+          | pattern_contains_const c (PVar) = false
+        fun inline_rewrite (ct as (c, _)) (prems, p, r) = 
+            if pattern_contains_const c p then 
+                raise Compile "parameterless rewrite cannot be used in pattern"
+            else (map (fn (Guard (a,b)) => Guard (subst_const ct a, subst_const ct b)) prems, p, subst_const ct r)
+        fun inline inlined rules =
+            (case find_rewrite rules of 
+                 NONE => (Inttab.make inlined, rules)
+               | SOME ct => 
+                 let
+                     val rules = map (inline_rewrite ct) (remove_rewrite ct rules)
+                     val inlined =  ct :: (map (fn (c', r) => (c', subst_const ct r)) inlined)
+                 in
+                     inline inlined rules
+                 end)           
     in
-	inline [] rules		
+        inline [] rules         
     end
 
 
@@ -153,88 +153,88 @@
 *)
 fun adjust_rules rules = 
     let
-	val arity = fold (fn (prems, p, t) => fn arity => fold collect_guard_arity prems (collect_term_arity t (collect_pattern_arity p arity))) rules Inttab.empty
-	val toplevel_arity = fold (fn (_, p, t) => fn arity => collect_pattern_toplevel_arity p arity) rules Inttab.empty
-	fun arity_of c = the (Inttab.lookup arity c)
-	fun toplevel_arity_of c = the (Inttab.lookup toplevel_arity c)
-	fun test_pattern PVar = ()
-	  | test_pattern (C as PConst (c, args)) = if (length args <> arity_of c) then raise Compile ("Constant inside pattern must have maximal arity") else (map test_pattern args; ())
-	fun adjust_rule (_, PVar, _) = raise Compile ("pattern may not be a variable")
-	  | adjust_rule (_, PConst (_, []), _) = raise Compile ("cannot deal with rewrites that take no parameters")
-	  | adjust_rule (rule as (prems, p as PConst (c, args),t)) = 
-	    let
-		val patternvars_counted = count_patternvars p
-		fun check_fv t = check_freevars patternvars_counted t
-		val _ = if not (check_fv t) then raise Compile ("unbound variables on right hand side of rule") else () 
-		val _ = if not (forall (fn (Guard (a,b)) => check_fv a andalso check_fv b) prems) then raise Compile ("unbound variables in guards") else () 
-		val _ = map test_pattern args	        
-		val len = length args
-		val arity = arity_of c
-		val lift = nlift 0
-		fun addapps_tm n t = if n=0 then t else addapps_tm (n-1) (App (t, Var (n-1)))
-		fun adjust_term n t = addapps_tm n (lift n t)
-		fun adjust_guard n (Guard (a,b)) = Guard (lift n a, lift n b)
-	    in
-		if len = arity then
-		    rule
-		else if arity >= len then  
-		    (map (adjust_guard (arity-len)) prems, PConst (c, args @ (rep (arity-len) PVar)), adjust_term (arity-len) t)
-		else (raise Compile "internal error in adjust_rule")
-	    end
-	fun beta_rule (prems, p, t) = ((prems, p, beta t) handle Match => raise Compile "beta_rule")
+        val arity = fold (fn (prems, p, t) => fn arity => fold collect_guard_arity prems (collect_term_arity t (collect_pattern_arity p arity))) rules Inttab.empty
+        val toplevel_arity = fold (fn (_, p, t) => fn arity => collect_pattern_toplevel_arity p arity) rules Inttab.empty
+        fun arity_of c = the (Inttab.lookup arity c)
+        fun toplevel_arity_of c = the (Inttab.lookup toplevel_arity c)
+        fun test_pattern PVar = ()
+          | test_pattern (C as PConst (c, args)) = if (length args <> arity_of c) then raise Compile ("Constant inside pattern must have maximal arity") else (map test_pattern args; ())
+        fun adjust_rule (_, PVar, _) = raise Compile ("pattern may not be a variable")
+          | adjust_rule (_, PConst (_, []), _) = raise Compile ("cannot deal with rewrites that take no parameters")
+          | adjust_rule (rule as (prems, p as PConst (c, args),t)) = 
+            let
+                val patternvars_counted = count_patternvars p
+                fun check_fv t = check_freevars patternvars_counted t
+                val _ = if not (check_fv t) then raise Compile ("unbound variables on right hand side of rule") else () 
+                val _ = if not (forall (fn (Guard (a,b)) => check_fv a andalso check_fv b) prems) then raise Compile ("unbound variables in guards") else () 
+                val _ = map test_pattern args           
+                val len = length args
+                val arity = arity_of c
+                val lift = nlift 0
+                fun addapps_tm n t = if n=0 then t else addapps_tm (n-1) (App (t, Var (n-1)))
+                fun adjust_term n t = addapps_tm n (lift n t)
+                fun adjust_guard n (Guard (a,b)) = Guard (lift n a, lift n b)
+            in
+                if len = arity then
+                    rule
+                else if arity >= len then  
+                    (map (adjust_guard (arity-len)) prems, PConst (c, args @ (rep (arity-len) PVar)), adjust_term (arity-len) t)
+                else (raise Compile "internal error in adjust_rule")
+            end
+        fun beta_rule (prems, p, t) = ((prems, p, beta t) handle Match => raise Compile "beta_rule")
     in
-	(arity, toplevel_arity, map (beta_rule o adjust_rule) rules)
-    end		    
+        (arity, toplevel_arity, map (beta_rule o adjust_rule) rules)
+    end             
 
 fun print_term module arity_of toplevel_arity_of pattern_var_count pattern_lazy_var_count =
 let
     fun str x = string_of_int x
     fun protect_blank s = if exists_string Symbol.is_ascii_blank s then "(" ^ s ^")" else s
-    val module_prefix = (case module of NONE => "" | SOME s => s^".")											  
+    val module_prefix = (case module of NONE => "" | SOME s => s^".")                                                                                     
     fun print_apps d f [] = f
       | print_apps d f (a::args) = print_apps d (module_prefix^"app "^(protect_blank f)^" "^(protect_blank (print_term d a))) args
     and print_call d (App (a, b)) args = print_call d a (b::args) 
       | print_call d (Const c) args = 
-	(case arity_of c of 
-	     NONE => print_apps d (module_prefix^"Const "^(str c)) args 
-	   | SOME 0 => module_prefix^"C"^(str c)
-	   | SOME a =>
-	     let
-		 val len = length args
-	     in
-		 if a <= len then 
-		     let
-			 val strict_a = (case toplevel_arity_of c of SOME sa => sa | NONE => a)
-			 val _ = if strict_a > a then raise Compile "strict" else ()
-			 val s = module_prefix^"c"^(str c)^(implode (map (fn t => " "^(protect_blank (print_term d t))) (List.take (args, strict_a))))
-			 val s = s^(implode (map (fn t => " (fn () => "^print_term d t^")") (List.drop (List.take (args, a), strict_a))))
-		     in
-			 print_apps d s (List.drop (args, a))
-		     end
-		 else 
-		     let
-			 fun mk_apps n t = if n = 0 then t else mk_apps (n-1) (App (t, Var (n - 1)))
-			 fun mk_lambdas n t = if n = 0 then t else mk_lambdas (n-1) (Abs t)
-			 fun append_args [] t = t
-			   | append_args (c::cs) t = append_args cs (App (t, c))
-		     in
-			 print_term d (mk_lambdas (a-len) (mk_apps (a-len) (nlift 0 (a-len) (append_args args (Const c)))))
-		     end
-	     end)
+        (case arity_of c of 
+             NONE => print_apps d (module_prefix^"Const "^(str c)) args 
+           | SOME 0 => module_prefix^"C"^(str c)
+           | SOME a =>
+             let
+                 val len = length args
+             in
+                 if a <= len then 
+                     let
+                         val strict_a = (case toplevel_arity_of c of SOME sa => sa | NONE => a)
+                         val _ = if strict_a > a then raise Compile "strict" else ()
+                         val s = module_prefix^"c"^(str c)^(implode (map (fn t => " "^(protect_blank (print_term d t))) (List.take (args, strict_a))))
+                         val s = s^(implode (map (fn t => " (fn () => "^print_term d t^")") (List.drop (List.take (args, a), strict_a))))
+                     in
+                         print_apps d s (List.drop (args, a))
+                     end
+                 else 
+                     let
+                         fun mk_apps n t = if n = 0 then t else mk_apps (n-1) (App (t, Var (n - 1)))
+                         fun mk_lambdas n t = if n = 0 then t else mk_lambdas (n-1) (Abs t)
+                         fun append_args [] t = t
+                           | append_args (c::cs) t = append_args cs (App (t, c))
+                     in
+                         print_term d (mk_lambdas (a-len) (mk_apps (a-len) (nlift 0 (a-len) (append_args args (Const c)))))
+                     end
+             end)
       | print_call d t args = print_apps d (print_term d t) args
     and print_term d (Var x) = 
-	if x < d then 
-	    "b"^(str (d-x-1)) 
-	else 
-	    let
-		val n = pattern_var_count - (x-d) - 1
-		val x = "x"^(str n)
-	    in
-		if n < pattern_var_count - pattern_lazy_var_count then 
-		    x
-		else 
-		    "("^x^" ())"
-	    end								
+        if x < d then 
+            "b"^(str (d-x-1)) 
+        else 
+            let
+                val n = pattern_var_count - (x-d) - 1
+                val x = "x"^(str n)
+            in
+                if n < pattern_var_count - pattern_lazy_var_count then 
+                    x
+                else 
+                    "("^x^" ())"
+            end                                                         
       | print_term d (Abs c) = module_prefix^"Abs (fn b"^(str d)^" => "^(print_term (d + 1) c)^")"
       | print_term d t = print_call d t []
 in
@@ -242,215 +242,215 @@
 end
 
 fun section n = if n = 0 then [] else (section (n-1))@[n-1]
-			 			
+                                                
 fun print_rule gnum arity_of toplevel_arity_of (guards, p, t) = 
-    let	
-	fun str x = Int.toString x		    
-	fun print_pattern top n PVar = (n+1, "x"^(str n))
-	  | print_pattern top n (PConst (c, [])) = (n, (if top then "c" else "C")^(str c)^(if top andalso gnum > 0 then "_"^(str gnum) else ""))
-	  | print_pattern top n (PConst (c, args)) = 
-	    let
-		val f = (if top then "c" else "C")^(str c)^(if top andalso gnum > 0 then "_"^(str gnum) else "")
-		val (n, s) = print_pattern_list 0 top (n, f) args
-	    in
-		(n, s)
-	    end
-	and print_pattern_list' counter top (n,p) [] = if top then (n,p) else (n,p^")")
-	  | print_pattern_list' counter top (n, p) (t::ts) = 
-	    let
-		val (n, t) = print_pattern false n t
-	    in
-		print_pattern_list' (counter + 1) top (n, if top then p^" (a"^(str counter)^" as ("^t^"))" else p^", "^t) ts
-	    end	
-	and print_pattern_list counter top (n, p) (t::ts) = 
-	    let
-		val (n, t) = print_pattern false n t
-	    in
-		print_pattern_list' (counter + 1) top (n, if top then p^" (a"^(str counter)^" as ("^t^"))" else p^" ("^t) ts
-	    end
-	val c = (case p of PConst (c, _) => c | _ => raise Match)
-	val (n, pattern) = print_pattern true 0 p
-	val lazy_vars = the (arity_of c) - the (toplevel_arity_of c)
-	fun print_tm tm = print_term NONE arity_of toplevel_arity_of n lazy_vars tm
+    let 
+        fun str x = Int.toString x                  
+        fun print_pattern top n PVar = (n+1, "x"^(str n))
+          | print_pattern top n (PConst (c, [])) = (n, (if top then "c" else "C")^(str c)^(if top andalso gnum > 0 then "_"^(str gnum) else ""))
+          | print_pattern top n (PConst (c, args)) = 
+            let
+                val f = (if top then "c" else "C")^(str c)^(if top andalso gnum > 0 then "_"^(str gnum) else "")
+                val (n, s) = print_pattern_list 0 top (n, f) args
+            in
+                (n, s)
+            end
+        and print_pattern_list' counter top (n,p) [] = if top then (n,p) else (n,p^")")
+          | print_pattern_list' counter top (n, p) (t::ts) = 
+            let
+                val (n, t) = print_pattern false n t
+            in
+                print_pattern_list' (counter + 1) top (n, if top then p^" (a"^(str counter)^" as ("^t^"))" else p^", "^t) ts
+            end 
+        and print_pattern_list counter top (n, p) (t::ts) = 
+            let
+                val (n, t) = print_pattern false n t
+            in
+                print_pattern_list' (counter + 1) top (n, if top then p^" (a"^(str counter)^" as ("^t^"))" else p^" ("^t) ts
+            end
+        val c = (case p of PConst (c, _) => c | _ => raise Match)
+        val (n, pattern) = print_pattern true 0 p
+        val lazy_vars = the (arity_of c) - the (toplevel_arity_of c)
+        fun print_tm tm = print_term NONE arity_of toplevel_arity_of n lazy_vars tm
         fun print_guard (Guard (a,b)) = "term_eq ("^(print_tm a)^") ("^(print_tm b)^")"
-	val else_branch = "c"^(str c)^"_"^(str (gnum+1))^(implode (map (fn i => " a"^(str i)) (section (the (arity_of c)))))
-	fun print_guards t [] = print_tm t
-	  | print_guards t (g::gs) = "if ("^(print_guard g)^")"^(implode (map (fn g => " andalso ("^(print_guard g)^")") gs))^" then ("^(print_tm t)^") else "^else_branch
+        val else_branch = "c"^(str c)^"_"^(str (gnum+1))^(implode (map (fn i => " a"^(str i)) (section (the (arity_of c)))))
+        fun print_guards t [] = print_tm t
+          | print_guards t (g::gs) = "if ("^(print_guard g)^")"^(implode (map (fn g => " andalso ("^(print_guard g)^")") gs))^" then ("^(print_tm t)^") else "^else_branch
     in
-	(if null guards then gnum else gnum+1, pattern^" = "^(print_guards t guards))
+        (if null guards then gnum else gnum+1, pattern^" = "^(print_guards t guards))
     end
 
 fun group_rules rules =
     let
-	fun add_rule (r as (_, PConst (c,_), _)) groups =
-	    let
-		val rs = (case Inttab.lookup groups c of NONE => [] | SOME rs => rs)
-	    in
-		Inttab.update (c, r::rs) groups
-	    end
-	  | add_rule _ _ = raise Compile "internal error group_rules"
+        fun add_rule (r as (_, PConst (c,_), _)) groups =
+            let
+                val rs = (case Inttab.lookup groups c of NONE => [] | SOME rs => rs)
+            in
+                Inttab.update (c, r::rs) groups
+            end
+          | add_rule _ _ = raise Compile "internal error group_rules"
     in
-	fold_rev add_rule rules Inttab.empty
+        fold_rev add_rule rules Inttab.empty
     end
 
 fun sml_prog name code rules = 
     let
-	val buffer = Unsynchronized.ref ""
-	fun write s = (buffer := (!buffer)^s)
-	fun writeln s = (write s; write "\n")
-	fun writelist [] = ()
-	  | writelist (s::ss) = (writeln s; writelist ss)
-	fun str i = Int.toString i
-	val (inlinetab, rules) = inline_rules rules
-	val (arity, toplevel_arity, rules) = adjust_rules rules
-	val rules = group_rules rules
-	val constants = Inttab.keys arity
-	fun arity_of c = Inttab.lookup arity c
-	fun toplevel_arity_of c = Inttab.lookup toplevel_arity c
-	fun rep_str s n = implode (rep n s)
-	fun indexed s n = s^(str n)
+        val buffer = Unsynchronized.ref ""
+        fun write s = (buffer := (!buffer)^s)
+        fun writeln s = (write s; write "\n")
+        fun writelist [] = ()
+          | writelist (s::ss) = (writeln s; writelist ss)
+        fun str i = Int.toString i
+        val (inlinetab, rules) = inline_rules rules
+        val (arity, toplevel_arity, rules) = adjust_rules rules
+        val rules = group_rules rules
+        val constants = Inttab.keys arity
+        fun arity_of c = Inttab.lookup arity c
+        fun toplevel_arity_of c = Inttab.lookup toplevel_arity c
+        fun rep_str s n = implode (rep n s)
+        fun indexed s n = s^(str n)
         fun string_of_tuple [] = ""
-	  | string_of_tuple (x::xs) = "("^x^(implode (map (fn s => ", "^s) xs))^")"
+          | string_of_tuple (x::xs) = "("^x^(implode (map (fn s => ", "^s) xs))^")"
         fun string_of_args [] = ""
-	  | string_of_args (x::xs) = x^(implode (map (fn s => " "^s) xs))
-	fun default_case gnum c = 
-	    let
-		val leftargs = implode (map (indexed " x") (section (the (arity_of c))))
-		val rightargs = section (the (arity_of c))
-		val strict_args = (case toplevel_arity_of c of NONE => the (arity_of c) | SOME sa => sa)
-		val xs = map (fn n => if n < strict_args then "x"^(str n) else "x"^(str n)^"()") rightargs
-		val right = (indexed "C" c)^" "^(string_of_tuple xs)
-		val message = "(\"unresolved lazy call: " ^ string_of_int c ^ "\")"
-		val right = if strict_args < the (arity_of c) then "raise AM_SML.Run "^message else right		
-	    in
-		(indexed "c" c)^(if gnum > 0 then "_"^(str gnum) else "")^leftargs^" = "^right
-	    end
+          | string_of_args (x::xs) = x^(implode (map (fn s => " "^s) xs))
+        fun default_case gnum c = 
+            let
+                val leftargs = implode (map (indexed " x") (section (the (arity_of c))))
+                val rightargs = section (the (arity_of c))
+                val strict_args = (case toplevel_arity_of c of NONE => the (arity_of c) | SOME sa => sa)
+                val xs = map (fn n => if n < strict_args then "x"^(str n) else "x"^(str n)^"()") rightargs
+                val right = (indexed "C" c)^" "^(string_of_tuple xs)
+                val message = "(\"unresolved lazy call: " ^ string_of_int c ^ "\")"
+                val right = if strict_args < the (arity_of c) then "raise AM_SML.Run "^message else right               
+            in
+                (indexed "c" c)^(if gnum > 0 then "_"^(str gnum) else "")^leftargs^" = "^right
+            end
 
-	fun eval_rules c = 
-	    let
-		val arity = the (arity_of c)
-		val strict_arity = (case toplevel_arity_of c of NONE => arity | SOME sa => sa)
-		fun eval_rule n = 
-		    let
-			val sc = string_of_int c
-			val left = fold (fn i => fn s => "AbstractMachine.App ("^s^(indexed ", x" i)^")") (section n) ("AbstractMachine.Const "^sc)
+        fun eval_rules c = 
+            let
+                val arity = the (arity_of c)
+                val strict_arity = (case toplevel_arity_of c of NONE => arity | SOME sa => sa)
+                fun eval_rule n = 
+                    let
+                        val sc = string_of_int c
+                        val left = fold (fn i => fn s => "AbstractMachine.App ("^s^(indexed ", x" i)^")") (section n) ("AbstractMachine.Const "^sc)
                         fun arg i = 
-			    let
-				val x = indexed "x" i
-				val x = if i < n then "(eval bounds "^x^")" else x
-				val x = if i < strict_arity then x else "(fn () => "^x^")"
-			    in
-				x
-			    end
-			val right = "c"^sc^" "^(string_of_args (map arg (section arity)))
-			val right = fold_rev (fn i => fn s => "Abs (fn "^(indexed "x" i)^" => "^s^")") (List.drop (section arity, n)) right		
-			val right = if arity > 0 then right else "C"^sc
-		    in
-			"  | eval bounds ("^left^") = "^right
-		    end
-	    in
-		map eval_rule (rev (section (arity + 1)))
-	    end
+                            let
+                                val x = indexed "x" i
+                                val x = if i < n then "(eval bounds "^x^")" else x
+                                val x = if i < strict_arity then x else "(fn () => "^x^")"
+                            in
+                                x
+                            end
+                        val right = "c"^sc^" "^(string_of_args (map arg (section arity)))
+                        val right = fold_rev (fn i => fn s => "Abs (fn "^(indexed "x" i)^" => "^s^")") (List.drop (section arity, n)) right             
+                        val right = if arity > 0 then right else "C"^sc
+                    in
+                        "  | eval bounds ("^left^") = "^right
+                    end
+            in
+                map eval_rule (rev (section (arity + 1)))
+            end
 
-	fun convert_computed_rules (c: int) : string list = 
-	    let
-		val arity = the (arity_of c)
-		fun eval_rule () = 
-		    let
-			val sc = string_of_int c
-			val left = fold (fn i => fn s => "AbstractMachine.App ("^s^(indexed ", x" i)^")") (section arity) ("AbstractMachine.Const "^sc)
+        fun convert_computed_rules (c: int) : string list = 
+            let
+                val arity = the (arity_of c)
+                fun eval_rule () = 
+                    let
+                        val sc = string_of_int c
+                        val left = fold (fn i => fn s => "AbstractMachine.App ("^s^(indexed ", x" i)^")") (section arity) ("AbstractMachine.Const "^sc)
                         fun arg i = "(convert_computed "^(indexed "x" i)^")" 
-			val right = "C"^sc^" "^(string_of_tuple (map arg (section arity)))		
-			val right = if arity > 0 then right else "C"^sc
-		    in
-			"  | convert_computed ("^left^") = "^right
-		    end
-	    in
-		[eval_rule ()]
-	    end
+                        val right = "C"^sc^" "^(string_of_tuple (map arg (section arity)))              
+                        val right = if arity > 0 then right else "C"^sc
+                    in
+                        "  | convert_computed ("^left^") = "^right
+                    end
+            in
+                [eval_rule ()]
+            end
         
-	fun mk_constr_type_args n = if n > 0 then " of Term "^(rep_str " * Term" (n-1)) else ""
-	val _ = writelist [                   
-		"structure "^name^" = struct",
-		"",
-		"datatype Term = Const of int | App of Term * Term | Abs of (Term -> Term)",
-		"         "^(implode (map (fn c => " | C"^(str c)^(mk_constr_type_args (the (arity_of c)))) constants)),
-		""]
-	fun make_constr c argprefix = "(C"^(str c)^" "^(string_of_tuple (map (fn i => argprefix^(str i)) (section (the (arity_of c)))))^")"
-	fun make_term_eq c = "  | term_eq "^(make_constr c "a")^" "^(make_constr c "b")^" = "^
+        fun mk_constr_type_args n = if n > 0 then " of Term "^(rep_str " * Term" (n-1)) else ""
+        val _ = writelist [                   
+                "structure "^name^" = struct",
+                "",
+                "datatype Term = Const of int | App of Term * Term | Abs of (Term -> Term)",
+                "         "^(implode (map (fn c => " | C"^(str c)^(mk_constr_type_args (the (arity_of c)))) constants)),
+                ""]
+        fun make_constr c argprefix = "(C"^(str c)^" "^(string_of_tuple (map (fn i => argprefix^(str i)) (section (the (arity_of c)))))^")"
+        fun make_term_eq c = "  | term_eq "^(make_constr c "a")^" "^(make_constr c "b")^" = "^
                              (case the (arity_of c) of 
-				  0 => "true"
-				| n => 
-				  let 
-				      val eqs = map (fn i => "term_eq a"^(str i)^" b"^(str i)) (section n)
-				      val (eq, eqs) = (List.hd eqs, map (fn s => " andalso "^s) (List.tl eqs))
-				  in
-				      eq^(implode eqs)
-				  end)
-	val _ = writelist [
-		"fun term_eq (Const c1) (Const c2) = (c1 = c2)",
-		"  | term_eq (App (a1,a2)) (App (b1,b2)) = term_eq a1 b1 andalso term_eq a2 b2"]
-	val _ = writelist (map make_term_eq constants)		
-	val _ = writelist [
-		"  | term_eq _ _ = false",
+                                  0 => "true"
+                                | n => 
+                                  let 
+                                      val eqs = map (fn i => "term_eq a"^(str i)^" b"^(str i)) (section n)
+                                      val (eq, eqs) = (List.hd eqs, map (fn s => " andalso "^s) (List.tl eqs))
+                                  in
+                                      eq^(implode eqs)
+                                  end)
+        val _ = writelist [
+                "fun term_eq (Const c1) (Const c2) = (c1 = c2)",
+                "  | term_eq (App (a1,a2)) (App (b1,b2)) = term_eq a1 b1 andalso term_eq a2 b2"]
+        val _ = writelist (map make_term_eq constants)          
+        val _ = writelist [
+                "  | term_eq _ _ = false",
                 "" 
-		] 
-	val _ = writelist [
-		"fun app (Abs a) b = a b",
-		"  | app a b = App (a, b)",
-		""]	
-	fun defcase gnum c = (case arity_of c of NONE => [] | SOME a => if a > 0 then [default_case gnum c] else [])
-	fun writefundecl [] = () 
-	  | writefundecl (x::xs) = writelist ((("and "^x)::(map (fn s => "  | "^s) xs)))
-	fun list_group c = (case Inttab.lookup rules c of 
-				NONE => [defcase 0 c]
-			      | SOME rs => 
-				let
-				    val rs = 
-					fold
-					    (fn r => 
-					     fn rs =>
-						let 
-						    val (gnum, l, rs) = 
-							(case rs of 
-							     [] => (0, [], []) 
-							   | (gnum, l)::rs => (gnum, l, rs))
-						    val (gnum', r) = print_rule gnum arity_of toplevel_arity_of r 
-						in 
-						    if gnum' = gnum then 
-							(gnum, r::l)::rs
-						    else
-							let
-							    val args = implode (map (fn i => " a"^(str i)) (section (the (arity_of c))))
-							    fun gnumc g = if g > 0 then "c"^(str c)^"_"^(str g)^args else "c"^(str c)^args
-							    val s = gnumc (gnum) ^ " = " ^ gnumc (gnum') 
-							in
-							    (gnum', [])::(gnum, s::r::l)::rs
-							end
-						end)
-					rs []
-				    val rs = (case rs of [] => [(0,defcase 0 c)] | (gnum,l)::rs => (gnum, (defcase gnum c)@l)::rs)
-				in
-				    rev (map (fn z => rev (snd z)) rs)
-				end)
-	val _ = map (fn z => (map writefundecl z; writeln "")) (map list_group constants)
-	val _ = writelist [
-		"fun convert (Const i) = AM_SML.Const i",
-		"  | convert (App (a, b)) = AM_SML.App (convert a, convert b)",
-                "  | convert (Abs _) = raise AM_SML.Run \"no abstraction in result allowed\""]	
-	fun make_convert c = 
-	    let
-		val args = map (indexed "a") (section (the (arity_of c)))
-		val leftargs = 
-		    case args of
-			[] => ""
-		      | (x::xs) => "("^x^(implode (map (fn s => ", "^s) xs))^")"
-		val args = map (indexed "convert a") (section (the (arity_of c)))
-		val right = fold (fn x => fn s => "AM_SML.App ("^s^", "^x^")") args ("AM_SML.Const "^(str c))
-	    in
-		"  | convert (C"^(str c)^" "^leftargs^") = "^right
-	    end 		
-	val _ = writelist (map make_convert constants)
+                ] 
+        val _ = writelist [
+                "fun app (Abs a) b = a b",
+                "  | app a b = App (a, b)",
+                ""]     
+        fun defcase gnum c = (case arity_of c of NONE => [] | SOME a => if a > 0 then [default_case gnum c] else [])
+        fun writefundecl [] = () 
+          | writefundecl (x::xs) = writelist ((("and "^x)::(map (fn s => "  | "^s) xs)))
+        fun list_group c = (case Inttab.lookup rules c of 
+                                NONE => [defcase 0 c]
+                              | SOME rs => 
+                                let
+                                    val rs = 
+                                        fold
+                                            (fn r => 
+                                             fn rs =>
+                                                let 
+                                                    val (gnum, l, rs) = 
+                                                        (case rs of 
+                                                             [] => (0, [], []) 
+                                                           | (gnum, l)::rs => (gnum, l, rs))
+                                                    val (gnum', r) = print_rule gnum arity_of toplevel_arity_of r 
+                                                in 
+                                                    if gnum' = gnum then 
+                                                        (gnum, r::l)::rs
+                                                    else
+                                                        let
+                                                            val args = implode (map (fn i => " a"^(str i)) (section (the (arity_of c))))
+                                                            fun gnumc g = if g > 0 then "c"^(str c)^"_"^(str g)^args else "c"^(str c)^args
+                                                            val s = gnumc (gnum) ^ " = " ^ gnumc (gnum') 
+                                                        in
+                                                            (gnum', [])::(gnum, s::r::l)::rs
+                                                        end
+                                                end)
+                                        rs []
+                                    val rs = (case rs of [] => [(0,defcase 0 c)] | (gnum,l)::rs => (gnum, (defcase gnum c)@l)::rs)
+                                in
+                                    rev (map (fn z => rev (snd z)) rs)
+                                end)
+        val _ = map (fn z => (map writefundecl z; writeln "")) (map list_group constants)
+        val _ = writelist [
+                "fun convert (Const i) = AM_SML.Const i",
+                "  | convert (App (a, b)) = AM_SML.App (convert a, convert b)",
+                "  | convert (Abs _) = raise AM_SML.Run \"no abstraction in result allowed\""]  
+        fun make_convert c = 
+            let
+                val args = map (indexed "a") (section (the (arity_of c)))
+                val leftargs = 
+                    case args of
+                        [] => ""
+                      | (x::xs) => "("^x^(implode (map (fn s => ", "^s) xs))^")"
+                val args = map (indexed "convert a") (section (the (arity_of c)))
+                val right = fold (fn x => fn s => "AM_SML.App ("^s^", "^x^")") args ("AM_SML.Const "^(str c))
+            in
+                "  | convert (C"^(str c)^" "^leftargs^") = "^right
+            end                 
+        val _ = writelist (map make_convert constants)
         val _ = writelist [
                 "",
                 "fun convert_computed (AbstractMachine.Abs b) = raise AM_SML.Run \"no abstraction in convert_computed allowed\"",
@@ -460,33 +460,33 @@
                 "  | convert_computed (AbstractMachine.Const c) = Const c",
                 "  | convert_computed (AbstractMachine.App (a, b)) = App (convert_computed a, convert_computed b)",
                 "  | convert_computed (AbstractMachine.Computed a) = raise AM_SML.Run \"no nesting in convert_computed allowed\""] 
-	val _ = writelist [
-		"",
-		"fun eval bounds (AbstractMachine.Abs m) = Abs (fn b => eval (b::bounds) m)",
-		"  | eval bounds (AbstractMachine.Var i) = AM_SML.list_nth (bounds, i)"]
-	val _ = map (writelist o eval_rules) constants
-	val _ = writelist [
+        val _ = writelist [
+                "",
+                "fun eval bounds (AbstractMachine.Abs m) = Abs (fn b => eval (b::bounds) m)",
+                "  | eval bounds (AbstractMachine.Var i) = AM_SML.list_nth (bounds, i)"]
+        val _ = map (writelist o eval_rules) constants
+        val _ = writelist [
                 "  | eval bounds (AbstractMachine.App (a, b)) = app (eval bounds a) (eval bounds b)",
                 "  | eval bounds (AbstractMachine.Const c) = Const c",
                 "  | eval bounds (AbstractMachine.Computed t) = convert_computed t"]                
-	val _ = writelist [		
-		"",
-		"fun export term = AM_SML.save_result (\""^code^"\", convert term)",
-		"",
+        val _ = writelist [             
+                "",
+                "fun export term = AM_SML.save_result (\""^code^"\", convert term)",
+                "",
                 "val _ = AM_SML.set_compiled_rewriter (fn t => (convert (eval [] t)))",
                 "",
-		"end"]
+                "end"]
     in
-	(arity, toplevel_arity, inlinetab, !buffer)
+        (arity, toplevel_arity, inlinetab, !buffer)
     end
 
 val guid_counter = Unsynchronized.ref 0
 fun get_guid () = 
     let
-	val c = !guid_counter
-	val _ = guid_counter := !guid_counter + 1
+        val c = !guid_counter
+        val _ = guid_counter := !guid_counter + 1
     in
-	(LargeInt.toString (Time.toMicroseconds (Time.now ()))) ^ (string_of_int c)
+        (LargeInt.toString (Time.toMicroseconds (Time.now ()))) ^ (string_of_int c)
     end
 
 
@@ -496,53 +496,53 @@
     
 fun compile cache_patterns const_arity eqs = 
     let
-	val guid = get_guid ()
-	val code = Real.toString (random ())
-	val module = "AMSML_"^guid
-	val (arity, toplevel_arity, inlinetab, source) = sml_prog module code eqs
-	val _ = case !dump_output of NONE => () | SOME p => writeTextFile p source
-	val _ = compiled_rewriter := NONE
-	val _ = use_source source
+        val guid = get_guid ()
+        val code = Real.toString (random ())
+        val module = "AMSML_"^guid
+        val (arity, toplevel_arity, inlinetab, source) = sml_prog module code eqs
+        val _ = case !dump_output of NONE => () | SOME p => writeTextFile p source
+        val _ = compiled_rewriter := NONE
+        val _ = use_source source
     in
-	case !compiled_rewriter of 
-	    NONE => raise Compile "broken link to compiled function"
-	  | SOME f => (module, code, arity, toplevel_arity, inlinetab, f)
+        case !compiled_rewriter of 
+            NONE => raise Compile "broken link to compiled function"
+          | SOME f => (module, code, arity, toplevel_arity, inlinetab, f)
     end
 
 
 fun run' (module, code, arity, toplevel_arity, inlinetab, compiled_fun) t = 
-    let	
-	val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms")
-	fun inline (Const c) = (case Inttab.lookup inlinetab c of NONE => Const c | SOME t => t)
-	  | inline (Var i) = Var i
-	  | inline (App (a, b)) = App (inline a, inline b)
-	  | inline (Abs m) = Abs (inline m)
-	val t = beta (inline t)
-	fun arity_of c = Inttab.lookup arity c		 	 
-	fun toplevel_arity_of c = Inttab.lookup toplevel_arity c
-	val term = print_term NONE arity_of toplevel_arity_of 0 0 t 
+    let 
+        val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms")
+        fun inline (Const c) = (case Inttab.lookup inlinetab c of NONE => Const c | SOME t => t)
+          | inline (Var i) = Var i
+          | inline (App (a, b)) = App (inline a, inline b)
+          | inline (Abs m) = Abs (inline m)
+        val t = beta (inline t)
+        fun arity_of c = Inttab.lookup arity c                   
+        fun toplevel_arity_of c = Inttab.lookup toplevel_arity c
+        val term = print_term NONE arity_of toplevel_arity_of 0 0 t 
         val source = "local open "^module^" in val _ = export ("^term^") end"
-	val _ = writeTextFile "Gencode_call.ML" source
-	val _ = clear_result ()
-	val _ = use_source source
+        val _ = writeTextFile "Gencode_call.ML" source
+        val _ = clear_result ()
+        val _ = use_source source
     in
-	case !saved_result of 
-	    NONE => raise Run "broken link to compiled code"
-	  | SOME (code', t) => (clear_result (); if code' = code then t else raise Run "link to compiled code was hijacked")
-    end		
+        case !saved_result of 
+            NONE => raise Run "broken link to compiled code"
+          | SOME (code', t) => (clear_result (); if code' = code then t else raise Run "link to compiled code was hijacked")
+    end         
 
 fun run (module, code, arity, toplevel_arity, inlinetab, compiled_fun) t = 
-    let	
-	val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms")
-	fun inline (Const c) = (case Inttab.lookup inlinetab c of NONE => Const c | SOME t => t)
-	  | inline (Var i) = Var i
-	  | inline (App (a, b)) = App (inline a, inline b)
-	  | inline (Abs m) = Abs (inline m)
-	  | inline (Computed t) = Computed t
+    let 
+        val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms")
+        fun inline (Const c) = (case Inttab.lookup inlinetab c of NONE => Const c | SOME t => t)
+          | inline (Var i) = Var i
+          | inline (App (a, b)) = App (inline a, inline b)
+          | inline (Abs m) = Abs (inline m)
+          | inline (Computed t) = Computed t
     in
-	compiled_fun (beta (inline t))
-    end	
+        compiled_fun (beta (inline t))
+    end 
 
 fun discard p = ()
-			 	  
+                                  
 end
--- a/src/Tools/Compute_Oracle/linker.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Tools/Compute_Oracle/linker.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,12 +1,12 @@
 (*  Title:      Tools/Compute_Oracle/linker.ML
     Author:     Steven Obua
 
-Linker.ML solves the problem that the computing oracle does not
-instantiate polymorphic rules. By going through the PCompute interface,
-all possible instantiations are resolved by compiling new programs, if
-necessary. The obvious disadvantage of this approach is that in the
-worst case for each new term to be rewritten, a new program may be
-compiled.
+This module solves the problem that the computing oracle does not
+instantiate polymorphic rules. By going through the PCompute
+interface, all possible instantiations are resolved by compiling new
+programs, if necessary. The obvious disadvantage of this approach is
+that in the worst case for each new term to be rewritten, a new
+program may be compiled.
 *)
 
 (*
@@ -292,33 +292,33 @@
     val polycs = filter Linker.is_polymorphic cs
 in
     if null (polycs) then
-	MonoPattern pat
+        MonoPattern pat
     else
-	PolyPattern (pat, Linker.empty polycs, [])
+        PolyPattern (pat, Linker.empty polycs, [])
 end
-	     
+             
 fun create_computer machine thy pats ths =
     let
         fun add (MonoThm th) ths = th::ths
           | add (PolyThm (_, _, ths')) ths = ths'@ths
-	fun addpat (MonoPattern p) pats = p::pats
-	  | addpat (PolyPattern (_, _, ps)) pats = ps@pats
+        fun addpat (MonoPattern p) pats = p::pats
+          | addpat (PolyPattern (_, _, ps)) pats = ps@pats
         val ths = fold_rev add ths []
-	val pats = fold_rev addpat pats []
+        val pats = fold_rev addpat pats []
     in
         Compute.make_with_cache machine thy pats ths
     end
 
 fun update_computer computer pats ths = 
     let
-	fun add (MonoThm th) ths = th::ths
-	  | add (PolyThm (_, _, ths')) ths = ths'@ths
-	fun addpat (MonoPattern p) pats = p::pats
-	  | addpat (PolyPattern (_, _, ps)) pats = ps@pats
-	val ths = fold_rev add ths []
-	val pats = fold_rev addpat pats []
+        fun add (MonoThm th) ths = th::ths
+          | add (PolyThm (_, _, ths')) ths = ths'@ths
+        fun addpat (MonoPattern p) pats = p::pats
+          | addpat (PolyPattern (_, _, ps)) pats = ps@pats
+        val ths = fold_rev add ths []
+        val pats = fold_rev addpat pats []
     in
-	Compute.update_with_cache computer pats ths
+        Compute.update_with_cache computer pats ths
     end
 
 fun conv_subst thy (subst : Type.tyenv) =
@@ -338,29 +338,29 @@
             in
                 (newmonos, PolyThm (th, instances, instanceths@newths))
             end
-	fun addpats monocs (pat as (MonoPattern _)) = pat
-	  | addpats monocs (PolyPattern (p, instances, instancepats)) =
-	    let
-		val (newsubsts, instances) = Linker.add_instances thy instances monocs
-		val _ = if not (null newsubsts) then changed := true else ()
-		val newpats = map (fn subst => Envir.subst_term_types subst p) newsubsts
-	    in
-		PolyPattern (p, instances, instancepats@newpats)
-	    end 
+        fun addpats monocs (pat as (MonoPattern _)) = pat
+          | addpats monocs (PolyPattern (p, instances, instancepats)) =
+            let
+                val (newsubsts, instances) = Linker.add_instances thy instances monocs
+                val _ = if not (null newsubsts) then changed := true else ()
+                val newpats = map (fn subst => Envir.subst_term_types subst p) newsubsts
+            in
+                PolyPattern (p, instances, instancepats@newpats)
+            end 
         fun step monocs ths =
             fold_rev (fn th =>
                       fn (newmonos, ths) =>
                          let 
-			     val (newmonos', th') = add monocs th 
-			 in
+                             val (newmonos', th') = add monocs th 
+                         in
                              (newmonos'@newmonos, th'::ths)
                          end)
                      ths ([], [])
         fun loop monocs pats ths =
             let 
-		val (monocs', ths') = step monocs ths 
-		val pats' = map (addpats monocs) pats
-	    in
+                val (monocs', ths') = step monocs ths 
+                val pats' = map (addpats monocs) pats
+            in
                 if null (monocs') then
                     (pats', ths')
                 else
@@ -406,18 +406,18 @@
 
 fun make_with_cache machine thy pats ths cs =
     let
-	val ths = remove_duplicates ths
-	val (monocs, ths) = fold_rev (fn th => 
-				      fn (monocs, ths) => 
-					 let val (m, t) = create_theorem th in 
-					     (m@monocs, t::ths)
-					 end)
-				     ths (cs, [])
-	val pats = map create_pattern pats
-	val (_, (pats, ths)) = add_monos thy monocs pats ths
-	val computer = create_computer machine thy pats ths
+        val ths = remove_duplicates ths
+        val (monocs, ths) = fold_rev (fn th => 
+                                      fn (monocs, ths) => 
+                                         let val (m, t) = create_theorem th in 
+                                             (m@monocs, t::ths)
+                                         end)
+                                     ths (cs, [])
+        val pats = map create_pattern pats
+        val (_, (pats, ths)) = add_monos thy monocs pats ths
+        val computer = create_computer machine thy pats ths
     in
-	PComputer (Theory.check_thy thy, computer, Unsynchronized.ref ths, Unsynchronized.ref pats)
+        PComputer (Theory.check_thy thy, computer, Unsynchronized.ref ths, Unsynchronized.ref pats)
     end
 
 fun make machine thy ths cs = make_with_cache machine thy [] ths cs
@@ -427,13 +427,13 @@
         val thy = Theory.deref thyref
         val (changed, (pats, ths)) = add_monos thy cs (!rpats) (!rths)
     in
-	if changed then
-	    (update_computer computer pats ths;
-	     rths := ths;
-	     rpats := pats;
-	     true)
-	else
-	    false
+        if changed then
+            (update_computer computer pats ths;
+             rths := ths;
+             rpats := pats;
+             true)
+        else
+            false
 
     end
 
@@ -441,37 +441,37 @@
 
 fun rewrite pc cts =
     let
-	val _ = add_instances' pc (map term_of cts)
-	val computer = (computer_of pc)
+        val _ = add_instances' pc (map term_of cts)
+        val computer = (computer_of pc)
     in
-	map (fn ct => Compute.rewrite computer ct) cts
+        map (fn ct => Compute.rewrite computer ct) cts
     end
 
 fun simplify pc th = Compute.simplify (computer_of pc) th
 
 fun make_theorem pc th vars = 
     let
-	val _ = add_instances' pc [prop_of th]
+        val _ = add_instances' pc [prop_of th]
 
     in
-	Compute.make_theorem (computer_of pc) th vars
+        Compute.make_theorem (computer_of pc) th vars
     end
 
 fun instantiate pc insts th = 
     let
-	val _ = add_instances' pc (map (term_of o snd) insts)
+        val _ = add_instances' pc (map (term_of o snd) insts)
     in
-	Compute.instantiate (computer_of pc) insts th
+        Compute.instantiate (computer_of pc) insts th
     end
 
 fun evaluate_prem pc prem_no th = Compute.evaluate_prem (computer_of pc) prem_no th
 
 fun modus_ponens pc prem_no th' th =
     let
-	val _ = add_instances' pc [prop_of th']
+        val _ = add_instances' pc [prop_of th']
     in
-	Compute.modus_ponens (computer_of pc) prem_no th' th
+        Compute.modus_ponens (computer_of pc) prem_no th' th
     end    
-		 							      			    
+                                                                                                    
 
 end
--- a/src/Tools/Compute_Oracle/report.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Tools/Compute_Oracle/report.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -11,12 +11,12 @@
 
 fun timeit f =
     let
-	val t1 = start_timing ()
-	val x = f ()
-	val t2 = #message (end_timing t1)
-	val _ = writeln ((report_space ()) ^ "--> "^t2)
+        val t1 = start_timing ()
+        val x = f ()
+        val t2 = #message (end_timing t1)
+        val _ = writeln ((report_space ()) ^ "--> "^t2)
     in
-	x	
+        x       
     end
 
 fun report s f = 
--- a/src/Tools/eqsubst.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/Tools/eqsubst.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -235,13 +235,13 @@
           val initenv =
             Envir.Envir {maxidx = ix2, tenv = Vartab.empty, tyenv = typinsttab};
           val useq = Unify.smash_unifiers thry [a] initenv
-	            handle UnequalLengths => Seq.empty
-		               | Term.TERM _ => Seq.empty;
+              handle UnequalLengths => Seq.empty
+                   | Term.TERM _ => Seq.empty;
           fun clean_unify' useq () =
               (case (Seq.pull useq) of
                  NONE => NONE
                | SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t)))
-	            handle UnequalLengths => NONE
+              handle UnequalLengths => NONE
                    | Term.TERM _ => NONE
         in
           (Seq.make (clean_unify' useq))
--- a/src/ZF/AC/AC15_WO6.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/AC15_WO6.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/AC15_WO6.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 The proofs needed to state that AC10, ..., AC15 are equivalent to the rest.
@@ -56,9 +55,9 @@
 
 lemma lemma3: 
      "\<forall>B \<in> {cons(0, x*nat). x \<in> A}. pairwise_disjoint(f`B) &   
-	     sets_of_size_between(f`B, 2, n) & Union(f`B)=B   
+             sets_of_size_between(f`B, 2, n) & Union(f`B)=B   
      ==> \<forall>B \<in> A. \<exists>! u. u \<in> f`cons(0, B*nat) & u \<subseteq> cons(0, B*nat) &   
-	     0 \<in> u & 2 \<lesssim> u & u \<lesssim> n"
+             0 \<in> u & 2 \<lesssim> u & u \<lesssim> n"
 apply (unfold sets_of_size_between_def)
 apply (rule ballI)
 apply (erule_tac x="cons(0, B*nat)" in ballE)
@@ -109,7 +108,7 @@
          n \<in> nat |]   
       ==> \<exists>f. \<forall>B \<in> A. f`B \<noteq> 0 & f`B \<subseteq> B & f`B \<lesssim> n"
 by (fast del: subsetI notI
-	 dest!: lemma3 theI intro!: lemma5_1 lemma5_2 lemma5_3)
+         dest!: lemma3 theI intro!: lemma5_1 lemma5_2 lemma5_3)
 
 
 (* ********************************************************************** *)
@@ -165,7 +164,7 @@
 apply (frule HH_subset_imp_eq)
 apply (erule ssubst)
 apply (blast dest!: HH_subset_x_imp_subset_Diff_UN [THEN not_emptyI2])
-	(*but can't use del: DiffE despite the obvious conflict*)
+        (*but can't use del: DiffE despite the obvious conflict*)
 done
 
 theorem AC15_WO6: "AC15 ==> WO6"
--- a/src/ZF/AC/AC16_WO4.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/AC16_WO4.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/AC16_WO4.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 The proof of AC16(n, k) ==> WO4(n-k)
@@ -212,7 +211,7 @@
       and GG_def:    "GG  == \<lambda>v \<in> LL. (THE w. w \<in> MM & v \<subseteq> w) - v"
       and s_def:     "s(u) == {v \<in> t_n. u \<in> v & k \<lesssim> v Int y}"
   assumes all_ex:    "\<forall>z \<in> {z \<in> Pow(x Un y) . z \<approx> succ(k)}.
-	               \<exists>! w. w \<in> t_n & z \<subseteq> w "
+                       \<exists>! w. w \<in> t_n & z \<subseteq> w "
     and disjoint[iff]:  "x Int y = 0"
     and "includes":  "t_n \<subseteq> {v \<in> Pow(x Un y). v \<approx> succ(k #+ m)}"
     and WO_R[iff]:      "well_ord(y,R)"
@@ -241,7 +240,7 @@
                     Card_cardinal [THEN Card_is_Ord, THEN nat_le_infinite_Ord,
                                    THEN le_imp_lepoll] 
             dest: well_ord_cardinal_eqpoll 
-		   eqpoll_sym  eqpoll_imp_lepoll
+                   eqpoll_sym  eqpoll_imp_lepoll
                    n_lesspoll_nat [THEN lesspoll_trans2]
                    well_ord_cardinal_eqpoll [THEN eqpoll_sym, 
                           THEN eqpoll_imp_lepoll, THEN lepoll_infinite])+
@@ -535,9 +534,9 @@
 apply (insert "includes")
 apply (rule eqpoll_sum_imp_Diff_lepoll)
 apply (blast del: subsetI
-	     dest!: ltD 
-	     intro!: eqpoll_sum_imp_Diff_lepoll in_LL_eqpoll_n
-	     intro: in_LL   unique_superset1 [THEN in_MM_eqpoll_n] 
+             dest!: ltD 
+             intro!: eqpoll_sum_imp_Diff_lepoll in_LL_eqpoll_n
+             intro: in_LL   unique_superset1 [THEN in_MM_eqpoll_n] 
                     ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, 
                                   THEN apply_type])+
 done
--- a/src/ZF/AC/AC16_lemmas.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/AC16_lemmas.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/AC16_lemmas.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 Lemmas used in the proofs concerning AC16
--- a/src/ZF/AC/AC17_AC1.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/AC17_AC1.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/AC1_AC17.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 The equivalence of AC0, AC1 and AC17
--- a/src/ZF/AC/AC18_AC19.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/AC18_AC19.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/AC18_AC19.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 The proof of AC1 ==> AC18 ==> AC19 ==> AC1
--- a/src/ZF/AC/AC7_AC9.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/AC7_AC9.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/AC7_AC9.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 The proofs needed to state that AC7, AC8 and AC9 are equivalent to the previous
@@ -77,7 +76,7 @@
 apply (rule AC7_AC6_lemma1)
 apply (erule allE) 
 apply (blast del: notI
-	     intro!: AC7_AC6_lemma2 intro: eqpoll_sym eqpoll_trans 
+             intro!: AC7_AC6_lemma2 intro: eqpoll_sym eqpoll_trans 
                     Sigma_fun_space_eqpoll)
 done
 
--- a/src/ZF/AC/AC_Equiv.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/AC_Equiv.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/AC_Equiv.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 Axioms AC1 -- AC19 come from "Equivalents of the Axiom of Choice, II"
@@ -29,14 +28,14 @@
 
 definition  
     "WO4(m) == \<forall>A. \<exists>a f. Ord(a) & domain(f)=a &   
-		         (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
+                         (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
 
 definition  
     "WO5 == \<exists>m \<in> nat. 1\<le>m & WO4(m)"
 
 definition  
     "WO6 == \<forall>A. \<exists>m \<in> nat. 1\<le>m & (\<exists>a f. Ord(a) & domain(f)=a
-		               & (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m))"
+                               & (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m))"
 
 definition  
     "WO7 == \<forall>A. Finite(A) <-> (\<forall>R. well_ord(A,R) --> well_ord(A,converse(R)))"
@@ -64,7 +63,7 @@
 
 definition
     "AC2 == \<forall>A. 0\<notin>A & pairwise_disjoint(A)   
-		   --> (\<exists>C. \<forall>B \<in> A. \<exists>y. B Int C = {y})"
+                   --> (\<exists>C. \<forall>B \<in> A. \<exists>y. B Int C = {y})"
 definition
     "AC3 == \<forall>A B. \<forall>f \<in> A->B. \<exists>g. g \<in> (\<Pi> x \<in> {a \<in> A. f`a\<noteq>0}. f`x)"
 
@@ -82,24 +81,24 @@
 
 definition
     "AC8 == \<forall>A. (\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> & B1\<approx>B2)   
-		   --> (\<exists>f. \<forall>B \<in> A. f`B \<in> bij(fst(B),snd(B)))"
+                   --> (\<exists>f. \<forall>B \<in> A. f`B \<in> bij(fst(B),snd(B)))"
 
 definition
     "AC9 == \<forall>A. (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) -->   
-		   (\<exists>f. \<forall>B1 \<in> A. \<forall>B2 \<in> A. f`<B1,B2> \<in> bij(B1,B2))"
+                   (\<exists>f. \<forall>B1 \<in> A. \<forall>B2 \<in> A. f`<B1,B2> \<in> bij(B1,B2))"
 
 definition
     "AC10(n) ==  \<forall>A. (\<forall>B \<in> A. ~Finite(B)) -->   
-		   (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
-		   sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B))"
+                   (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
+                   sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B))"
 
 definition
     "AC11 == \<exists>n \<in> nat. 1\<le>n & AC10(n)"
 
 definition
     "AC12 == \<forall>A. (\<forall>B \<in> A. ~Finite(B)) -->
-  	         (\<exists>n \<in> nat. 1\<le>n & (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
-	              sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B)))"
+                 (\<exists>n \<in> nat. 1\<le>n & (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
+                      sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B)))"
 
 definition
     "AC13(m) == \<forall>A. 0\<notin>A --> (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m)"
@@ -114,12 +113,12 @@
 definition
     "AC16(n, k)  == 
        \<forall>A. ~Finite(A) -->   
-	   (\<exists>T. T \<subseteq> {X \<in> Pow(A). X\<approx>succ(n)} &   
-	   (\<forall>X \<in> {X \<in> Pow(A). X\<approx>succ(k)}. \<exists>! Y. Y \<in> T & X \<subseteq> Y))"
+           (\<exists>T. T \<subseteq> {X \<in> Pow(A). X\<approx>succ(n)} &   
+           (\<forall>X \<in> {X \<in> Pow(A). X\<approx>succ(k)}. \<exists>! Y. Y \<in> T & X \<subseteq> Y))"
 
 definition
     "AC17 == \<forall>A. \<forall>g \<in> (Pow(A)-{0} -> A) -> Pow(A)-{0}.   
-		   \<exists>f \<in> Pow(A)-{0} -> A. f`(g`f) \<in> g`f"
+                   \<exists>f \<in> Pow(A)-{0} -> A. f`(g`f) \<in> g`f"
 
 locale AC18 =
   assumes AC18: "A\<noteq>0 & (\<forall>a \<in> A. B(a) \<noteq> 0) -->
@@ -129,7 +128,7 @@
 
 definition
     "AC19 == \<forall>A. A\<noteq>0 & 0\<notin>A --> ((\<Inter>a \<in> A. \<Union>b \<in> a. b) =   
-		   (\<Union>f \<in> (\<Pi> B \<in> A. B). \<Inter>a \<in> A. f`a))"
+                   (\<Union>f \<in> (\<Pi> B \<in> A. B). \<Inter>a \<in> A. f`a))"
 
 
 
--- a/src/ZF/AC/Cardinal_aux.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/Cardinal_aux.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,7 @@
 (*  Title:      ZF/AC/Cardinal_aux.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
-Auxiliary lemmas concerning cardinalities
+Auxiliary lemmas concerning cardinalities.
 *)
 
 theory Cardinal_aux imports AC_Equiv begin
--- a/src/ZF/AC/DC.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/DC.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,7 @@
 (*  Title:      ZF/AC/DC.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
-The proofs concerning the Axiom of Dependent Choice
+The proofs concerning the Axiom of Dependent Choice.
 *)
 
 theory DC
@@ -83,8 +82,8 @@
 definition
   DC  :: "i => o"  where
     "DC(a) == \<forall>X R. R \<subseteq> Pow(X)*X  &
-		    (\<forall>Y \<in> Pow(X). Y \<prec> a --> (\<exists>x \<in> X. <Y,x> \<in> R)) 
-		    --> (\<exists>f \<in> a->X. \<forall>b<a. <f``b,f`b> \<in> R)"
+                    (\<forall>Y \<in> Pow(X). Y \<prec> a --> (\<exists>x \<in> X. <Y,x> \<in> R)) 
+                    --> (\<exists>f \<in> a->X. \<forall>b<a. <f``b,f`b> \<in> R)"
 
 definition
   DC0 :: o  where
@@ -94,7 +93,7 @@
 definition
   ff  :: "[i, i, i, i] => i"  where
     "ff(b, X, Q, R) ==
-	   transrec(b, %c r. THE x. first(x, {x \<in> X. <r``c, x> \<in> R}, Q))"
+           transrec(b, %c r. THE x. first(x, {x \<in> X. <r``c, x> \<in> R}, Q))"
 
 
 locale DC0_imp =
@@ -180,7 +179,7 @@
 apply (rule rev_bexI, assumption)
 apply (subgoal_tac "0 \<in> y", force)
 apply (force simp add: RR_def
-	     intro: ltD elim!: nat_0_le [THEN leE])
+             intro: ltD elim!: nat_0_le [THEN leE])
 (** LEVEL 7, other subgoal **)
 apply (drule bspec [OF _ nat_succI], assumption)
 apply (subgoal_tac "f ` succ (succ (x)) \<in> succ (k) ->X")
@@ -218,7 +217,7 @@
  apply (simp add: RR_def)
 apply (drule lemma2, assumption+)
 apply (blast dest!: domain_of_fun 
-	     intro: nat_into_Ord OrdmemD [THEN subsetD])
+             intro: nat_into_Ord OrdmemD [THEN subsetD])
 done
 
 lemma (in DC0_imp) lemma3:
@@ -268,10 +267,10 @@
 
     RR = {<z1,z2>:Fin(XX)*XX. 
            (domain(z2)=succ(\<Union>f \<in> z1. domain(f)) &
-	    (\<forall>f \<in> z1. restrict(z2, domain(f)) = f)) |      
-	   (~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f)) &
-	                (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) &           
-	    z2={<0,x>})}                                          
+            (\<forall>f \<in> z1. restrict(z2, domain(f)) = f)) |      
+           (~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f)) &
+                        (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) &           
+            z2={<0,x>})}                                          
                                                                           
    Then XX and RR satisfy the hypotheses of DC(omega).                    
    So applying DC:                                                        
@@ -291,26 +290,26 @@
             (\<Union>n \<in> nat. {f \<in> succ(n)->X. \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})"
 apply (rule nat_0I [THEN UN_I])
 apply (force simp add: singleton_0 [symmetric]
-	     intro!: singleton_fun [THEN Pi_type])
+             intro!: singleton_fun [THEN Pi_type])
 done
 
 
 locale imp_DC0 =
   fixes XX and RR and x and R and f and allRR
   defines XX_def: "XX == (\<Union>n \<in> nat.
-		      {f \<in> succ(n)->domain(R). \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})"
+                      {f \<in> succ(n)->domain(R). \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})"
       and RR_def:
-	 "RR == {<z1,z2>:Fin(XX)*XX. 
-		  (domain(z2)=succ(\<Union>f \<in> z1. domain(f))  
-		    & (\<forall>f \<in> z1. restrict(z2, domain(f)) = f))
-		  | (~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f))  
-		     & (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}"
+         "RR == {<z1,z2>:Fin(XX)*XX. 
+                  (domain(z2)=succ(\<Union>f \<in> z1. domain(f))  
+                    & (\<forall>f \<in> z1. restrict(z2, domain(f)) = f))
+                  | (~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f))  
+                     & (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}"
       and allRR_def:
-	"allRR == \<forall>b<nat.
-		   <f``b, f`b> \<in>  
-		    {<z1,z2>\<in>Fin(XX)*XX. (domain(z2)=succ(\<Union>f \<in> z1. domain(f))
-				    & (\<Union>f \<in> z1. domain(f)) = b  
-				    & (\<forall>f \<in> z1. restrict(z2,domain(f)) = f))}"
+        "allRR == \<forall>b<nat.
+                   <f``b, f`b> \<in>  
+                    {<z1,z2>\<in>Fin(XX)*XX. (domain(z2)=succ(\<Union>f \<in> z1. domain(f))
+                                    & (\<Union>f \<in> z1. domain(f)) = b  
+                                    & (\<forall>f \<in> z1. restrict(z2,domain(f)) = f))}"
 
 lemma (in imp_DC0) lemma4:
      "[| range(R) \<subseteq> domain(R);  x \<in> domain(R) |]   
@@ -417,10 +416,10 @@
 prefer 2 apply (blast intro: cons_fun_type2) 
 apply (rule conjI)
 prefer 2 apply (fast del: ballI subsetI
-		 elim: trans [OF _ subst_context, THEN domain_cons_eq_succ]
-		       subst_context
-		       all_in_image_restrict_eq [simplified XX_def]
-		       trans equalityD1)
+                 elim: trans [OF _ subst_context, THEN domain_cons_eq_succ]
+                       subst_context
+                       all_in_image_restrict_eq [simplified XX_def]
+                       trans equalityD1)
 (*one remaining subgoal*)
 apply (rule ballI)
 apply (erule succE)
--- a/src/ZF/AC/HH.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/HH.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/HH.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 Some properties of the recursive definition of HH used in the proofs of
@@ -104,7 +103,7 @@
 apply (rule_tac x = "\<lambda>j \<in> i. HH (f, x, j) " in exI)
 apply (simp (no_asm_simp))
 apply (fast del: DiffE
-	    elim!: HH_eq_imp_arg_eq Ord_in_Ord HH_subset_x_lt_too 
+            elim!: HH_eq_imp_arg_eq Ord_in_Ord HH_subset_x_lt_too 
             intro!: lam_type ballI ltI intro: bexI)
 done
 
@@ -113,7 +112,7 @@
 prefer 2 apply assumption 
 apply (fast del: DiffE
             intro!: Ord_Hartog 
-	    dest!: HH_subset_x_imp_lepoll 
+            dest!: HH_subset_x_imp_lepoll 
             elim!: Hartog_lepoll_selfE)
 done
 
@@ -205,7 +204,7 @@
      "f \<in> (\<Pi> X \<in> Pow(x)-{0}. F(X))   
       ==> (\<lambda>X \<in> Pow(x)-{0}. {f`X}) \<in> (\<Pi> X \<in> Pow(x)-{0}. {{z}. z \<in> F(X)})"
 by (fast del: DiffI DiffE
-	    intro!: lam_type singleton_eq_iff [THEN iffD2] dest: apply_type)
+            intro!: lam_type singleton_eq_iff [THEN iffD2] dest: apply_type)
 
 (*FIXME: both uses have the form ...[THEN bij_converse_bij], so 
   simplification is needed!*)
--- a/src/ZF/AC/Hartog.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/Hartog.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/Hartog.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 Hartog's function.
--- a/src/ZF/AC/ROOT.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/ROOT.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/ROOT.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1995  University of Cambridge
 
--- a/src/ZF/AC/WO1_AC.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/WO1_AC.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/WO1_AC.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 The proofs of WO1 ==> AC1 and WO1 ==> AC10(n) for n >= 1
--- a/src/ZF/AC/WO1_WO7.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/WO1_WO7.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/WO1_WO7.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, CU Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -82,7 +81,7 @@
 apply (unfold WO1_def LEMMA_def, clarify) 
 apply (blast dest: well_ord_converse_Memrel
                    Ord_ordertype [THEN converse_Memrel_not_well_ord]
-	     intro: ordertype_ord_iso ord_iso_is_bij bij_is_inj lepoll_Finite
+             intro: ordertype_ord_iso ord_iso_is_bij bij_is_inj lepoll_Finite
                     lepoll_def [THEN def_imp_iff, THEN iffD2] )
 done
 
--- a/src/ZF/AC/WO2_AC16.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/WO2_AC16.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,17 +1,16 @@
 (*  Title:      ZF/AC/WO2_AC16.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
-  The proof of WO2 ==> AC16(k #+ m, k)
+The proof of WO2 ==> AC16(k #+ m, k)
   
-  The main part of the proof is the inductive reasoning concerning
-  properties of constructed family T_gamma.
-  The proof deals with three cases for ordinals: 0, succ and limit ordinal.
-  The first instance is trivial, the third not difficult, but the second
-  is very complicated requiring many lemmas.
-  We also need to prove that at any stage gamma the set 
-  (s - Union(...) - k_gamma)   (Rubin & Rubin page 15)
-  contains m distinct elements (in fact is equipollent to s)
+The main part of the proof is the inductive reasoning concerning
+properties of constructed family T_gamma.
+The proof deals with three cases for ordinals: 0, succ and limit ordinal.
+The first instance is trivial, the third not difficult, but the second
+is very complicated requiring many lemmas.
+We also need to prove that at any stage gamma the set 
+(s - Union(...) - k_gamma)   (Rubin & Rubin page 15)
+contains m distinct elements (in fact is equipollent to s)
 *)
 
 theory WO2_AC16 imports AC_Equiv AC16_lemmas Cardinal_aux begin
@@ -300,9 +299,9 @@
 
 lemma dbl_Diff_eqpoll:
      "[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};   
-	 Card(a); ~ Finite(a); A\<approx>a;   
-	 k \<in> nat;  y<a;   
-	 h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)}) |]   
+         Card(a); ~ Finite(a); A\<approx>a;   
+         k \<in> nat;  y<a;   
+         h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)}) |]   
       ==> A - Union(recfunAC16(f, h, y, a)) - h`y\<approx>a"
 apply (rule dbl_Diff_eqpoll_Card, simp_all)
 apply (simp add: Union_recfunAC16_lesspoll)
@@ -430,14 +429,14 @@
 
 lemma ex_next_Ord:
      "[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};   
-	 Card(a); ~ Finite(a); A\<approx>a;   
-	 k \<in> nat; m \<in> nat; y<a;   
-	 h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});   
-	 f \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k #+ m)});   
-	 ~ (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y) |]   
+         Card(a); ~ Finite(a); A\<approx>a;   
+         k \<in> nat; m \<in> nat; y<a;   
+         h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});   
+         f \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k #+ m)});   
+         ~ (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y) |]   
       ==> \<exists>c<a. h`y \<subseteq> f`c &   
-	        (\<forall>b<a. h`b \<subseteq> f`c -->   
-		(\<forall>T \<in> recfunAC16(f, h, y, a). ~ h`b \<subseteq> T))"
+                (\<forall>b<a. h`b \<subseteq> f`c -->   
+                (\<forall>T \<in> recfunAC16(f, h, y, a). ~ h`b \<subseteq> T))"
 apply (drule ex_next_set, assumption+)
 apply (erule bexE)
 apply (rule_tac x="converse(f)`X" in oexI)
@@ -569,7 +568,7 @@
        in bij_is_surj [THEN surj_image_eq, THEN subst], assumption)
 apply (rule lemma_simp_induct)
 apply (blast del: conjI notI
-	     intro!: main_induct eqpoll_imp_lepoll [THEN lepoll_infinite] ) 
+             intro!: main_induct eqpoll_imp_lepoll [THEN lepoll_infinite] ) 
 apply (blast intro: bij_is_fun [THEN surj_image, THEN surj_is_fun])
 apply (erule eqpoll_imp_lepoll [THEN lepoll_infinite, 
                                 THEN infinite_Card_is_InfCard, 
--- a/src/ZF/AC/WO6_WO1.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/AC/WO6_WO1.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/AC/WO6_WO1.thy
-    ID:         $Id$
     Author:     Krzysztof Grabczewski
 
 Proofs needed to state that formulations WO1,...,WO6 are all equivalent.
@@ -58,7 +57,7 @@
 definition
   gg2 :: "[i, i, i, i] => i"  where
      "gg2(f,a,b,s) ==
-	      \<lambda>g \<in> a++a. if g<a then vv2(f,b,g,s) else ww2(f,b,g--a,s)"
+              \<lambda>g \<in> a++a. if g<a then vv2(f,b,g,s) else ww2(f,b,g--a,s)"
 
 
 lemma WO2_WO3: "WO2 ==> WO3"
@@ -315,9 +314,9 @@
 
 lemma vv2_subset: 
      "[| \<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 -->             
-		       domain(uu(f, b, g, d)) \<approx> succ(m);
-	 \<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;
-	 (\<Union>b<a. f`b)=y;  b<a;  g<a;  m \<in> nat;  s \<in> f`b |] 
+                       domain(uu(f, b, g, d)) \<approx> succ(m);
+         \<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;
+         (\<Union>b<a. f`b)=y;  b<a;  g<a;  m \<in> nat;  s \<in> f`b |] 
       ==> vv2(f,b,g,s) \<subseteq> f`g"
 apply (simp add: vv2_def)
 apply (blast intro: uu_Least_is_fun [THEN apply_type])
--- a/src/ZF/Bin.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Bin.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Bin.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
 
@@ -25,7 +24,7 @@
 datatype
   "bin" = Pls
         | Min
-        | Bit ("w: bin", "b: bool")	(infixl "BIT" 90)
+        | Bit ("w: bin", "b: bool")     (infixl "BIT" 90)
 
 use "Tools/numeral_syntax.ML"
 
@@ -70,7 +69,7 @@
     "bin_minus (Min)       = Pls BIT 1"
   bin_minus_BIT:
     "bin_minus (w BIT b) = cond(b, bin_pred(NCons(bin_minus(w),0)),
-				bin_minus(w) BIT 0)"
+                                bin_minus(w) BIT 0)"
 
 primrec (*sum*)
   bin_adder_Pls:
@@ -90,7 +89,7 @@
   "adding (v,x,Pls)     = v BIT x"
   "adding (v,x,Min)     = bin_pred(v BIT x)"
   "adding (v,x,w BIT y) = NCons(bin_adder (v, cond(x and y, bin_succ(w), w)), 
-				x xor y)"
+                                x xor y)"
 *)
 
 definition
@@ -105,7 +104,7 @@
     "bin_mult (Min,w)     = bin_minus(w)"
   bin_mult_BIT:
     "bin_mult (v BIT b,w) = cond(b, bin_add(NCons(bin_mult(v,w),0),w),
-				 NCons(bin_mult(v,w),0))"
+                                 NCons(bin_mult(v,w),0))"
 
 setup NumeralSyntax.setup
 
--- a/src/ZF/Cardinal.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Cardinal.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/Cardinal.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
-
 *)
 
 header{*Cardinal Numbers Without the Axiom of Choice*}
@@ -60,7 +58,7 @@
 lemma Banach_last_equation:
     "g: Y->X
      ==> g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) =        
-	 X - lfp(X, %W. X - g``(Y - f``W))" 
+         X - lfp(X, %W. X - g``(Y - f``W))" 
 apply (rule_tac P = "%u. ?v = X-u" 
        in decomp_bnd_mono [THEN lfp_unfold, THEN ssubst])
 apply (simp add: double_complement  fun_is_rel [THEN image_subset])
--- a/src/ZF/CardinalArith.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/CardinalArith.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/CardinalArith.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
-
 *)
 
 header{*Cardinal Arithmetic Without the Axiom of Choice*}
@@ -24,9 +22,9 @@
 definition
   csquare_rel   :: "i=>i"  where
     "csquare_rel(K) ==   
-	  rvimage(K*K,   
-		  lam <x,y>:K*K. <x Un y, x, y>, 
-		  rmult(K,Memrel(K), K*K, rmult(K,Memrel(K), K,Memrel(K))))"
+          rvimage(K*K,   
+                  lam <x,y>:K*K. <x Un y, x, y>, 
+                  rmult(K,Memrel(K), K*K, rmult(K,Memrel(K), K,Memrel(K))))"
 
 definition
   jump_cardinal :: "i=>i"  where
--- a/src/ZF/Coind/ECR.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Coind/ECR.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Coind/ECR.thy
-    ID:         $Id$
     Author:     Jacob Frost, Cambridge University Computer Laboratory
     Copyright   1995  University of Cambridge
 *)
@@ -29,8 +28,8 @@
 definition
   hastyenv :: "[i,i] => o"  where
     "hastyenv(ve,te) ==                        
-	 ve_dom(ve) = te_dom(te) &          
-	 (\<forall>x \<in> ve_dom(ve). <ve_app(ve,x),te_app(te,x)> \<in> HasTyRel)"
+         ve_dom(ve) = te_dom(te) &          
+         (\<forall>x \<in> ve_dom(ve). <ve_app(ve,x),te_app(te,x)> \<in> HasTyRel)"
 
 (* Specialised co-induction rule *)
 
--- a/src/ZF/Coind/Static.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Coind/Static.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Coind/Static.thy
-    ID:         $Id$
     Author:     Jacob Frost, Cambridge University Computer Laboratory
     Copyright   1995  University of Cambridge
 *)
@@ -23,7 +22,7 @@
    "isofenv(ve,te) ==                
       ve_dom(ve) = te_dom(te) &            
       (\<forall>x \<in> ve_dom(ve).                          
-	\<exists>c \<in> Const. ve_app(ve,x) = v_const(c) & isof(c,te_app(te,x)))"
+        \<exists>c \<in> Const. ve_app(ve,x) = v_const(c) & isof(c,te_app(te,x)))"
   
 
 (*** Elaboration ***)
--- a/src/ZF/Constructible/AC_in_L.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/AC_in_L.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/AC_in_L.thy
-    ID: $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -227,7 +226,7 @@
     --{*wellordering on (environment, formula) pairs*}
    "env_form_r(f,r,A) ==
       rmult(list(A), rlist(A, r),
-	    formula, measure(formula, enum(f)))"
+            formula, measure(formula, enum(f)))"
 
 definition
   env_form_map :: "[i,i,i,i]=>i" where
@@ -327,10 +326,10 @@
       lets us remove the premise @{term "Limit(i)"} from some theorems.*}
     "rlimit(i,r) ==
        if Limit(i) then 
-	 {z: Lset(i) * Lset(i).
-	  \<exists>x' x. z = <x',x> &
-		 (lrank(x') < lrank(x) |
-		  (lrank(x') = lrank(x) & <x',x> \<in> r(succ(lrank(x)))))}
+         {z: Lset(i) * Lset(i).
+          \<exists>x' x. z = <x',x> &
+                 (lrank(x') < lrank(x) |
+                  (lrank(x') = lrank(x) & <x',x> \<in> r(succ(lrank(x)))))}
        else 0"
 
 definition
--- a/src/ZF/Constructible/DPow_absolute.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/DPow_absolute.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/DPow_absolute.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
--- a/src/ZF/Constructible/Datatype_absolute.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Datatype_absolute.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Datatype_absolute.thy
-    ID: $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -27,7 +26,7 @@
      "[|i \<in> nat; j \<in> nat; bnd_mono(D,h)|] ==> i \<le> j --> h^i(0) \<subseteq> h^j(0)"
 apply (rule_tac m=i and n=j in diff_induct, simp_all)
 apply (blast del: subsetI
-	     intro: bnd_mono_iterates_subset bnd_monoD2 [of concl: h]) 
+             intro: bnd_mono_iterates_subset bnd_monoD2 [of concl: h]) 
 done
 
 lemma directed_iterates: "bnd_mono(D,h) ==> directed({h^n (0). n\<in>nat})"
@@ -178,7 +177,7 @@
 lemma list_fun_bnd_mono: "bnd_mono(univ(A), \<lambda>X. {0} + A*X)"
 apply (rule bnd_monoI)
  apply (intro subset_refl zero_subset_univ A_subset_univ 
-	      sum_subset_univ Sigma_subset_univ) 
+              sum_subset_univ Sigma_subset_univ) 
 apply (rule subset_refl sum_mono Sigma_mono | assumption)+
 done
 
@@ -226,7 +225,7 @@
      "bnd_mono(univ(0), \<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))"
 apply (rule bnd_monoI)
  apply (intro subset_refl zero_subset_univ A_subset_univ 
-	      sum_subset_univ Sigma_subset_univ nat_subset_univ) 
+              sum_subset_univ Sigma_subset_univ nat_subset_univ) 
 apply (rule subset_refl sum_mono Sigma_mono | assumption)+
 done
 
@@ -678,8 +677,8 @@
 
 theorem (in M_eclose) transrec_closed:
      "[|transrec_replacement(M,MH,i);  relation2(M,MH,H);
-	Ord(i);  M(i);
-	\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|]
+        Ord(i);  M(i);
+        \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|]
       ==> M(transrec(i,H))"
 by (simp add: trans_wfrec_closed transrec_replacement_def is_transrec_def
         transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel)
@@ -768,7 +767,7 @@
   is_Member :: "[i=>o,i,i,i] => o" where
      --{* because @{term "Member(x,y) \<equiv> Inl(Inl(\<langle>x,y\<rangle>))"}*}
     "is_Member(M,x,y,Z) ==
-	\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)"
+        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)"
 
 lemma (in M_trivial) Member_abs [simp]:
      "[|M(x); M(y); M(Z)|] ==> is_Member(M,x,y,Z) <-> (Z = Member(x,y))"
@@ -782,7 +781,7 @@
   is_Equal :: "[i=>o,i,i,i] => o" where
      --{* because @{term "Equal(x,y) \<equiv> Inl(Inr(\<langle>x,y\<rangle>))"}*}
     "is_Equal(M,x,y,Z) ==
-	\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)"
+        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)"
 
 lemma (in M_trivial) Equal_abs [simp]:
      "[|M(x); M(y); M(Z)|] ==> is_Equal(M,x,y,Z) <-> (Z = Equal(x,y))"
@@ -795,7 +794,7 @@
   is_Nand :: "[i=>o,i,i,i] => o" where
      --{* because @{term "Nand(x,y) \<equiv> Inr(Inl(\<langle>x,y\<rangle>))"}*}
     "is_Nand(M,x,y,Z) ==
-	\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)"
+        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)"
 
 lemma (in M_trivial) Nand_abs [simp]:
      "[|M(x); M(y); M(Z)|] ==> is_Nand(M,x,y,Z) <-> (Z = Nand(x,y))"
@@ -931,7 +930,7 @@
            is_c, \<lambda>u v. c(u, v, h`succ(depth(u))`u, h`succ(depth(v))`v));
         Relation1(M, formula,
            is_d, \<lambda>u. d(u, h ` succ(depth(u)) ` u));
- 	M(h) |]
+        M(h) |]
       ==> Relation1(M, formula,
                          is_formula_case (M, is_a, is_b, is_c, is_d),
                          formula_rec_case(a, b, c, d, h))"
@@ -947,9 +946,9 @@
   fixes a and is_a and b and is_b and c and is_c and d and is_d and MH
   defines
       "MH(u::i,f,z) ==
-	\<forall>fml[M]. is_formula(M,fml) -->
+        \<forall>fml[M]. is_formula(M,fml) -->
              is_lambda
-	 (M, fml, is_formula_case (M, is_a, is_b, is_c(f), is_d(f)), z)"
+         (M, fml, is_formula_case (M, is_a, is_b, is_c(f), is_d(f)), z)"
 
   assumes a_closed: "[|x\<in>nat; y\<in>nat|] ==> M(a(x,y))"
       and a_rel:    "Relation2(M, nat, nat, is_a, a)"
@@ -969,8 +968,8 @@
       and fr_lam_replace:
            "M(g) ==>
             strong_replacement
-	      (M, \<lambda>x y. x \<in> formula &
-		  y = \<langle>x, formula_rec_case(a,b,c,d,g,x)\<rangle>)";
+              (M, \<lambda>x y. x \<in> formula &
+                  y = \<langle>x, formula_rec_case(a,b,c,d,g,x)\<rangle>)";
 
 lemma (in Formula_Rec) formula_rec_case_closed:
     "[|M(g); p \<in> formula|] ==> M(formula_rec_case(a, b, c, d, g, p))"
--- a/src/ZF/Constructible/Formula.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Formula.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Formula.thy
-    ID: $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -975,7 +974,7 @@
    ==> \<forall>env \<in> list(A). sats(A,p, env@[0]) <-> sats(A,p,env)"
 apply (induct_tac p)
 apply (simp_all del: app_Cons add: app_Cons [symmetric]
-		add: nth_zero_eq_0 nth_append not_lt_iff_le nth_eq_0)
+                add: nth_zero_eq_0 nth_append not_lt_iff_le nth_eq_0)
 done
 
 lemma sats_app_zeroes_iff:
@@ -992,7 +991,7 @@
               (\<forall>a\<in>A. sats(A,p,Cons(a,env')) <-> sats(A,p,Cons(a,env)))"
 apply (rule_tac x="env @ repeat(0,arity(p))" in bexI) 
 apply (simp del: app_Cons add: app_Cons [symmetric]
-	    add: length_repeat sats_app_zeroes_iff, typecheck)
+            add: length_repeat sats_app_zeroes_iff, typecheck)
 done
 
 
--- a/src/ZF/Constructible/Internalize.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Internalize.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Internalize.thy
-    ID: $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -323,7 +322,7 @@
 subsubsection{*The Operator @{term is_Member}, Internalized*}
 
 (*    "is_Member(M,x,y,Z) ==
-	\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" *)
+        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" *)
 definition
   Member_fm :: "[i,i,i]=>i" where
     "Member_fm(x,y,Z) ==
@@ -356,7 +355,7 @@
 subsubsection{*The Operator @{term is_Equal}, Internalized*}
 
 (*    "is_Equal(M,x,y,Z) ==
-	\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" *)
+        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" *)
 definition
   Equal_fm :: "[i,i,i]=>i" where
     "Equal_fm(x,y,Z) ==
@@ -389,7 +388,7 @@
 subsubsection{*The Operator @{term is_Nand}, Internalized*}
 
 (*    "is_Nand(M,x,y,Z) ==
-	\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" *)
+        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" *)
 definition
   Nand_fm :: "[i,i,i]=>i" where
     "Nand_fm(x,y,Z) ==
@@ -600,10 +599,10 @@
      And(p, 
       And(pair_fm(2,0,3),
        Exists(Exists(Exists(
-	And(pair_fm(5,a#+7,2),
-	 And(upair_fm(5,5,1),
-	  And(pre_image_fm(r#+7,1,0),
-	   And(restriction_fm(f#+7,0,4), Member(2,r#+7)))))))))))))))"
+        And(pair_fm(5,a#+7,2),
+         And(upair_fm(5,5,1),
+          And(pre_image_fm(r#+7,1,0),
+           And(restriction_fm(f#+7,0,4), Member(2,r#+7)))))))))))))))"
 
 lemma is_recfun_type [TC]:
      "[| p \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat |] 
--- a/src/ZF/Constructible/L_axioms.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/L_axioms.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1309,7 +1309,7 @@
 subsubsection{*Finite Ordinals: The Predicate ``Is A Natural Number''*}
 
 (*     "finite_ordinal(M,a) == 
-	ordinal(M,a) & ~ limit_ordinal(M,a) & 
+        ordinal(M,a) & ~ limit_ordinal(M,a) & 
         (\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x))" *)
 definition
   finite_ordinal_fm :: "i=>i" where
--- a/src/ZF/Constructible/MetaExists.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/MetaExists.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/MetaExists.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
--- a/src/ZF/Constructible/Normal.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Normal.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Normal.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
--- a/src/ZF/Constructible/Rank.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Rank.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Rank.thy
-    ID:   $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -14,25 +13,25 @@
 assumes well_ord_iso_separation:
      "[| M(A); M(f); M(r) |]
       ==> separation (M, \<lambda>x. x\<in>A --> (\<exists>y[M]. (\<exists>p[M].
-		     fun_apply(M,f,x,y) & pair(M,y,x,p) & p \<in> r)))"
+                     fun_apply(M,f,x,y) & pair(M,y,x,p) & p \<in> r)))"
   and obase_separation:
      --{*part of the order type formalization*}
      "[| M(A); M(r) |]
       ==> separation(M, \<lambda>a. \<exists>x[M]. \<exists>g[M]. \<exists>mx[M]. \<exists>par[M].
-	     ordinal(M,x) & membership(M,x,mx) & pred_set(M,A,a,r,par) &
-	     order_isomorphism(M,par,r,x,mx,g))"
+             ordinal(M,x) & membership(M,x,mx) & pred_set(M,A,a,r,par) &
+             order_isomorphism(M,par,r,x,mx,g))"
   and obase_equals_separation:
      "[| M(A); M(r) |]
       ==> separation (M, \<lambda>x. x\<in>A --> ~(\<exists>y[M]. \<exists>g[M].
-			      ordinal(M,y) & (\<exists>my[M]. \<exists>pxr[M].
-			      membership(M,y,my) & pred_set(M,A,x,r,pxr) &
-			      order_isomorphism(M,pxr,r,y,my,g))))"
+                              ordinal(M,y) & (\<exists>my[M]. \<exists>pxr[M].
+                              membership(M,y,my) & pred_set(M,A,x,r,pxr) &
+                              order_isomorphism(M,pxr,r,y,my,g))))"
   and omap_replacement:
      "[| M(A); M(r) |]
       ==> strong_replacement(M,
              \<lambda>a z. \<exists>x[M]. \<exists>g[M]. \<exists>mx[M]. \<exists>par[M].
-	     ordinal(M,x) & pair(M,a,x,z) & membership(M,x,mx) &
-	     pred_set(M,A,a,r,par) & order_isomorphism(M,par,r,x,mx,g))"
+             ordinal(M,x) & pair(M,a,x,z) & membership(M,x,mx) &
+             pred_set(M,A,a,r,par) & order_isomorphism(M,par,r,x,mx,g))"
 
 
 text{*Inductive argument for Kunen's Lemma I 6.1, etc.
@@ -146,7 +145,7 @@
   omap :: "[i=>o,i,i,i] => o" where
     --{*the function that maps wosets to order types*}
    "omap(M,A,r,f) == 
-	\<forall>z[M].
+        \<forall>z[M].
          z \<in> f <-> (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> & Ord(x) & 
                         g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))"
 
@@ -213,7 +212,7 @@
  apply (blast intro: Ord_in_Ord) 
 apply (rename_tac y a g)
 apply (frule ord_iso_sym [THEN ord_iso_is_bij, THEN bij_is_fun, 
-			  THEN apply_funtype],  assumption)  
+                          THEN apply_funtype],  assumption)  
 apply (rule_tac x="converse(g)`y" in bexI)
  apply (frule_tac a="converse(g) ` y" in ord_iso_restrict_pred, assumption) 
 apply (safe elim!: predE) 
@@ -284,17 +283,17 @@
        M(A); M(r); M(f); M(i); b \<in> A |] ==> Ord(f `` Order.pred(A,b,r))"
 apply (frule wellordered_is_trans_on, assumption)
 apply (rule OrdI) 
-	prefer 2 apply (simp add: image_iff omap_iff Ord_def, blast) 
+        prefer 2 apply (simp add: image_iff omap_iff Ord_def, blast) 
 txt{*Hard part is to show that the image is a transitive set.*}
 apply (simp add: Transset_def, clarify) 
 apply (simp add: image_iff pred_iff apply_iff [OF omap_funtype [of A r f i]])
 apply (rename_tac c j, clarify)
 apply (frule omap_funtype [of A r f, THEN apply_funtype], assumption+)
 apply (subgoal_tac "j \<in> i") 
-	prefer 2 apply (blast intro: Ord_trans Ord_otype)
+        prefer 2 apply (blast intro: Ord_trans Ord_otype)
 apply (subgoal_tac "converse(f) ` j \<in> obase(M,A,r)") 
-	prefer 2 
-	apply (blast dest: wellordered_omap_bij [THEN bij_converse_bij, 
+        prefer 2 
+        apply (blast dest: wellordered_omap_bij [THEN bij_converse_bij, 
                                       THEN bij_is_fun, THEN apply_funtype])
 apply (rule_tac x="converse(f) ` j" in bexI) 
  apply (simp add: right_inverse_bij [OF wellordered_omap_bij]) 
@@ -423,9 +422,9 @@
     "is_oadd_fun(M,i,j,x,f) == 
        (\<forall>sj msj. M(sj) --> M(msj) --> 
                  successor(M,j,sj) --> membership(M,sj,msj) --> 
-	         M_is_recfun(M, 
-		     %x g y. \<exists>gx[M]. image(M,g,x,gx) & union(M,i,gx,y),
-		     msj, x, f))"
+                 M_is_recfun(M, 
+                     %x g y. \<exists>gx[M]. image(M,g,x,gx) & union(M,i,gx,y),
+                     msj, x, f))"
 
 definition
   is_oadd :: "[i=>o,i,i,i] => o" where
@@ -434,30 +433,30 @@
         (~ ordinal(M,i) & ordinal(M,j) & k=j) |
         (ordinal(M,i) & ~ ordinal(M,j) & k=i) |
         (ordinal(M,i) & ordinal(M,j) & 
-	 (\<exists>f fj sj. M(f) & M(fj) & M(sj) & 
-		    successor(M,j,sj) & is_oadd_fun(M,i,sj,sj,f) & 
-		    fun_apply(M,f,j,fj) & fj = k))"
+         (\<exists>f fj sj. M(f) & M(fj) & M(sj) & 
+                    successor(M,j,sj) & is_oadd_fun(M,i,sj,sj,f) & 
+                    fun_apply(M,f,j,fj) & fj = k))"
 
 definition
  (*NEEDS RELATIVIZATION*)
   omult_eqns :: "[i,i,i,i] => o" where
     "omult_eqns(i,x,g,z) ==
             Ord(x) & 
-	    (x=0 --> z=0) &
+            (x=0 --> z=0) &
             (\<forall>j. x = succ(j) --> z = g`j ++ i) &
             (Limit(x) --> z = \<Union>(g``x))"
 
 definition
   is_omult_fun :: "[i=>o,i,i,i] => o" where
     "is_omult_fun(M,i,j,f) == 
-	    (\<exists>df. M(df) & is_function(M,f) & 
+            (\<exists>df. M(df) & is_function(M,f) & 
                   is_domain(M,f,df) & subset(M, j, df)) & 
             (\<forall>x\<in>j. omult_eqns(i,x,f,f`x))"
 
 definition
   is_omult :: "[i=>o,i,i,i] => o" where
     "is_omult(M,i,j,k) == 
-	\<exists>f fj sj. M(f) & M(fj) & M(sj) & 
+        \<exists>f fj sj. M(f) & M(fj) & M(sj) & 
                   successor(M,j,sj) & is_omult_fun(M,i,sj,f) & 
                   fun_apply(M,f,j,fj) & fj = k"
 
@@ -468,14 +467,14 @@
     strong_replacement(M, 
          \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & 
                   (\<exists>f[M]. \<exists>fx[M]. is_oadd_fun(M,i,j,x,f) & 
-		           image(M,f,x,fx) & y = i Un fx))"
+                           image(M,f,x,fx) & y = i Un fx))"
 
  and omult_strong_replacement':
    "[| M(i); M(j) |] ==>
     strong_replacement(M, 
          \<lambda>x z. \<exists>y[M]. z = <x,y> &
-	     (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) & 
-	     y = (THE z. omult_eqns(i, x, g, z))))" 
+             (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) & 
+             y = (THE z. omult_eqns(i, x, g, z))))" 
 
 
 
@@ -483,7 +482,7 @@
 lemma (in M_ord_arith) is_oadd_fun_iff:
    "[| a\<le>j; M(i); M(j); M(a); M(f) |] 
     ==> is_oadd_fun(M,i,j,a,f) <->
-	f \<in> a \<rightarrow> range(f) & (\<forall>x. M(x) --> x < a --> f`x = i Un f``x)"
+        f \<in> a \<rightarrow> range(f) & (\<forall>x. M(x) --> x < a --> f`x = i Un f``x)"
 apply (frule lt_Ord) 
 apply (simp add: is_oadd_fun_def Memrel_closed Un_closed 
              relation2_def is_recfun_abs [of "%x g. i Un g``x"]
@@ -498,8 +497,8 @@
     "[| M(i); M(j) |] ==>
      strong_replacement(M, 
             \<lambda>x z. \<exists>y[M]. z = <x,y> &
-		  (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. i Un g``x,g) & 
-		  y = i Un g``x))" 
+                  (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. i Un g``x,g) & 
+                  y = i Un g``x))" 
 apply (insert oadd_strong_replacement [of i j]) 
 apply (simp add: is_oadd_fun_def relation2_def
                  is_recfun_abs [of "%x g. i Un g``x"])  
@@ -705,7 +704,7 @@
 lemma (in M_wfrank) wfrank_separation':
      "M(r) ==>
       separation
-	   (M, \<lambda>x. ~ (\<exists>f[M]. is_recfun(r^+, x, %x f. range(f), f)))"
+           (M, \<lambda>x. ~ (\<exists>f[M]. is_recfun(r^+, x, %x f. range(f), f)))"
 apply (insert wfrank_separation [of r])
 apply (simp add: relation2_def is_recfun_abs [of "%x. range"])
 done
@@ -713,8 +712,8 @@
 lemma (in M_wfrank) wfrank_strong_replacement':
      "M(r) ==>
       strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>f[M]. 
-		  pair(M,x,y,z) & is_recfun(r^+, x, %x f. range(f), f) &
-		  y = range(f))"
+                  pair(M,x,y,z) & is_recfun(r^+, x, %x f. range(f), f) &
+                  y = range(f))"
 apply (insert wfrank_strong_replacement [of r])
 apply (simp add: relation2_def is_recfun_abs [of "%x. range"])
 done
--- a/src/ZF/Constructible/Rank_Separation.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Rank_Separation.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Rank_Separation.thy
-    ID:   $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -117,8 +116,8 @@
 lemma M_ordertype_axioms_L: "M_ordertype_axioms(L)"
   apply (rule M_ordertype_axioms.intro)
        apply (assumption | rule well_ord_iso_separation
-	 obase_separation obase_equals_separation
-	 omap_replacement)+
+         obase_separation obase_equals_separation
+         omap_replacement)+
   done
 
 theorem M_ordertype_L: "PROP M_ordertype(L)"
--- a/src/ZF/Constructible/Reflection.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Reflection.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Reflection.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -157,7 +156,7 @@
       ==> \<exists>z. M(z) & P(<y,z>)"
 apply (simp add: ClEx_def M_def)
 apply (blast dest: Cl_reflects
-	     intro: Limit_is_Ord Pair_in_Mset)
+             intro: Limit_is_Ord Pair_in_Mset)
 done
 
 text{*Class @{text ClEx} indeed consists of reflecting ordinals...*}
@@ -298,8 +297,8 @@
     (\<lambda>a. (Ord(a) &
           ClEx(\<lambda>x. ~ (snd(x) \<subseteq> fst(fst(x)) --> snd(x) \<in> snd(fst(x))), a)) &
           ClEx(\<lambda>x. \<forall>z. M(z) --> z \<subseteq> fst(x) --> z \<in> snd(x), a),
-	    \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> z \<subseteq> x --> z \<in> y), 
-	    \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x --> z \<in> y)" 
+            \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> z \<subseteq> x --> z \<in> y), 
+            \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x --> z \<in> y)" 
 by fast
 
 text{*Example 2''.  We expand the subset relation.*}
--- a/src/ZF/Constructible/Relative.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Relative.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Relative.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -79,12 +78,12 @@
   big_inter :: "[i=>o,i,i] => o" where
     "big_inter(M,A,z) ==
              (A=0 --> z=0) &
-	     (A\<noteq>0 --> (\<forall>x[M]. x \<in> z <-> (\<forall>y[M]. y\<in>A --> x \<in> y)))"
+             (A\<noteq>0 --> (\<forall>x[M]. x \<in> z <-> (\<forall>y[M]. y\<in>A --> x \<in> y)))"
 
 definition
   cartprod :: "[i=>o,i,i,i] => o" where
     "cartprod(M,A,B,z) ==
-	\<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))"
+        \<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))"
 
 definition
   is_sum :: "[i=>o,i,i,i] => o" where
@@ -104,18 +103,18 @@
 definition
   is_converse :: "[i=>o,i,i] => o" where
     "is_converse(M,r,z) ==
-	\<forall>x[M]. x \<in> z <->
+        \<forall>x[M]. x \<in> z <->
              (\<exists>w[M]. w\<in>r & (\<exists>u[M]. \<exists>v[M]. pair(M,u,v,w) & pair(M,v,u,x)))"
 
 definition
   pre_image :: "[i=>o,i,i,i] => o" where
     "pre_image(M,r,A,z) ==
-	\<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))"
+        \<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))"
 
 definition
   is_domain :: "[i=>o,i,i] => o" where
     "is_domain(M,r,z) ==
-	\<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w)))"
+        \<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w)))"
 
 definition
   image :: "[i=>o,i,i,i] => o" where
@@ -129,12 +128,12 @@
       unfortunately needs an instance of separation in order to prove
         @{term "M(converse(r))"}.*}
     "is_range(M,r,z) ==
-	\<forall>y[M]. y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w)))"
+        \<forall>y[M]. y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w)))"
 
 definition
   is_field :: "[i=>o,i,i] => o" where
     "is_field(M,r,z) ==
-	\<exists>dr[M]. \<exists>rr[M]. is_domain(M,r,dr) & is_range(M,r,rr) &
+        \<exists>dr[M]. \<exists>rr[M]. is_domain(M,r,dr) & is_range(M,r,rr) &
                         union(M,dr,rr,z)"
 
 definition
@@ -145,7 +144,7 @@
 definition
   is_function :: "[i=>o,i] => o" where
     "is_function(M,r) ==
-	\<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
+        \<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
            pair(M,x,y,p) --> pair(M,x,y',p') --> p\<in>r --> p'\<in>r --> y=y'"
 
 definition
@@ -176,7 +175,7 @@
 definition
   injection :: "[i=>o,i,i,i] => o" where
     "injection(M,A,B,f) ==
-	typed_function(M,A,B,f) &
+        typed_function(M,A,B,f) &
         (\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
           pair(M,x,y,p) --> pair(M,x',y,p') --> p\<in>f --> p'\<in>f --> x=x')"
 
@@ -193,7 +192,7 @@
 definition
   restriction :: "[i=>o,i,i,i] => o" where
     "restriction(M,r,A,z) ==
-	\<forall>x[M]. x \<in> z <-> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))"
+        \<forall>x[M]. x \<in> z <-> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))"
 
 definition
   transitive_set :: "[i=>o,i] => o" where
@@ -208,20 +207,20 @@
   limit_ordinal :: "[i=>o,i] => o" where
     --{*a limit ordinal is a non-empty, successor-closed ordinal*}
     "limit_ordinal(M,a) ==
-	ordinal(M,a) & ~ empty(M,a) &
+        ordinal(M,a) & ~ empty(M,a) &
         (\<forall>x[M]. x\<in>a --> (\<exists>y[M]. y\<in>a & successor(M,x,y)))"
 
 definition
   successor_ordinal :: "[i=>o,i] => o" where
     --{*a successor ordinal is any ordinal that is neither empty nor limit*}
     "successor_ordinal(M,a) ==
-	ordinal(M,a) & ~ empty(M,a) & ~ limit_ordinal(M,a)"
+        ordinal(M,a) & ~ empty(M,a) & ~ limit_ordinal(M,a)"
 
 definition
   finite_ordinal :: "[i=>o,i] => o" where
     --{*an ordinal is finite if neither it nor any of its elements are limit*}
     "finite_ordinal(M,a) ==
-	ordinal(M,a) & ~ limit_ordinal(M,a) &
+        ordinal(M,a) & ~ limit_ordinal(M,a) &
         (\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x))"
 
 definition
@@ -291,7 +290,7 @@
 definition
   extensionality :: "(i=>o) => o" where
     "extensionality(M) ==
-	\<forall>x[M]. \<forall>y[M]. (\<forall>z[M]. z \<in> x <-> z \<in> y) --> x=y"
+        \<forall>x[M]. \<forall>y[M]. (\<forall>z[M]. z \<in> x <-> z \<in> y) --> x=y"
 
 definition
   separation :: "[i=>o, i=>o] => o" where
@@ -300,7 +299,7 @@
         to @{text M}.  We do not have separation as a scheme; every instance
         that we need must be assumed (and later proved) separately.*}
     "separation(M,P) ==
-	\<forall>z[M]. \<exists>y[M]. \<forall>x[M]. x \<in> y <-> x \<in> z & P(x)"
+        \<forall>z[M]. \<exists>y[M]. \<forall>x[M]. x \<in> y <-> x \<in> z & P(x)"
 
 definition
   upair_ax :: "(i=>o) => o" where
@@ -317,7 +316,7 @@
 definition
   univalent :: "[i=>o, i, [i,i]=>o] => o" where
     "univalent(M,A,P) ==
-	\<forall>x[M]. x\<in>A --> (\<forall>y[M]. \<forall>z[M]. P(x,y) & P(x,z) --> y=z)"
+        \<forall>x[M]. x\<in>A --> (\<forall>y[M]. \<forall>z[M]. P(x,y) & P(x,z) --> y=z)"
 
 definition
   replacement :: "[i=>o, [i,i]=>o] => o" where
@@ -334,7 +333,7 @@
 definition
   foundation_ax :: "(i=>o) => o" where
     "foundation_ax(M) ==
-	\<forall>x[M]. (\<exists>y[M]. y\<in>x) --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & z \<in> y))"
+        \<forall>x[M]. (\<exists>y[M]. y\<in>x) --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & z \<in> y))"
 
 
 subsection{*A trivial consistency proof for $V_\omega$ *}
@@ -513,12 +512,12 @@
 definition
   pred_set :: "[i=>o,i,i,i,i] => o" where
     "pred_set(M,A,x,r,B) ==
-	\<forall>y[M]. y \<in> B <-> (\<exists>p[M]. p\<in>r & y \<in> A & pair(M,y,x,p))"
+        \<forall>y[M]. y \<in> B <-> (\<exists>p[M]. p\<in>r & y \<in> A & pair(M,y,x,p))"
 
 definition
   membership :: "[i=>o,i,i] => o" where --{*membership relation*}
     "membership(M,A,r) ==
-	\<forall>p[M]. p \<in> r <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>A & x\<in>y & pair(M,x,y,p)))"
+        \<forall>p[M]. p \<in> r <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>A & x\<in>y & pair(M,x,y,p)))"
 
 
 subsection{*Introducing a Transitive Class Model*}
@@ -528,8 +527,8 @@
 locale M_trivial =
   fixes M
   assumes transM:           "[| y\<in>x; M(x) |] ==> M(y)"
-      and upair_ax:	    "upair_ax(M)"
-      and Union_ax:	    "Union_ax(M)"
+      and upair_ax:         "upair_ax(M)"
+      and Union_ax:         "Union_ax(M)"
       and power_ax:         "power_ax(M)"
       and replacement:      "replacement(M,P)"
       and M_nat [iff]:      "M(nat)"           (*i.e. the axiom of infinity*)
@@ -951,8 +950,8 @@
   primrec
       "natnumber_aux(M,0) = (\<lambda>x\<in>nat. if empty(M,x) then 1 else 0)"
       "natnumber_aux(M,succ(n)) =
-	   (\<lambda>x\<in>nat. if (\<exists>y[M]. natnumber_aux(M,n)`y=1 & successor(M,y,x))
-		     then 1 else 0)"
+           (\<lambda>x\<in>nat. if (\<exists>y[M]. natnumber_aux(M,n)`y=1 & successor(M,y,x))
+                     then 1 else 0)"
 
   definition
     natnumber :: "[i=>o,i,i] => o"
@@ -983,7 +982,7 @@
   and comp_separation:
      "[| M(r); M(s) |]
       ==> separation(M, \<lambda>xz. \<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
-		  pair(M,x,z,xz) & pair(M,x,y,xy) & pair(M,y,z,yz) &
+                  pair(M,x,z,xz) & pair(M,x,y,xy) & pair(M,y,z,yz) &
                   xy\<in>s & yz\<in>r)"
   and pred_separation:
      "[| M(r); M(x) |] ==> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & pair(M,y,x,p))"
--- a/src/ZF/Constructible/Satisfies_absolute.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Satisfies_absolute.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Satisfies_absolute.thy
-    ID:  $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -180,8 +179,8 @@
    --{*Merely a useful abbreviation for the sequel.*}
   "is_depth_apply(M,h,p,z) ==
     \<exists>dp[M]. \<exists>sdp[M]. \<exists>hsdp[M]. 
-	finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
-	fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z)"
+        finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
+        fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z)"
 
 lemma (in M_datatypes) is_depth_apply_abs [simp]:
      "[|M(h); p \<in> formula; M(z)|] 
@@ -205,9 +204,9 @@
    "satisfies_is_a(M,A) == 
     \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) -->
              is_lambda(M, lA, 
-		\<lambda>env z. is_bool_of_o(M, 
+                \<lambda>env z. is_bool_of_o(M, 
                       \<exists>nx[M]. \<exists>ny[M]. 
- 		       is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z),
+                       is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z),
                 zz)"
 
 definition
@@ -235,8 +234,8 @@
    "satisfies_is_c(M,A,h) == 
     \<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) -->
              is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M]. 
-		 (\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & 
-		 (\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & 
+                 (\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & 
+                 (\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & 
                  (\<exists>pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)),
                 zz)"
 
@@ -294,7 +293,7 @@
    Member_replacement:
     "[|M(A); x \<in> nat; y \<in> nat|]
      ==> strong_replacement
-	 (M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. 
+         (M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. 
               env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & 
               is_bool_of_o(M, nx \<in> ny, bo) &
               pair(M, env, bo, z))"
@@ -302,7 +301,7 @@
    Equal_replacement:
     "[|M(A); x \<in> nat; y \<in> nat|]
      ==> strong_replacement
-	 (M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. 
+         (M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. 
               env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & 
               is_bool_of_o(M, nx = ny, bo) &
               pair(M, env, bo, z))"
@@ -310,7 +309,7 @@
    Nand_replacement:
     "[|M(A); M(rp); M(rq)|]
      ==> strong_replacement
-	 (M, \<lambda>env z. \<exists>rpe[M]. \<exists>rqe[M]. \<exists>andpq[M]. \<exists>notpq[M]. 
+         (M, \<lambda>env z. \<exists>rpe[M]. \<exists>rqe[M]. \<exists>andpq[M]. \<exists>notpq[M]. 
                fun_apply(M,rp,env,rpe) & fun_apply(M,rq,env,rqe) & 
                is_and(M,rpe,rqe,andpq) & is_not(M,andpq,notpq) & 
                env \<in> list(A) & pair(M, env, notpq, z))"
@@ -318,14 +317,14 @@
   Forall_replacement:
    "[|M(A); M(rp)|]
     ==> strong_replacement
-	(M, \<lambda>env z. \<exists>bo[M]. 
-	      env \<in> list(A) & 
-	      is_bool_of_o (M, 
-			    \<forall>a[M]. \<forall>co[M]. \<forall>rpco[M]. 
-			       a\<in>A --> is_Cons(M,a,env,co) -->
-			       fun_apply(M,rp,co,rpco) --> number1(M, rpco), 
+        (M, \<lambda>env z. \<exists>bo[M]. 
+              env \<in> list(A) & 
+              is_bool_of_o (M, 
+                            \<forall>a[M]. \<forall>co[M]. \<forall>rpco[M]. 
+                               a\<in>A --> is_Cons(M,a,env,co) -->
+                               fun_apply(M,rp,co,rpco) --> number1(M, rpco), 
                             bo) &
-	      pair(M,env,bo,z))"
+              pair(M,env,bo,z))"
  and
   formula_rec_replacement: 
       --{*For the @{term transrec}*}
@@ -346,29 +345,29 @@
 lemma (in M_satisfies) Member_replacement':
     "[|M(A); x \<in> nat; y \<in> nat|]
      ==> strong_replacement
-	 (M, \<lambda>env z. env \<in> list(A) &
-		     z = \<langle>env, bool_of_o(nth(x, env) \<in> nth(y, env))\<rangle>)"
+         (M, \<lambda>env z. env \<in> list(A) &
+                     z = \<langle>env, bool_of_o(nth(x, env) \<in> nth(y, env))\<rangle>)"
 by (insert Member_replacement, simp) 
 
 lemma (in M_satisfies) Equal_replacement':
     "[|M(A); x \<in> nat; y \<in> nat|]
      ==> strong_replacement
-	 (M, \<lambda>env z. env \<in> list(A) &
-		     z = \<langle>env, bool_of_o(nth(x, env) = nth(y, env))\<rangle>)"
+         (M, \<lambda>env z. env \<in> list(A) &
+                     z = \<langle>env, bool_of_o(nth(x, env) = nth(y, env))\<rangle>)"
 by (insert Equal_replacement, simp) 
 
 lemma (in M_satisfies) Nand_replacement':
     "[|M(A); M(rp); M(rq)|]
      ==> strong_replacement
-	 (M, \<lambda>env z. env \<in> list(A) & z = \<langle>env, not(rp`env and rq`env)\<rangle>)"
+         (M, \<lambda>env z. env \<in> list(A) & z = \<langle>env, not(rp`env and rq`env)\<rangle>)"
 by (insert Nand_replacement, simp) 
 
 lemma (in M_satisfies) Forall_replacement':
    "[|M(A); M(rp)|]
     ==> strong_replacement
-	(M, \<lambda>env z.
-	       env \<in> list(A) &
-	       z = \<langle>env, bool_of_o (\<forall>a\<in>A. rp ` Cons(a,env) = 1)\<rangle>)"
+        (M, \<lambda>env z.
+               env \<in> list(A) &
+               z = \<langle>env, bool_of_o (\<forall>a\<in>A. rp ` Cons(a,env) = 1)\<rangle>)"
 by (insert Forall_replacement, simp) 
 
 lemma (in M_satisfies) a_closed:
@@ -408,8 +407,8 @@
  "[|M(A); M(f)|] ==> 
   Relation2 (M, formula, formula, 
                satisfies_is_c(M,A,f),
-	       \<lambda>u v. satisfies_c(A, u, v, f ` succ(depth(u)) ` u, 
-					  f ` succ(depth(v)) ` v))"
+               \<lambda>u v. satisfies_c(A, u, v, f ` succ(depth(u)) ` u, 
+                                          f ` succ(depth(v)) ` v))"
 apply (simp add: Relation2_def satisfies_is_c_def satisfies_c_def)
 apply (auto del: iffI intro!: lambda_abs2 
             simp add: Relation1_def formula_into_M) 
@@ -467,9 +466,9 @@
 lemma (in M_satisfies) Formula_Rec_axioms_M:
    "M(A) ==>
     Formula_Rec_axioms(M, satisfies_a(A), satisfies_is_a(M,A), 
-			  satisfies_b(A), satisfies_is_b(M,A), 
-			  satisfies_c(A), satisfies_is_c(M,A), 
-			  satisfies_d(A), satisfies_is_d(M,A))"
+                          satisfies_b(A), satisfies_is_b(M,A), 
+                          satisfies_c(A), satisfies_is_c(M,A), 
+                          satisfies_d(A), satisfies_is_d(M,A))"
 apply (rule Formula_Rec_axioms.intro)
 apply (assumption | 
        rule a_closed a_rel b_closed b_rel c_closed c_rel d_closed d_rel
@@ -481,9 +480,9 @@
 theorem (in M_satisfies) Formula_Rec_M: 
     "M(A) ==>
      PROP Formula_Rec(M, satisfies_a(A), satisfies_is_a(M,A), 
-			 satisfies_b(A), satisfies_is_b(M,A), 
-			 satisfies_c(A), satisfies_is_c(M,A), 
-			 satisfies_d(A), satisfies_is_d(M,A))"
+                         satisfies_b(A), satisfies_is_b(M,A), 
+                         satisfies_c(A), satisfies_is_c(M,A), 
+                         satisfies_d(A), satisfies_is_d(M,A))"
   apply (rule Formula_Rec.intro)
    apply (rule M_satisfies.axioms, rule M_satisfies_axioms)
   apply (erule Formula_Rec_axioms_M) 
@@ -513,8 +512,8 @@
 (* is_depth_apply(M,h,p,z) ==
     \<exists>dp[M]. \<exists>sdp[M]. \<exists>hsdp[M]. 
       2        1         0
-	finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
-	fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z) *)
+        finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
+        fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z) *)
 definition
   depth_apply_fm :: "[i,i,i]=>i" where
     "depth_apply_fm(h,p,z) ==
@@ -554,9 +553,9 @@
 (* satisfies_is_a(M,A) == 
     \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) -->
              is_lambda(M, lA, 
-		\<lambda>env z. is_bool_of_o(M, 
+                \<lambda>env z. is_bool_of_o(M, 
                       \<exists>nx[M]. \<exists>ny[M]. 
- 		       is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z),
+                       is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z),
                 zz)  *)
 
 definition
@@ -656,8 +655,8 @@
 (* satisfies_is_c(M,A,h) == 
     \<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) -->
              is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M]. 
-		 (\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & 
-		 (\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & 
+                 (\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & 
+                 (\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & 
                  (\<exists>pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)),
                 zz) *)
 
@@ -836,7 +835,7 @@
 lemma Member_replacement:
     "[|L(A); x \<in> nat; y \<in> nat|]
      ==> strong_replacement
-	 (L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. 
+         (L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. 
               env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) & 
               is_bool_of_o(L, nx \<in> ny, bo) &
               pair(L, env, bo, z))"
@@ -866,7 +865,7 @@
 lemma Equal_replacement:
     "[|L(A); x \<in> nat; y \<in> nat|]
      ==> strong_replacement
-	 (L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. 
+         (L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. 
               env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) & 
               is_bool_of_o(L, nx = ny, bo) &
               pair(L, env, bo, z))"
@@ -882,10 +881,10 @@
 
 lemma Nand_Reflects:
     "REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B \<and>
-	       (\<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L].
-		 fun_apply(L, rp, u, rpe) \<and> fun_apply(L, rq, u, rqe) \<and>
-		 is_and(L, rpe, rqe, andpq) \<and> is_not(L, andpq, notpq) \<and>
-		 u \<in> list(A) \<and> pair(L, u, notpq, x)),
+               (\<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L].
+                 fun_apply(L, rp, u, rpe) \<and> fun_apply(L, rq, u, rqe) \<and>
+                 is_and(L, rpe, rqe, andpq) \<and> is_not(L, andpq, notpq) \<and>
+                 u \<in> list(A) \<and> pair(L, u, notpq, x)),
     \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and>
      (\<exists>rpe \<in> Lset(i). \<exists>rqe \<in> Lset(i). \<exists>andpq \<in> Lset(i). \<exists>notpq \<in> Lset(i).
        fun_apply(##Lset(i), rp, u, rpe) \<and> fun_apply(##Lset(i), rq, u, rqe) \<and>
@@ -898,7 +897,7 @@
 lemma Nand_replacement:
     "[|L(A); L(rp); L(rq)|]
      ==> strong_replacement
-	 (L, \<lambda>env z. \<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L]. 
+         (L, \<lambda>env z. \<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L]. 
                fun_apply(L,rp,env,rpe) & fun_apply(L,rq,env,rqe) & 
                is_and(L,rpe,rqe,andpq) & is_not(L,andpq,notpq) & 
                env \<in> list(A) & pair(L, env, notpq, z))"
@@ -923,9 +922,9 @@
  \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>bo \<in> Lset(i). u \<in> list(A) \<and>
         is_bool_of_o (##Lset(i),
  \<forall>a \<in> Lset(i). \<forall>co \<in> Lset(i). \<forall>rpco \<in> Lset(i). a \<in> A \<longrightarrow>
-	    is_Cons(##Lset(i),a,u,co) \<longrightarrow> fun_apply(##Lset(i),rp,co,rpco) \<longrightarrow> 
-	    number1(##Lset(i),rpco),
-		       bo) \<and> pair(##Lset(i),u,bo,x))]"
+            is_Cons(##Lset(i),a,u,co) \<longrightarrow> fun_apply(##Lset(i),rp,co,rpco) \<longrightarrow> 
+            number1(##Lset(i),rpco),
+                       bo) \<and> pair(##Lset(i),u,bo,x))]"
 apply (unfold is_bool_of_o_def) 
 apply (intro FOL_reflections function_reflections Cons_reflection)
 done
@@ -933,14 +932,14 @@
 lemma Forall_replacement:
    "[|L(A); L(rp)|]
     ==> strong_replacement
-	(L, \<lambda>env z. \<exists>bo[L]. 
-	      env \<in> list(A) & 
-	      is_bool_of_o (L, 
-			    \<forall>a[L]. \<forall>co[L]. \<forall>rpco[L]. 
-			       a\<in>A --> is_Cons(L,a,env,co) -->
-			       fun_apply(L,rp,co,rpco) --> number1(L, rpco), 
+        (L, \<lambda>env z. \<exists>bo[L]. 
+              env \<in> list(A) & 
+              is_bool_of_o (L, 
+                            \<forall>a[L]. \<forall>co[L]. \<forall>rpco[L]. 
+                               a\<in>A --> is_Cons(L,a,env,co) -->
+                               fun_apply(L,rp,co,rpco) --> number1(L, rpco), 
                             bo) &
-	      pair(L,env,bo,z))"
+              pair(L,env,bo,z))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{A,list(A),B,rp}" 
          in gen_separation_multi [OF Forall_Reflects],
@@ -978,18 +977,18 @@
  "REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B &
      mem_formula(L,u) &
      (\<exists>c[L].
-	 is_formula_case
-	  (L, satisfies_is_a(L,A), satisfies_is_b(L,A),
-	   satisfies_is_c(L,A,g), satisfies_is_d(L,A,g),
-	   u, c) &
-	 pair(L,u,c,x)),
+         is_formula_case
+          (L, satisfies_is_a(L,A), satisfies_is_b(L,A),
+           satisfies_is_c(L,A,g), satisfies_is_d(L,A,g),
+           u, c) &
+         pair(L,u,c,x)),
   \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & mem_formula(##Lset(i),u) &
      (\<exists>c \<in> Lset(i).
-	 is_formula_case
-	  (##Lset(i), satisfies_is_a(##Lset(i),A), satisfies_is_b(##Lset(i),A),
-	   satisfies_is_c(##Lset(i),A,g), satisfies_is_d(##Lset(i),A,g),
-	   u, c) &
-	 pair(##Lset(i),u,c,x))]"
+         is_formula_case
+          (##Lset(i), satisfies_is_a(##Lset(i),A), satisfies_is_b(##Lset(i),A),
+           satisfies_is_c(##Lset(i),A,g), satisfies_is_d(##Lset(i),A,g),
+           u, c) &
+         pair(##Lset(i),u,c,x))]"
 by (intro FOL_reflections function_reflections mem_formula_reflection
           is_formula_case_reflection satisfies_is_a_reflection
           satisfies_is_b_reflection satisfies_is_c_reflection
@@ -1022,7 +1021,7 @@
 lemma M_satisfies_axioms_L: "M_satisfies_axioms(L)"
   apply (rule M_satisfies_axioms.intro)
        apply (assumption | rule
-	 Member_replacement Equal_replacement 
+         Member_replacement Equal_replacement 
          Nand_replacement Forall_replacement
          formula_rec_replacement formula_rec_lambda_replacement)+
   done
--- a/src/ZF/Constructible/Separation.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Separation.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Separation.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -296,10 +295,10 @@
 lemma M_basic_axioms_L: "M_basic_axioms(L)"
   apply (rule M_basic_axioms.intro)
        apply (assumption | rule
-	 Inter_separation Diff_separation cartprod_separation image_separation
-	 converse_separation restrict_separation
-	 comp_separation pred_separation Memrel_separation
-	 funspace_succ_replacement is_recfun_separation)+
+         Inter_separation Diff_separation cartprod_separation image_separation
+         converse_separation restrict_separation
+         comp_separation pred_separation Memrel_separation
+         funspace_succ_replacement is_recfun_separation)+
   done
 
 theorem M_basic_L: "PROP M_basic(L)"
--- a/src/ZF/Constructible/WF_absolute.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/WF_absolute.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/WF_absolute.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -56,28 +55,28 @@
 
 lemma rtrancl_alt_eq_rtrancl: "rtrancl_alt(field(r),r) = r^*"
 by (blast del: subsetI
-	  intro: rtrancl_alt_subset_rtrancl rtrancl_subset_rtrancl_alt)
+          intro: rtrancl_alt_subset_rtrancl rtrancl_subset_rtrancl_alt)
 
 
 definition
   rtran_closure_mem :: "[i=>o,i,i,i] => o" where
     --{*The property of belonging to @{text "rtran_closure(r)"}*}
     "rtran_closure_mem(M,A,r,p) ==
-	      \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. 
+              \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. 
                omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
-	       (\<exists>f[M]. typed_function(M,n',A,f) &
-		(\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
-		  fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
-		  (\<forall>j[M]. j\<in>n --> 
-		    (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. 
-		      fun_apply(M,f,j,fj) & successor(M,j,sj) &
-		      fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"
+               (\<exists>f[M]. typed_function(M,n',A,f) &
+                (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
+                  fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
+                  (\<forall>j[M]. j\<in>n --> 
+                    (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. 
+                      fun_apply(M,f,j,fj) & successor(M,j,sj) &
+                      fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"
 
 definition
   rtran_closure :: "[i=>o,i,i] => o" where
     "rtran_closure(M,r,s) == 
         \<forall>A[M]. is_field(M,r,A) -->
- 	 (\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))"
+         (\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))"
 
 definition
   tran_closure :: "[i=>o,i,i] => o" where
@@ -96,12 +95,12 @@
 
 locale M_trancl = M_basic +
   assumes rtrancl_separation:
-	 "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))"
+         "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))"
       and wellfounded_trancl_separation:
-	 "[| M(r); M(Z) |] ==> 
-	  separation (M, \<lambda>x. 
-	      \<exists>w[M]. \<exists>wx[M]. \<exists>rp[M]. 
-	       w \<in> Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \<in> rp)"
+         "[| M(r); M(Z) |] ==> 
+          separation (M, \<lambda>x. 
+              \<exists>w[M]. \<exists>wx[M]. \<exists>rp[M]. 
+               w \<in> Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \<in> rp)"
 
 
 lemma (in M_trancl) rtran_closure_rtrancl:
--- a/src/ZF/Constructible/WFrec.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/WFrec.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/WFrec.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -22,8 +21,8 @@
 text{*Expresses @{text is_recfun} as a recursion equation*}
 lemma is_recfun_iff_equation:
      "is_recfun(r,a,H,f) <->
-	   f \<in> r -`` {a} \<rightarrow> range(f) &
-	   (\<forall>x \<in> r-``{a}. f`x = H(x, restrict(f, r-``{x})))"  
+           f \<in> r -`` {a} \<rightarrow> range(f) &
+           (\<forall>x \<in> r-``{a}. f`x = H(x, restrict(f, r-``{x})))"  
 apply (rule iffI) 
  apply (simp add: is_recfun_type apply_recfun Ball_def vimage_singleton_iff, 
         clarify)  
@@ -174,8 +173,8 @@
    "[| wellfounded(M,r); trans(r); M(r);
        \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g));  M(Y);
        \<forall>b[M]. 
-	   b \<in> Y <->
-	   (\<exists>x[M]. <x,a1> \<in> r &
+           b \<in> Y <->
+           (\<exists>x[M]. <x,a1> \<in> r &
             (\<exists>y[M]. b = \<langle>x,y\<rangle> & (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g))));
           \<langle>x,a1\<rangle> \<in> r; is_recfun(r,x,H,f); M(f) |]
        ==> restrict(Y, r -`` {x}) = f"
@@ -187,7 +186,7 @@
 apply (rule iffI)
  apply (frule_tac y="<y,z>" in transM, assumption)
  apply (clarsimp simp add: vimage_singleton_iff is_recfun_type [THEN apply_iff]
-			   apply_recfun is_recfun_cut) 
+                           apply_recfun is_recfun_cut) 
 txt{*Opposite inclusion: something in f, show in Y*}
 apply (frule_tac y="<y,z>" in transM, assumption)  
 apply (simp add: vimage_singleton_iff) 
@@ -276,7 +275,7 @@
    "M_is_recfun(M,MH,r,a,f) == 
      \<forall>z[M]. z \<in> f <-> 
             (\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M]. 
-	       pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) &
+               pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) &
                pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
                xa \<in> r & MH(x, f_r_sx, y))"
 
--- a/src/ZF/Constructible/Wellorderings.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Constructible/Wellorderings.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Constructible/Wellorderings.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 *)
 
@@ -23,30 +22,30 @@
 definition
   transitive_rel :: "[i=>o,i,i]=>o" where
     "transitive_rel(M,A,r) == 
-	\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> (\<forall>z[M]. z\<in>A --> 
+        \<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> (\<forall>z[M]. z\<in>A --> 
                           <x,y>\<in>r --> <y,z>\<in>r --> <x,z>\<in>r))"
 
 definition
   linear_rel :: "[i=>o,i,i]=>o" where
     "linear_rel(M,A,r) == 
-	\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> <x,y>\<in>r | x=y | <y,x>\<in>r)"
+        \<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> <x,y>\<in>r | x=y | <y,x>\<in>r)"
 
 definition
   wellfounded :: "[i=>o,i]=>o" where
     --{*EVERY non-empty set has an @{text r}-minimal element*}
     "wellfounded(M,r) == 
-	\<forall>x[M]. x\<noteq>0 --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
+        \<forall>x[M]. x\<noteq>0 --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
 definition
   wellfounded_on :: "[i=>o,i,i]=>o" where
     --{*every non-empty SUBSET OF @{text A} has an @{text r}-minimal element*}
     "wellfounded_on(M,A,r) == 
-	\<forall>x[M]. x\<noteq>0 --> x\<subseteq>A --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
+        \<forall>x[M]. x\<noteq>0 --> x\<subseteq>A --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
 
 definition
   wellordered :: "[i=>o,i,i]=>o" where
     --{*linear and wellfounded on @{text A}*}
     "wellordered(M,A,r) == 
-	transitive_rel(M,A,r) & linear_rel(M,A,r) & wellfounded_on(M,A,r)"
+        transitive_rel(M,A,r) & linear_rel(M,A,r) & wellfounded_on(M,A,r)"
 
 
 subsubsection {*Trivial absoluteness proofs*}
@@ -221,7 +220,7 @@
     "[| wellordered(M,A,r);  B<=A |] ==> wellordered(M,B,r)"
 apply (unfold wellordered_def)
 apply (blast intro: linear_rel_subset transitive_rel_subset 
-		    wellfounded_on_subset)
+                    wellfounded_on_subset)
 done
 
 lemma (in M_basic) wellfounded_on_asym:
--- a/src/ZF/Finite.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Finite.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Finite.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
 
@@ -13,8 +12,8 @@
 (*The natural numbers as a datatype*)
 rep_datatype
   elimination    natE
-  induction	 nat_induct
-  case_eqns	 nat_case_0 nat_case_succ
+  induction      nat_induct
+  case_eqns      nat_case_0 nat_case_succ
   recursor_eqns  recursor_0 recursor_succ
 
 
--- a/src/ZF/Induct/Comb.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Induct/Comb.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Induct/Comb.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson
     Copyright   1994  University of Cambridge
 *)
@@ -39,7 +38,7 @@
   "p ---> q" == "<p,q> \<in> contract^*"
 
 syntax (xsymbols)
-  "comb.app"    :: "[i, i] => i"    	     (infixl "\<bullet>" 90)
+  "comb.app"    :: "[i, i] => i"             (infixl "\<bullet>" 90)
 
 inductive
   domains "contract" \<subseteq> "comb \<times> comb"
--- a/src/ZF/Induct/FoldSet.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Induct/FoldSet.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/Induct/FoldSet.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
 
-
 A "fold" functional for finite sets.  For n non-negative we have
 fold f e {x1,...,xn} = f x1 (... (f xn e)) where f is at
 least left-commutative.  
@@ -17,7 +15,7 @@
   intros
     emptyI: "e\<in>B ==> <0, e>\<in>fold_set(A, B, f,e)"
     consI:  "[| x\<in>A; x \<notin>C;  <C,y> : fold_set(A, B,f,e); f(x,y):B |]
-		==>  <cons(x,C), f(x,y)>\<in>fold_set(A, B, f, e)"
+                ==>  <cons(x,C), f(x,y)>\<in>fold_set(A, B, f, e)"
   type_intros Fin.intros
   
 definition
--- a/src/ZF/Induct/Multiset.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Induct/Multiset.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Induct/Multiset.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
 
 A definitional theory of multisets,
@@ -50,7 +49,7 @@
 definition
   mdiff  :: "[i, i] => i" (infixl "-#" 65)  where
   "M -# N ==  normalize(\<lambda>x \<in> mset_of(M).
-			if x \<in> mset_of(N) then M`x #- N`x else M`x)"
+                        if x \<in> mset_of(N) then M`x #- N`x else M`x)"
 
 definition
   (* set of elements of a multiset *)
@@ -94,7 +93,7 @@
 
 definition
   multirel :: "[i, i] => i"  where
-  "multirel(A, r) == multirel1(A, r)^+" 		
+  "multirel(A, r) == multirel1(A, r)^+"                 
 
   (* ordinal multiset orderings *)
 
@@ -380,7 +379,7 @@
 
 lemma setsum_mcount_Int:
      "Finite(A) ==> setsum(%a. $# mcount(N, a), A Int mset_of(N))
-		  = setsum(%a. $# mcount(N, a), A)"
+                  = setsum(%a. $# mcount(N, a), A)"
 apply (induct rule: Finite_induct)
  apply auto
 apply (subgoal_tac "Finite (B Int mset_of (N))")
--- a/src/ZF/InfDatatype.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/InfDatatype.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
-(*  Title:      ZF/InfDatatype.ML
-    ID:         $Id$
+(*  Title:      ZF/InfDatatype.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
-
 *)
 
 header{*Infinite-Branching Datatype Definitions*}
@@ -64,7 +62,7 @@
      "InfCard(K) ==> K -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))"
 apply (frule InfCard_is_Card [THEN Card_is_Ord])
 apply (blast del: subsetI
-	     intro: fun_Vcsucc Ord_cardinal_le i_subset_Vfrom 
+             intro: fun_Vcsucc Ord_cardinal_le i_subset_Vfrom 
                    lt_csucc [THEN leI, THEN le_imp_subset, THEN subset_trans]) 
 done
 
--- a/src/ZF/IntDiv_ZF.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/IntDiv_ZF.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      ZF/IntDiv.thy
-    ID:         $Id$
+(*  Title:      ZF/IntDiv_ZF.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1999  University of Cambridge
 
@@ -8,25 +7,24 @@
     fun posDivAlg (a,b) =
       if a<b then (0,a)
       else let val (q,r) = posDivAlg(a, 2*b)
-	       in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
-	   end
+               in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
+           end
 
     fun negDivAlg (a,b) =
       if 0<=a+b then (~1,a+b)
       else let val (q,r) = negDivAlg(a, 2*b)
-	       in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
-	   end;
+               in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
+           end;
 
     fun negateSnd (q,r:int) = (q,~r);
 
     fun divAlg (a,b) = if 0<=a then 
-			  if b>0 then posDivAlg (a,b) 
-			   else if a=0 then (0,0)
-				else negateSnd (negDivAlg (~a,~b))
-		       else 
-			  if 0<b then negDivAlg (a,b)
-			  else        negateSnd (posDivAlg (~a,~b));
-
+                          if b>0 then posDivAlg (a,b) 
+                           else if a=0 then (0,0)
+                                else negateSnd (negDivAlg (~a,~b))
+                       else 
+                          if 0<b then negDivAlg (a,b)
+                          else        negateSnd (posDivAlg (~a,~b));
 *)
 
 header{*The Division Operators Div and Mod*}
@@ -53,8 +51,8 @@
 (*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*)
     "posDivAlg(ab) ==
        wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)),
-	     ab,
-	     %<a,b> f. if (a$<b | b$<=#0) then <#0,a>
+             ab,
+             %<a,b> f. if (a$<b | b$<=#0) then <#0,a>
                        else adjust(b, f ` <a,#2$*b>))"
 
 
@@ -64,8 +62,8 @@
 (*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*)
     "negDivAlg(ab) ==
        wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)),
-	     ab,
-	     %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b>
+             ab,
+             %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b>
                        else adjust(b, f ` <a,#2$*b>))"
 
 (*for the general case b\<noteq>0*)
--- a/src/ZF/Int_ZF.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Int_ZF.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/Int.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
-
 *)
 
 header{*The Integers as Equivalence Classes Over Pairs of Natural Numbers*}
@@ -55,7 +53,7 @@
   --{*could be replaced by an absolute value function from int to int?*}
     "zmagnitude(z) ==
      THE m. m\<in>nat & ((~ znegative(z) & z = $# m) |
-		       (znegative(z) & $- z = $# m))"
+                       (znegative(z) & $- z = $# m))"
 
 definition
   raw_zmult   ::      "[i,i]=>i"  where
--- a/src/ZF/List_ZF.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/List_ZF.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
-(*  Title:      ZF/List
-    ID:         $Id$
+(*  Title:      ZF/List_ZF.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
-
 *)
 
 header{*Lists in Zermelo-Fraenkel Set Theory*}
@@ -96,14 +94,14 @@
 (* Function `take' returns the first n elements of a list *)
   take     :: "[i,i]=>i"  where
   "take(n, as) == list_rec(lam n:nat. [],
-		%a l r. lam n:nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
+                %a l r. lam n:nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
 
 definition
   nth :: "[i, i]=>i"  where
   --{*returns the (n+1)th element of a list, or 0 if the
    list is too short.*}
   "nth(n, as) == list_rec(lam n:nat. 0,
- 		          %a l r. lam n:nat. nat_case(a, %m. r`m, n), as) ` n"
+                          %a l r. lam n:nat. nat_case(a, %m. r`m, n), as) ` n"
 
 definition
   list_update :: "[i, i, i]=>i"  where
--- a/src/ZF/Nat_ZF.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Nat_ZF.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
-(*  Title:      ZF/Nat.thy
-    ID:         $Id$
+(*  Title:      ZF/Nat_ZF.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
-
 *)
 
 header{*The Natural numbers As a Least Fixed Point*}
@@ -156,7 +154,7 @@
 lemmas complete_induct = Ord_induct [OF _ Ord_nat, case_names less, consumes 1]
 
 lemmas complete_induct_rule =  
-	complete_induct [rule_format, case_names less, consumes 1]
+        complete_induct [rule_format, case_names less, consumes 1]
 
 
 lemma nat_induct_from_lemma [rule_format]: 
--- a/src/ZF/Order.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Order.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Order.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
 
@@ -17,15 +16,15 @@
   counterparts. *}
 
 definition
-  part_ord :: "[i,i]=>o"          	(*Strict partial ordering*)  where
+  part_ord :: "[i,i]=>o"                (*Strict partial ordering*)  where
    "part_ord(A,r) == irrefl(A,r) & trans[A](r)"
 
 definition
-  linear   :: "[i,i]=>o"          	(*Strict total ordering*)  where
+  linear   :: "[i,i]=>o"                (*Strict total ordering*)  where
    "linear(A,r) == (ALL x:A. ALL y:A. <x,y>:r | x=y | <y,x>:r)"
 
 definition
-  tot_ord  :: "[i,i]=>o"          	(*Strict total ordering*)  where
+  tot_ord  :: "[i,i]=>o"                (*Strict total ordering*)  where
    "tot_ord(A,r) == part_ord(A,r) & linear(A,r)"
 
 definition
@@ -41,25 +40,25 @@
   "Partial_order(r) \<equiv> partial_order_on(field(r), r)"
 
 definition
-  well_ord :: "[i,i]=>o"          	(*Well-ordering*)  where
+  well_ord :: "[i,i]=>o"                (*Well-ordering*)  where
    "well_ord(A,r) == tot_ord(A,r) & wf[A](r)"
 
 definition
-  mono_map :: "[i,i,i,i]=>i"      	(*Order-preserving maps*)  where
+  mono_map :: "[i,i,i,i]=>i"            (*Order-preserving maps*)  where
    "mono_map(A,r,B,s) ==
-	      {f: A->B. ALL x:A. ALL y:A. <x,y>:r --> <f`x,f`y>:s}"
+              {f: A->B. ALL x:A. ALL y:A. <x,y>:r --> <f`x,f`y>:s}"
 
 definition
-  ord_iso  :: "[i,i,i,i]=>i"		(*Order isomorphisms*)  where
+  ord_iso  :: "[i,i,i,i]=>i"            (*Order isomorphisms*)  where
    "ord_iso(A,r,B,s) ==
-	      {f: bij(A,B). ALL x:A. ALL y:A. <x,y>:r <-> <f`x,f`y>:s}"
+              {f: bij(A,B). ALL x:A. ALL y:A. <x,y>:r <-> <f`x,f`y>:s}"
 
 definition
-  pred     :: "[i,i,i]=>i"		(*Set of predecessors*)  where
+  pred     :: "[i,i,i]=>i"              (*Set of predecessors*)  where
    "pred(A,x,r) == {y:A. <y,x>:r}"
 
 definition
-  ord_iso_map :: "[i,i,i,i]=>i"      	(*Construction for linearity theorem*)  where
+  ord_iso_map :: "[i,i,i,i]=>i"         (*Construction for linearity theorem*)  where
    "ord_iso_map(A,r,B,s) ==
      \<Union>x\<in>A. \<Union>y\<in>B. \<Union>f \<in> ord_iso(pred(A,x,r), r, pred(B,y,s), s). {<x,y>}"
 
@@ -537,7 +536,7 @@
    prefer 4
    apply (rule converse_ord_iso_map [THEN subst])
    apply (simp add: ord_iso_map_mono_map
-		    ord_iso_map_subset [THEN converse_converse])
+                    ord_iso_map_subset [THEN converse_converse])
 apply (blast intro!: domain_ord_iso_map range_ord_iso_map
              intro: well_ord_subset ord_iso_map_mono_map)+
 done
@@ -581,8 +580,8 @@
 apply (erule Diff_eq_0_iff [THEN iffD1])
 (*The other case: the domain equals an initial segment*)
 apply (blast del: domainI subsetI
-	     elim!: predE
-	     intro!: domain_ord_iso_map_subset
+             elim!: predE
+             intro!: domain_ord_iso_map_subset
              intro: subsetI)+
 done
 
--- a/src/ZF/OrderType.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/OrderType.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/OrderType.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
-
 *)
 
 header{*Order Types and Ordinal Arithmetic*}
@@ -161,7 +159,7 @@
 apply (rule OrdI [OF _ Ord_is_Transset])
 apply (unfold pred_def Transset_def)
 apply (blast intro: trans_onD
-	     dest!: ordermap_unfold [THEN equalityD1])+ 
+             dest!: ordermap_unfold [THEN equalityD1])+ 
 done
 
 lemma Ord_ordertype: 
@@ -172,7 +170,7 @@
 prefer 2 apply (blast intro: Ord_ordermap)
 apply (unfold Transset_def well_ord_def)
 apply (blast intro: trans_onD
-	     dest!: ordermap_unfold [THEN equalityD1])
+             dest!: ordermap_unfold [THEN equalityD1])
 done
 
 
--- a/src/ZF/Perm.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Perm.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      ZF/perm
-    ID:         $Id$
+(*  Title:      ZF/Perm.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1991  University of Cambridge
 
@@ -485,7 +484,7 @@
      ==> (f Un g) : surj(A Un C, B Un D)"
 apply (simp add: surj_def fun_disjoint_Un) 
 apply (blast dest!: domain_of_fun 
-	     intro!: fun_disjoint_apply1 fun_disjoint_apply2)
+             intro!: fun_disjoint_apply1 fun_disjoint_apply2)
 done
 
 (*A simple, high-level proof; the version for injections follows from it,
--- a/src/ZF/QUniv.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/QUniv.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/QUniv.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
-
 *)
 
 header{*A Small Universe for Lazy Recursive Types*}
@@ -11,15 +9,15 @@
 
 (*Disjoint sums as a datatype*)
 rep_datatype 
-  elimination	sumE
-  induction	TrueI
-  case_eqns	case_Inl case_Inr
+  elimination   sumE
+  induction     TrueI
+  case_eqns     case_Inl case_Inr
 
 (*Variant disjoint sums as a datatype*)
 rep_datatype 
-  elimination	qsumE
-  induction	TrueI
-  case_eqns	qcase_QInl qcase_QInr
+  elimination   qsumE
+  induction     TrueI
+  case_eqns     qcase_QInl qcase_QInr
 
 definition
   quniv :: "i => i"  where
--- a/src/ZF/Resid/Redex.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Resid/Redex.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,5 @@
 (*  Title:      ZF/Resid/Redex.thy
-    ID:         $Id$
-    Author:     Ole Rasmussen
-    Copyright   1995  University of Cambridge
-    Logic Image: ZF
+    Author:     Ole Rasmussen, University of Cambridge
 *)
 
 theory Redex imports Main begin
@@ -38,12 +35,12 @@
 
   "union_aux(Fun(u)) =
      (\<lambda>t \<in> redexes. redexes_case(%j. 0, %y. Fun(union_aux(u)`y),
-	 			  %b y z. 0, t))"
+                                  %b y z. 0, t))"
 
   "union_aux(App(b,f,a)) =
      (\<lambda>t \<in> redexes.
         redexes_case(%j. 0, %y. 0,
-		     %c z u. App(b or c, union_aux(f)`z, union_aux(a)`u), t))"
+                     %c z u. App(b or c, union_aux(f)`z, union_aux(a)`u), t))"
 
 definition
   union  (infixl "un" 70) where
--- a/src/ZF/Resid/Substitution.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Resid/Substitution.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,5 @@
 (*  Title:      ZF/Resid/Substitution.thy
-    ID:         $Id$
-    Author:     Ole Rasmussen
-    Copyright   1995  University of Cambridge
-    Logic Image: ZF
+    Author:     Ole Rasmussen, University of Cambridge
 *)
 
 theory Substitution imports Redex begin
@@ -33,7 +30,7 @@
 primrec
   "subst_aux(Var(i)) =
      (\<lambda>r \<in> redexes. \<lambda>k \<in> nat. if k<i then Var(i #- 1)
-				else if k=i then r else Var(i))"
+                                else if k=i then r else Var(i))"
   "subst_aux(Fun(t)) =
      (\<lambda>r \<in> redexes. \<lambda>k \<in> nat. Fun(subst_aux(t) ` lift(r) ` succ(k)))"
 
@@ -41,7 +38,7 @@
      (\<lambda>r \<in> redexes. \<lambda>k \<in> nat. App(b, subst_aux(f)`r`k, subst_aux(a)`r`k))"
 
 definition
-  subst_rec     :: "[i,i,i]=> i"	(**NOTE THE ARGUMENT ORDER BELOW**)  where
+  subst_rec     :: "[i,i,i]=> i"        (**NOTE THE ARGUMENT ORDER BELOW**)  where
     "subst_rec(u,r,k) == subst_aux(r)`u`k"
 
 abbreviation
@@ -216,8 +213,8 @@
 lemma subst_rec_subst_rec [rule_format]:
      "v \<in> redexes ==>
         \<forall>m \<in> nat. \<forall>n \<in> nat. \<forall>u \<in> redexes. \<forall>w \<in> redexes. m\<le>n -->
-	  subst_rec(subst_rec(w,u,n),subst_rec(lift_rec(w,m),v,succ(n)),m) =
-	  subst_rec(w,subst_rec(u,v,m),n)"
+          subst_rec(subst_rec(w,u,n),subst_rec(lift_rec(w,m),v,succ(n)),m) =
+          subst_rec(w,subst_rec(u,v,m),n)"
 apply (erule redexes.induct)
 apply (simp_all add: lift_lift [symmetric] lift_rec_subst_rec_lt, safe)
 apply (rename_tac n' u w)
--- a/src/ZF/Sum.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Sum.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/sum.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
-
 *)
 
 header{*Disjoint Sums*}
@@ -161,8 +159,7 @@
 by auto 
 
 lemma case_case: "z: A+B ==>    
-        
-	case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) =  
+        case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) =  
         case(%x. c(c'(x)), %y. d(d'(y)), z)"
 by auto
 
--- a/src/ZF/Tools/datatype_package.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Tools/datatype_package.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -2,7 +2,7 @@
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
 
-Datatype/Codatatype Definitions
+Datatype/Codatatype Definitions.
 
 The functor will be instantiated for normal sums/products (datatype defs)
                          and non-standard sums/products (codatatype defs)
@@ -55,9 +55,9 @@
           (fn t => "Datatype set not previously declared as constant: " ^
                    Syntax.string_of_term_global @{theory IFOL} t);
         val rec_names = (*nat doesn't have to be added*)
-	    @{const_name nat} :: map (#1 o dest_Const) rec_hds
-	val u = if co then @{const QUniv.quniv} else @{const Univ.univ}
-	val cs = (fold o fold) (fn (_, _, _, prems) => prems |> (fold o fold_aterms)
+            @{const_name nat} :: map (#1 o dest_Const) rec_hds
+        val u = if co then @{const QUniv.quniv} else @{const Univ.univ}
+        val cs = (fold o fold) (fn (_, _, _, prems) => prems |> (fold o fold_aterms)
           (fn t as Const (a, _) => if a mem_string rec_names then I else insert (op =) t
             | _ => I)) con_ty_lists [];
     in  u $ Ind_Syntax.union_params (hd rec_tms, cs)  end;
@@ -292,7 +292,7 @@
          rtac case_trans 1,
          REPEAT (resolve_tac [refl, split_trans, Su.case_inl RS trans, Su.case_inr RS trans] 1)]);
 
-  val free_iffs = map standard (con_defs RL [@{thm def_swap_iff}]);
+  val free_iffs = map Drule.standard (con_defs RL [@{thm def_swap_iff}]);
 
   val case_eqns = map prove_case_eqn (flat con_ty_lists ~~ case_args ~~ tl con_defs);
 
@@ -338,7 +338,7 @@
   val constructors =
       map (head_of o #1 o Logic.dest_equals o #prop o rep_thm) (tl con_defs);
 
-  val free_SEs = map standard (Ind_Syntax.mk_free_SEs free_iffs);
+  val free_SEs = map Drule.standard (Ind_Syntax.mk_free_SEs free_iffs);
 
   val {intrs, elim, induct, mutual_induct, ...} = ind_result
 
--- a/src/ZF/Tools/inductive_package.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Tools/inductive_package.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -193,9 +193,9 @@
         [rtac (@{thm Collect_subset} RS @{thm bnd_monoI}) 1,
          REPEAT (ares_tac (@{thms basic_monos} @ monos) 1)]);
 
-  val dom_subset = standard (big_rec_def RS Fp.subs);
+  val dom_subset = Drule.standard (big_rec_def RS Fp.subs);
 
-  val unfold = standard ([big_rec_def, bnd_mono] MRS Fp.Tarski);
+  val unfold = Drule.standard ([big_rec_def, bnd_mono] MRS Fp.Tarski);
 
   (********)
   val dummy = writeln "  Proving the introduction rules...";
@@ -205,7 +205,7 @@
   val Part_trans =
       case rec_names of
            [_] => asm_rl
-         | _   => standard (@{thm Part_subset} RS @{thm subset_trans});
+         | _   => Drule.standard (@{thm Part_subset} RS @{thm subset_trans});
 
   (*To type-check recursive occurrences of the inductive sets, possibly
     enclosed in some monotonic operator M.*)
@@ -503,7 +503,7 @@
      val Const (@{const_name Trueprop}, _) $ (pred_var $ _) = concl_of induct0
 
      val induct = CP.split_rule_var(pred_var, elem_type-->FOLogic.oT, induct0)
-                  |> standard
+                  |> Drule.standard
      and mutual_induct = CP.remove_split mutual_induct_fsplit
 
      val ([induct', mutual_induct'], thy') =
@@ -514,7 +514,7 @@
     in ((thy', induct'), mutual_induct')
     end;  (*of induction_rules*)
 
-  val raw_induct = standard ([big_rec_def, bnd_mono] MRS Fp.induct)
+  val raw_induct = Drule.standard ([big_rec_def, bnd_mono] MRS Fp.induct)
 
   val ((thy2, induct), mutual_induct) =
     if not coind then induction_rules raw_induct thy1
--- a/src/ZF/Tools/numeral_syntax.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Tools/numeral_syntax.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Tools/numeral_syntax.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 
 Concrete syntax for generic numerals.  Assumes consts and syntax of
@@ -37,12 +36,12 @@
 
 fun mk_bin i =
   let fun bin_of_int 0  = []
-  	    | bin_of_int ~1 = [~1]
-    	| bin_of_int n  = (n mod 2) :: bin_of_int (n div 2);
+        | bin_of_int ~1 = [~1]
+        | bin_of_int n  = (n mod 2) :: bin_of_int (n div 2);
 
       fun term_of [] = @{const Pls}
-	    | term_of [~1] = @{const Min}
-	    | term_of (b :: bs) = @{const Bit} $ term_of bs $ mk_bit b;
+            | term_of [~1] = @{const Min}
+            | term_of (b :: bs) = @{const Bit} $ term_of bs $ mk_bit b;
   in
     term_of (bin_of_int i)
   end;
@@ -72,9 +71,9 @@
   let
     val rev_digs = bin_of t;
     val (sign, zs) =
-	(case rev rev_digs of
-	     ~1 :: bs => ("-", prefix_len (equal 1) bs)
-	   | bs =>       ("",  prefix_len (equal 0) bs));
+        (case rev rev_digs of
+             ~1 :: bs => ("-", prefix_len (equal 1) bs)
+           | bs =>       ("",  prefix_len (equal 0) bs));
     val num = string_of_int (abs (integ_of rev_digs));
   in
     "#" ^ sign ^ implode (replicate zs "0") ^ num
--- a/src/ZF/UNITY/AllocBase.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/AllocBase.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/UNITY/AllocBase.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2001  University of Cambridge
-
 *)
 
 header{*Common declarations for Chandy and Charpentier's Allocator*}
@@ -355,7 +353,7 @@
 
 lemma var_not_Finite: "~Finite(var)"
 apply (insert nat_not_Finite) 
-apply (blast intro: lepoll_Finite [OF 	nat_lepoll_var]) 
+apply (blast intro: lepoll_Finite [OF nat_lepoll_var]) 
 done
 
 lemma not_Finite_imp_exist: "~Finite(A) ==> \<exists>x. x\<in>A"
--- a/src/ZF/UNITY/AllocImpl.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/AllocImpl.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,8 @@
-(*  Title: ZF/UNITY/AllocImpl.thy
-    ID:    $Id$
+(*  Title:      ZF/UNITY/AllocImpl.thy
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2002  University of Cambridge
 
-Single-client allocator implementation
+Single-client allocator implementation.
 Charpentier and Chandy, section 7 (page 17).
 *)
 
@@ -27,27 +26,27 @@
 definition
   "alloc_giv_act ==
        {<s, t> \<in> state*state.
-	\<exists>k. k = length(s`giv) &
+        \<exists>k. k = length(s`giv) &
             t = s(giv := s`giv @ [nth(k, s`ask)],
-		  available_tok := s`available_tok #- nth(k, s`ask)) &
-	    k < length(s`ask) & nth(k, s`ask) le s`available_tok}"
+                  available_tok := s`available_tok #- nth(k, s`ask)) &
+            k < length(s`ask) & nth(k, s`ask) le s`available_tok}"
 
 definition
   "alloc_rel_act ==
        {<s, t> \<in> state*state.
         t = s(available_tok := s`available_tok #+ nth(s`NbR, s`rel),
-	      NbR := succ(s`NbR)) &
-  	s`NbR < length(s`rel)}"
+              NbR := succ(s`NbR)) &
+        s`NbR < length(s`rel)}"
 
 definition
   (*The initial condition s`giv=[] is missing from the
     original definition: S. O. Ehmety *)
   "alloc_prog ==
        mk_program({s:state. s`available_tok=NbT & s`NbR=0 & s`giv=Nil},
-		  {alloc_giv_act, alloc_rel_act},
-		  \<Union>G \<in> preserves(lift(available_tok)) \<inter>
-		        preserves(lift(NbR)) \<inter>
-		        preserves(lift(giv)). Acts(G))"
+                  {alloc_giv_act, alloc_rel_act},
+                  \<Union>G \<in> preserves(lift(available_tok)) \<inter>
+                        preserves(lift(NbR)) \<inter>
+                        preserves(lift(giv)). Acts(G))"
 
 
 lemma available_tok_value_type [simp,TC]: "s\<in>state ==> s`available_tok \<in> nat"
@@ -246,10 +245,10 @@
 
 lemma alloc_prog_NbR_LeadsTo_lemma:
      "[| G \<in> program; alloc_prog ok G;
-	 alloc_prog \<squnion> G \<in> Incr(lift(rel)); k\<in>nat |]
+         alloc_prog \<squnion> G \<in> Incr(lift(rel)); k\<in>nat |]
       ==> alloc_prog \<squnion> G \<in>
-	    {s\<in>state. k \<le> length(s`rel)} \<inter> {s\<in>state. succ(s`NbR) = k}
-	    LeadsTo {s\<in>state. k \<le> s`NbR}"
+            {s\<in>state. k \<le> length(s`rel)} \<inter> {s\<in>state. succ(s`NbR) = k}
+            LeadsTo {s\<in>state. k \<le> s`NbR}"
 apply (subgoal_tac "alloc_prog \<squnion> G \<in> Stable ({s\<in>state. k \<le> length (s`rel)})")
 apply (drule_tac [2] a = k and g1 = length in imp_Increasing_comp [THEN Increasing_imp_Stable])
 apply (rule_tac [2] mono_length)
@@ -268,9 +267,9 @@
     "[| G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel));
         k\<in>nat; n \<in> nat; n < k |]
       ==> alloc_prog \<squnion> G \<in>
-	    {s\<in>state . k \<le> length(s ` rel)} \<inter> {s\<in>state . s ` NbR = n}
-	       LeadsTo {x \<in> state. k \<le> length(x`rel)} \<inter>
-		 (\<Union>m \<in> greater_than(n). {x \<in> state. x ` NbR=m})"
+            {s\<in>state . k \<le> length(s ` rel)} \<inter> {s\<in>state . s ` NbR = n}
+               LeadsTo {x \<in> state. k \<le> length(x`rel)} \<inter>
+                 (\<Union>m \<in> greater_than(n). {x \<in> state. x ` NbR=m})"
 apply (unfold greater_than_def)
 apply (rule_tac A' = "{x \<in> state. k \<le> length(x`rel)} \<inter> {x \<in> state. n < x`NbR}"
        in LeadsTo_weaken_R)
@@ -374,10 +373,10 @@
 "[| G \<in> program; alloc_prog ok G;
     alloc_prog \<squnion> G \<in> Incr(lift(ask)); k\<in>nat |]
  ==> alloc_prog \<squnion> G \<in>
-	{s\<in>state. nth(length(s`giv), s`ask) \<le> s`available_tok} \<inter>
-	{s\<in>state.  k < length(s`ask)} \<inter>
-	{s\<in>state. length(s`giv) = k}
-	LeadsTo {s\<in>state. k < length(s`giv)}"
+        {s\<in>state. nth(length(s`giv), s`ask) \<le> s`available_tok} \<inter>
+        {s\<in>state.  k < length(s`ask)} \<inter>
+        {s\<in>state. length(s`giv) = k}
+        LeadsTo {s\<in>state. k < length(s`giv)}"
 apply (subgoal_tac "alloc_prog \<squnion> G \<in> {s\<in>state. nth (length(s`giv), s`ask) \<le> s`available_tok} \<inter> {s\<in>state. k < length(s`ask) } \<inter> {s\<in>state. length(s`giv) = k} LeadsTo {s\<in>state. ~ k <length(s`ask) } Un {s\<in>state. length(s`giv) \<noteq> k}")
 prefer 2 apply (blast intro: alloc_prog_giv_Ensures_lemma [THEN LeadsTo_Basis])
 apply (subgoal_tac "alloc_prog \<squnion> G \<in> Stable ({s\<in>state. k < length(s`ask) }) ")
@@ -482,8 +481,8 @@
 lemma (in alloc_progress) tokens_take_NbR_lemma2:
      "k \<in> tokbag
       ==> alloc_prog \<squnion> G \<in>
-	    {s\<in>state. tokens(s`giv) = k}
-	    LeadsTo {s\<in>state. k \<le> tokens(take(s`NbR, s`rel))}"
+            {s\<in>state. tokens(s`giv) = k}
+            LeadsTo {s\<in>state. k \<le> tokens(take(s`NbR, s`rel))}"
 apply (rule LeadsTo_Trans)
  apply (rule_tac [2] tokens_take_NbR_lemma)
  prefer 2 apply assumption
@@ -495,10 +494,10 @@
 lemma (in alloc_progress) length_giv_disj:
      "[| k \<in> tokbag; n \<in> nat |]
       ==> alloc_prog \<squnion> G \<in>
-	    {s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
-	    LeadsTo
-	      {s\<in>state. (length(s`giv) = n & tokens(s`giv) = k &
-			 k \<le> tokens(take(s`NbR, s`rel))) | n < length(s`giv)}"
+            {s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
+            LeadsTo
+              {s\<in>state. (length(s`giv) = n & tokens(s`giv) = k &
+                         k \<le> tokens(take(s`NbR, s`rel))) | n < length(s`giv)}"
 apply (rule single_LeadsTo_I, safe)
 apply (rule_tac a1 = "s`giv" in Increasing_imp_Stable [THEN PSP_StableI])
 apply (rule alloc_prog_Increasing_giv [THEN guaranteesD])
@@ -515,10 +514,10 @@
 lemma (in alloc_progress) length_giv_disj2:
      "[|k \<in> tokbag; n \<in> nat|]
       ==> alloc_prog \<squnion> G \<in>
-	    {s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
-	    LeadsTo
-	      {s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
-			n < length(s`giv)}"
+            {s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
+            LeadsTo
+              {s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
+                        n < length(s`giv)}"
 apply (rule LeadsTo_weaken_R)
 apply (rule Always_LeadsToD [OF alloc_prog_Always_lemma length_giv_disj], auto)
 done
@@ -528,10 +527,10 @@
 lemma (in alloc_progress) length_giv_disj3:
      "n \<in> nat
       ==> alloc_prog \<squnion> G \<in>
-	    {s\<in>state. length(s`giv) = n}
-	    LeadsTo
-	      {s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
-			n < length(s`giv)}"
+            {s\<in>state. length(s`giv) = n}
+            LeadsTo
+              {s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
+                        n < length(s`giv)}"
 apply (rule LeadsTo_weaken_L)
 apply (rule_tac I = nat in LeadsTo_UN)
 apply (rule_tac k = i in length_giv_disj2)
@@ -566,11 +565,11 @@
 lemma (in alloc_progress) length_ask_giv2:
      "[|k \<in> nat;  n < k|]
       ==> alloc_prog \<squnion> G \<in>
-	    {s\<in>state. length(s`ask) = k & length(s`giv) = n}
-	    LeadsTo
-	      {s\<in>state. (nth(length(s`giv), s`ask) \<le> s`available_tok &
-			 length(s`giv) < length(s`ask) & length(s`giv) = n) |
-			n < length(s`giv)}"
+            {s\<in>state. length(s`ask) = k & length(s`giv) = n}
+            LeadsTo
+              {s\<in>state. (nth(length(s`giv), s`ask) \<le> s`available_tok &
+                         length(s`giv) < length(s`ask) & length(s`giv) = n) |
+                        n < length(s`giv)}"
 apply (rule LeadsTo_weaken_R)
 apply (rule Always_LeadsToD [OF safety length_ask_giv], assumption+, clarify)
 apply (simp add: INT_iff)
@@ -583,8 +582,8 @@
 lemma (in alloc_progress) extend_giv:
      "[| k \<in> nat;  n < k|]
       ==> alloc_prog \<squnion> G \<in>
-	    {s\<in>state. length(s`ask) = k & length(s`giv) = n}
-	    LeadsTo {s\<in>state. n < length(s`giv)}"
+            {s\<in>state. length(s`ask) = k & length(s`giv) = n}
+            LeadsTo {s\<in>state. n < length(s`giv)}"
 apply (rule LeadsTo_Un_duplicate)
 apply (rule LeadsTo_cancel1)
 apply (rule_tac [2] alloc_prog_giv_LeadsTo_lemma)
@@ -625,7 +624,7 @@
      "h \<in> list(tokbag)
       ==> alloc_prog \<squnion> G
             \<in> {s\<in>state. <h, s`ask> \<in> prefix(tokbag)} LeadsTo
-	      {s\<in>state. <h, s`giv> \<in> prefix(tokbag)}"
+              {s\<in>state. <h, s`giv> \<in> prefix(tokbag)}"
 apply (rule single_LeadsTo_I)
  prefer 2 apply simp
 apply (rename_tac s0)
--- a/src/ZF/UNITY/ClientImpl.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/ClientImpl.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,8 @@
 (*  Title:      ZF/UNITY/ClientImpl.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2002  University of Cambridge
 
-Distributed Resource Management System:  Client Implementation
+Distributed Resource Management System:  Client Implementation.
 *)
 
 theory ClientImpl imports AllocBase Guar begin
@@ -38,7 +37,7 @@
       of the action to be ignored **)
 definition
   "client_tok_act == {<s,t> \<in> state*state. t=s |
-		     t = s(tok:=succ(s`tok mod NbT))}"
+                     t = s(tok:=succ(s`tok mod NbT))}"
 
 definition
   "client_ask_act == {<s,t> \<in> state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}"
@@ -46,11 +45,11 @@
 definition
   "client_prog ==
    mk_program({s \<in> state. s`tok \<le> NbT & s`giv = Nil &
-	               s`ask = Nil & s`rel = Nil},
+                       s`ask = Nil & s`rel = Nil},
                     {client_rel_act, client_tok_act, client_ask_act},
                    \<Union>G \<in> preserves(lift(rel)) Int
-			 preserves(lift(ask)) Int
-	                 preserves(lift(tok)).  Acts(G))"
+                         preserves(lift(ask)) Int
+                         preserves(lift(tok)).  Acts(G))"
 
 
 declare type_assumes [simp] default_val_assumes [simp]
--- a/src/ZF/UNITY/Comp.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Comp.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/UNITY/Comp.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -47,7 +46,7 @@
 definition
   localize  :: "[i=>i, i] => i"  where
   "localize(f,F) == mk_program(Init(F), Acts(F),
-		       AllowedActs(F) Int (UN G:preserves(f). Acts(G)))"
+                       AllowedActs(F) Int (UN G:preserves(f). Acts(G)))"
 
   
 (*** component and strict_component relations ***)
--- a/src/ZF/UNITY/Constrains.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Constrains.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,4 +1,4 @@
-(*  ID:         $Id$
+(*  Title:      ZF/UNITY/Constrains.thy
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   2001  University of Cambridge
 *)
@@ -315,7 +315,7 @@
      ==> F \<in> ({s \<in> A. x(s):M}) Co (\<Union>m \<in> M. B(m))"
 apply (unfold Constrains_def, auto)
 apply (rule_tac A1 = "reachable (F) Int A" 
-	in UNITY.elimination [THEN constrains_weaken_L])
+        in UNITY.elimination [THEN constrains_weaken_L])
 apply (auto intro: constrains_weaken_L)
 done
 
@@ -397,7 +397,7 @@
 apply (simp (no_asm) add: Always_eq_includes_reachable)
 apply (rule equalityI, auto) 
 apply (blast intro: invariantI rev_subsetD [OF _ Init_into_reachable] 
-		    rev_subsetD [OF _ invariant_includes_reachable]  
+                    rev_subsetD [OF _ invariant_includes_reachable]  
              dest: invariant_type [THEN subsetD])+
 done
 
--- a/src/ZF/UNITY/Distributor.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Distributor.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,10 +1,9 @@
-(*  Title: ZF/UNITY/Distributor
-    ID:         $Id$
+(*  Title:      ZF/UNITY/Distributor.thy
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2002  University of Cambridge
 
 A multiple-client allocator from a single-client allocator:
-Distributor specification
+Distributor specification.
 *)
 
 theory Distributor imports AllocBase Follows  Guar GenPrefix begin
@@ -24,7 +23,7 @@
           lift(Out(n))
               Fols
           (%s. sublist(s`In, {k \<in> nat. k<length(s`iIn) & nth(k, s`iIn) = n}))
-	  Wrt prefix(A)/list(A))"
+          Wrt prefix(A)/list(A))"
 
 definition
   distr_allowed_acts :: "[i=>i]=>i"  where
@@ -80,7 +79,7 @@
 
 lemma (in distr) D_ok_iff:
      "G \<in> program ==>
-	D ok G <-> ((\<forall>n \<in> nat. G \<in> preserves(lift(Out(n)))) & D \<in> Allowed(G))"
+        D ok G <-> ((\<forall>n \<in> nat. G \<in> preserves(lift(Out(n)))) & D \<in> Allowed(G))"
 apply (cut_tac distr_spec)
 apply (auto simp add: INT_iff distr_spec_def distr_allowed_acts_def
                       Allowed_def ok_iff_Allowed)
@@ -159,9 +158,9 @@
 apply (auto simp add: distr_spec_def distr_follows_def)
 apply (drule guaranteesD, assumption)
 apply (simp_all cong add: Follows_cong
-		add: refl_prefix
-		   mono_bag_of [THEN subset_Follows_comp, THEN subsetD, 
-				unfolded metacomp_def])
+                add: refl_prefix
+                   mono_bag_of [THEN subset_Follows_comp, THEN subsetD, 
+                                unfolded metacomp_def])
 done
 
 end
--- a/src/ZF/UNITY/FP.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/FP.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/UNITY/FP.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   2001  University of Cambridge
 
--- a/src/ZF/UNITY/Follows.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Follows.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      ZF/UNITY/Follows
-    ID:         $Id \<in> Follows.thy,v 1.1 2003/05/28 16:13:42 paulson Exp $
+(*  Title:      ZF/UNITY/Follows.thy
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2002  University of Cambridge
 
@@ -279,8 +278,8 @@
 apply (unfold Follows_def Increasing_def Stable_def)
 apply (simp add: INT_iff, auto)
 apply (rule_tac [3] C = "{s \<in> state. f(s)=g(s)}"
-	and A = "{s \<in> state. <k, h (s)> \<in> r}"
-	and A' = "{s \<in> state. <k, f(s)> \<in> r}" in Always_LeadsTo_weaken)
+        and A = "{s \<in> state. <k, h (s)> \<in> r}"
+        and A' = "{s \<in> state. <k, f(s)> \<in> r}" in Always_LeadsTo_weaken)
 apply (erule_tac A = "{s \<in> state. <k,f(s) > \<in> r}"
            and A' = "{s \<in> state. <k,f(s) > \<in> r}" in Always_Constrains_weaken)
 apply auto
--- a/src/ZF/UNITY/GenPrefix.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/GenPrefix.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,10 +1,10 @@
-(*  ID:         $Id$
+(*  Title:      ZF/UNITY/GenPrefix.thy
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2001  University of Cambridge
 
-   <xs,ys>:gen_prefix(r)
-     if ys = xs' @ zs where length(xs) = length(xs')
-     and corresponding elements of xs, xs' are pairwise related by r
+<xs,ys>:gen_prefix(r)
+  if ys = xs' @ zs where length(xs) = length(xs')
+  and corresponding elements of xs, xs' are pairwise related by r
 
 Based on Lex/Prefix
 *)
@@ -31,10 +31,10 @@
     Nil:     "<[],[]>:gen_prefix(A, r)"
 
     prepend: "[| <xs,ys>:gen_prefix(A, r);  <x,y>:r; x:A; y:A |]
-	      ==> <Cons(x,xs), Cons(y,ys)>: gen_prefix(A, r)"
+              ==> <Cons(x,xs), Cons(y,ys)>: gen_prefix(A, r)"
 
     append:  "[| <xs,ys>:gen_prefix(A, r); zs:list(A) |]
-	      ==> <xs, ys@zs>:gen_prefix(A, r)"
+              ==> <xs, ys@zs>:gen_prefix(A, r)"
     type_intros app_type list.Nil list.Cons
 
 definition
--- a/src/ZF/UNITY/Guar.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Guar.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/UNITY/Guar.thy
-    ID:         $Id \<in> Guar.thy,v 1.3 2001/11/15 14:51:43 ehmety Exp $
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   2001  University of Cambridge
 
--- a/src/ZF/UNITY/Increasing.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Increasing.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      ZF/UNITY/Increasing
-    ID:         $Id \<in> Increasing.thy,v 1.2 2003/06/02 09:17:52 paulson Exp $
+(*  Title:      ZF/UNITY/Increasing.thy
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2001  University of Cambridge
 
--- a/src/ZF/UNITY/Merge.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Merge.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,10 +1,9 @@
-(*  Title: ZF/UNITY/Merge
-    ID:         $Id$
+(*  Title:      ZF/UNITY/Merge.thy
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2002  University of Cambridge
 
-A multiple-client allocator from a single-client allocator:
-Merge specification
+A multiple-client allocator from a single-client allocator: Merge
+specification.
 *)
 
 theory Merge imports AllocBase Follows  Guar GenPrefix begin
@@ -37,7 +36,7 @@
   merge_follows :: "[i, i=>i, i, i] =>i"  where
     "merge_follows(A, In, Out, iOut) ==
      (\<Inter>n \<in> Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A))
-		   guarantees
+                   guarantees
      (\<Inter>n \<in> Nclients. 
         (%s. sublist(s`Out, {k \<in> nat. k < length(s`iOut) &
                       nth(k, s`iOut) = n})) Fols lift(In(n))
@@ -116,14 +115,14 @@
 lemma (in merge) M_ok_iff: 
      "G \<in> program ==>  
        M ok G <-> (G \<in> preserves(lift(Out)) &   
- 	           G \<in> preserves(lift(iOut)) & M \<in> Allowed(G))"
+                   G \<in> preserves(lift(iOut)) & M \<in> Allowed(G))"
 apply (cut_tac merge_spec)
 apply (auto simp add: merge_Allowed ok_iff_Allowed)
 done
 
 lemma (in merge) merge_Always_Out_eq_iOut: 
      "[| G \<in> preserves(lift(Out)); G \<in> preserves(lift(iOut));  
- 	 M \<in> Allowed(G) |]
+         M \<in> Allowed(G) |]
       ==> M \<squnion> G \<in> Always({s \<in> state. length(s`Out)=length(s`iOut)})"
 apply (frule preserves_type [THEN subsetD])
 apply (subgoal_tac "G \<in> program")
@@ -135,7 +134,7 @@
 
 lemma (in merge) merge_Bounded: 
      "[| G \<in> preserves(lift(iOut)); G \<in> preserves(lift(Out));  
-	 M \<in> Allowed(G) |] ==>  
+         M \<in> Allowed(G) |] ==>  
        M \<squnion> G: Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iOut). elt<Nclients})"
 apply (frule preserves_type [THEN subsetD])
 apply (frule M_ok_iff)
@@ -174,7 +173,7 @@
 
 theorem (in merge) merge_bag_Follows: 
     "M \<in> (\<Inter>n \<in> Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A))  
-	    guarantees   
+            guarantees   
       (%s. bag_of(s`Out)) Fols  
       (%s. msetsum(%i. bag_of(s`In(i)),Nclients, A)) Wrt MultLe(A, r)/Mult(A)"
 apply (cut_tac merge_spec)
@@ -189,9 +188,9 @@
 apply (drule guaranteesD, assumption)
   apply (simp add: merge_spec_def merge_follows_def, blast)
 apply (simp cong add: Follows_cong
-	    add: refl_prefix
-	       mono_bag_of [THEN subset_Follows_comp, THEN subsetD, 
-			    unfolded metacomp_def])
+            add: refl_prefix
+               mono_bag_of [THEN subset_Follows_comp, THEN subsetD, 
+                            unfolded metacomp_def])
 done
 
 end
--- a/src/ZF/UNITY/Monotonicity.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Monotonicity.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,9 @@
 (*  Title:      ZF/UNITY/Monotonicity.thy
-    ID:         $Id \<in> Monotonicity.thy,v 1.1 2003/05/28 16:13:42 paulson Exp $
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2002  University of Cambridge
 
-Monotonicity of an operator (meta-function) with respect to arbitrary set relations.
+Monotonicity of an operator (meta-function) with respect to arbitrary
+set relations.
 *)
 
 header{*Monotonicity of an Operator WRT a Relation*}
--- a/src/ZF/UNITY/MultisetSum.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/MultisetSum.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/UNITY/MultusetSum.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety
 *)
 
@@ -122,7 +121,7 @@
      "[| I \<in> Fin(K); \<forall>i \<in> K. C(i) \<in> Fin(A) |] ==>  
       (\<forall>x \<in> A. multiset(f(x)) & mset_of(f(x))\<subseteq>B) -->   
       (\<forall>i \<in> I. \<forall>j \<in> I. i\<noteq>j --> C(i) Int C(j) = 0) -->  
-	msetsum(f, \<Union>i \<in> I. C(i), B) = msetsum (%i. msetsum(f, C(i),B), I, B)"
+        msetsum(f, \<Union>i \<in> I. C(i), B) = msetsum (%i. msetsum(f, C(i),B), I, B)"
 apply (erule Fin_induct, auto)
 apply (subgoal_tac "\<forall>i \<in> y. x \<noteq> i")
  prefer 2 apply blast
--- a/src/ZF/UNITY/Mutex.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Mutex.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,8 @@
-(*  ID:         $Id$
+(*  Title:      ZF/UNITY/Mutex.thy
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   2001  University of Cambridge
 
-Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra
+Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra.
 
 Variables' types are introduced globally so that type verification
 reduces to the usual ZF typechecking \<in> an ill-tyed expression will
@@ -76,11 +76,11 @@
 
 definition
   "IU == {s:state. (s`u = 1<->(#1 $<= s`m & s`m $<= #3))
-	             & (s`m = #3 --> s`p=0)}"
+                     & (s`m = #3 --> s`p=0)}"
 
 definition
   "IV == {s:state. (s`v = 1 <-> (#1 $<= s`n & s`n $<= #3))
-	              & (s`n = #3 --> s`p=1)}"
+                      & (s`n = #3 --> s`p=1)}"
 
   (** The faulty invariant (for U alone) **)
 
--- a/src/ZF/UNITY/State.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/State.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/State.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   2001  University of Cambridge
 
--- a/src/ZF/UNITY/SubstAx.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/SubstAx.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,4 +1,4 @@
-(*  ID:         $Id$
+(*  Title:      ZF/UNITY/SubstAx.thy
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   2001  University of Cambridge
 
--- a/src/ZF/UNITY/UNITY.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/UNITY.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/UNITY/UNITY.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   2001  University of Cambridge
 *)
@@ -20,8 +19,8 @@
 definition
   program  :: i  where
   "program == {<init, acts, allowed>:
-	       Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)).
-	       id(state) \<in> acts & id(state) \<in> allowed}"
+               Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)).
+               id(state) \<in> acts & id(state) \<in> allowed}"
 
 definition
   mk_program :: "[i,i,i]=>i"  where
--- a/src/ZF/UNITY/Union.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/Union.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/UNITY/Union.thy
-    ID:         $Id$
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   2001  University of Cambridge
 
@@ -34,7 +33,7 @@
 definition
   Join :: "[i, i] => i"    (infixl "Join" 65)  where
   "F Join G == mk_program (Init(F) Int Init(G), Acts(F) Un Acts(G),
-				AllowedActs(F) Int AllowedActs(G))"
+                                AllowedActs(F) Int AllowedActs(G))"
 definition
   (*Characterizes safety properties.  Used with specifying AllowedActs*)
   safety_prop :: "i => o"  where
--- a/src/ZF/UNITY/WFair.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/UNITY/WFair.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/WFair.thy
-    ID:         $Id$
     Author:     Sidi Ehmety, Computer Laboratory
     Copyright   1998  University of Cambridge
 *)
@@ -19,7 +18,7 @@
     is generic to all forms of fairness.*) 
   transient :: "i=>i"  where
   "transient(A) =={F:program. (EX act: Acts(F). A<=domain(act) &
-			       act``A <= state-A) & st_set(A)}"
+                               act``A <= state-A) & st_set(A)}"
 
 definition
   ensures :: "[i,i] => i"       (infixl "ensures" 60)  where
@@ -37,7 +36,7 @@
     Basis:  "[| F:A ensures B;  A:Pow(D); B:Pow(D) |] ==> <A,B>:leads(D, F)"
     Trans:  "[| <A,B> : leads(D, F); <B,C> : leads(D, F) |] ==>  <A,C>:leads(D, F)"
     Union:   "[| S:Pow({A:S. <A, B>:leads(D, F)}); B:Pow(D); S:Pow(Pow(D)) |] ==> 
-	      <Union(S),B>:leads(D, F)"
+              <Union(S),B>:leads(D, F)"
 
   monos        Pow_mono
   type_intros  Union_Pow_iff [THEN iffD2] UnionI PowI
--- a/src/ZF/Univ.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Univ.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,9 @@
-(*  Title:      ZF/univ.thy
-    ID:         $Id$
+(*  Title:      ZF/Univ.thy
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 
-Standard notation for Vset(i) is V(i), but users might want V for a variable
+Standard notation for Vset(i) is V(i), but users might want V for a
+variable.
 
 NOTE: univ(A) could be a translation; would simplify many proofs!
   But Ind_Syntax.univ refers to the constant "Univ.univ"
@@ -25,12 +25,12 @@
 definition
   Vrec        :: "[i, [i,i]=>i] =>i"  where
     "Vrec(a,H) == transrec(rank(a), %x g. lam z: Vset(succ(x)).
- 		 	   H(z, lam w:Vset(x). g`rank(w)`w)) ` a"
+                           H(z, lam w:Vset(x). g`rank(w)`w)) ` a"
 
 definition
   Vrecursor   :: "[[i,i]=>i, i] =>i"  where
     "Vrecursor(H,a) == transrec(rank(a), %x g. lam z: Vset(succ(x)).
-				H(lam w:Vset(x). g`rank(w)`w, z)) ` a"
+                                H(lam w:Vset(x). g`rank(w)`w, z)) ` a"
 
 definition
   univ        :: "i=>i"  where
--- a/src/ZF/ZF.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/ZF.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/ZF.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson and Martin D Coen, CU Computer Laboratory
     Copyright   1993  University of Cambridge
 *)
@@ -285,7 +284,7 @@
   field_def:    "field(r) == domain(r) Un range(r)"
   relation_def: "relation(r) == \<forall>z\<in>r. \<exists>x y. z = <x,y>"
   function_def: "function(r) ==
-		    \<forall>x y. <x,y>:r --> (\<forall>y'. <x,y'>:r --> y=y')"
+                    \<forall>x y. <x,y>:r --> (\<forall>y'. <x,y'>:r --> y=y')"
   image_def:    "r `` A  == {y : range(r) . \<exists>x\<in>A. <x,y> : r}"
   vimage_def:   "r -`` A == converse(r)``A"
 
--- a/src/ZF/Zorn.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/Zorn.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/Zorn.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
-
 *)
 
 header{*Zorn's Lemma*}
@@ -120,7 +118,7 @@
 apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE],
        assumption+)
 apply (blast del: subsetI
-	     intro: increasing_trans subsetI, blast)
+             intro: increasing_trans subsetI, blast)
 txt{*second induction step*}
 apply (rule impI [THEN ballI])
 apply (rule Union_lemma0 [THEN disjE])
@@ -172,7 +170,7 @@
 apply (erule ssubst)
 apply (rule increasingD2 [THEN equalityI], assumption)
 apply (blast del: subsetI
-	     intro: subsetI TFin_UnionI TFin.nextI TFin_is_subset)+
+             intro: subsetI TFin_UnionI TFin.nextI TFin_is_subset)+
 done
 
 
@@ -412,8 +410,8 @@
                     <= Union ({y \<in> TFin (S,next) . x \<notin> y}) ")
  prefer 2
  apply (simp del: Union_iff
-	     add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset]
-	     Pow_iff cons_subset_iff subset_refl choice_Diff [THEN DiffD2])
+             add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset]
+             Pow_iff cons_subset_iff subset_refl choice_Diff [THEN DiffD2])
 apply (subgoal_tac "x \<in> next ` Union ({y \<in> TFin (S,next) . x \<notin> y}) ")
  prefer 2
  apply (blast intro!: Collect_subset [THEN TFin_UnionI] TFin.nextI)
@@ -480,8 +478,8 @@
       assume "a : field(r)" "r -`` {a} : C" "b : field(r)" "r -`` {b} : C"
       hence "r -`` {a} \<subseteq> r -`` {b} | r -`` {b} \<subseteq> r -`` {a}" using 2 by auto
       then show "<a, b> : r | <b, a> : r"
-	using `Preorder(r)` `a : field(r)` `b : field(r)`
-	by (simp add: subset_vimage1_vimage1_iff)
+        using `Preorder(r)` `a : field(r)` `b : field(r)`
+        by (simp add: subset_vimage1_vimage1_iff)
     qed
     then obtain u where uA: "u : field(r)" "ALL a:?A. <a, u> : r"
       using u
@@ -494,14 +492,14 @@
       fix a B
       assume aB: "B : C" "a : B"
       with 1 obtain x where "x : field(r)" "B = r -`` {x}"
-	apply -
-	apply (drule subsetD) apply assumption
-	apply (erule imageE)
-	apply (erule lamE)
-	apply simp
-	done
+        apply -
+        apply (drule subsetD) apply assumption
+        apply (erule imageE)
+        apply (erule lamE)
+        apply simp
+        done
       then show "<a, u> : r" using uA aB `Preorder(r)`
-	by (auto simp: preorder_on_def refl_def) (blast dest: trans_onD)+
+        by (auto simp: preorder_on_def refl_def) (blast dest: trans_onD)+
     qed
     then show "EX U:field(r). ALL A:C. A \<subseteq> r -`` {U}"
       using `u : field(r)` ..
--- a/src/ZF/ex/Group.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/ex/Group.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,6 +1,4 @@
-(* Title:  ZF/ex/Group.thy
-
-*)
+(*  Title:      ZF/ex/Group.thy *)
 
 header {* Groups *}
 
@@ -114,9 +112,9 @@
     { 
       assume eq: "x \<cdot> y = x \<cdot> z"
       with G l_inv_ex obtain x_inv where xG: "x_inv \<in> carrier(G)"
-	and l_inv: "x_inv \<cdot> x = \<one>" by fast
+        and l_inv: "x_inv \<cdot> x = \<one>" by fast
       from G eq xG have "(x_inv \<cdot> x) \<cdot> y = (x_inv \<cdot> x) \<cdot> z"
-	by (simp add: m_assoc)
+        by (simp add: m_assoc)
       with G show "y = z" by (simp add: l_inv)
     next
       assume eq: "y = z"
@@ -882,7 +880,7 @@
     proof (simp add: r_congruent_def sym_def, clarify)
       fix x y
       assume [simp]: "x \<in> carrier(G)" "y \<in> carrier(G)" 
-	and "inv x \<cdot> y \<in> H"
+        and "inv x \<cdot> y \<in> H"
       hence "inv (inv x \<cdot> y) \<in> H" by (simp add: m_inv_closed) 
       thus "inv y \<cdot> x \<in> H" by (simp add: inv_mult_group)
     qed
@@ -891,7 +889,7 @@
     proof (simp add: r_congruent_def trans_def, clarify)
       fix x y z
       assume [simp]: "x \<in> carrier(G)" "y \<in> carrier(G)" "z \<in> carrier(G)"
-	and "inv x \<cdot> y \<in> H" and "inv y \<cdot> z \<in> H"
+        and "inv x \<cdot> y \<in> H" and "inv y \<cdot> z \<in> H"
       hence "(inv x \<cdot> y) \<cdot> (inv y \<cdot> z) \<in> H" by simp
       hence "inv x \<cdot> (y \<cdot> inv y) \<cdot> z \<in> H" by (simp add: m_assoc del: inv) 
       thus "inv x \<cdot> z \<in> H" by simp
--- a/src/ZF/ex/LList.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/ex/LList.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,9 +1,8 @@
 (*  Title:      ZF/ex/LList.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
 
-Codatatype definition of Lazy Lists
+Codatatype definition of Lazy Lists.
 
 Equality for llist(A) as a greatest fixed point
 
@@ -65,10 +64,10 @@
         Union_least [intro!]
         UN_least [intro!]
         Un_least [intro!]
-	Inter_greatest [intro!]
+        Inter_greatest [intro!]
         Int_greatest [intro!]
         RepFun_subset [intro!]
-	Un_upper1 [intro!]
+        Un_upper1 [intro!]
         Un_upper2 [intro!]
         Int_lower1 [intro!]
         Int_lower2 [intro!]
--- a/src/ZF/ex/Limit.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/ex/Limit.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,22 +1,20 @@
-(*  Title:      ZF/ex/Limit
-    ID:         $Id$
+(*  Title:      ZF/ex/Limit.thy
     Author:     Sten Agerholm
+    Author:     Lawrence C Paulson
 
-    A formalization of the inverse limit construction of domain theory.
+A formalization of the inverse limit construction of domain theory.
 
-    The following paper comments on the formalization:
+The following paper comments on the formalization:
 
-    "A Comparison of HOL-ST and Isabelle/ZF" by Sten Agerholm
-    In Proceedings of the First Isabelle Users Workshop, Technical 
-    Report No. 379, University of Cambridge Computer Laboratory, 1995.
-
-    This is a condensed version of:
+"A Comparison of HOL-ST and Isabelle/ZF" by Sten Agerholm
+In Proceedings of the First Isabelle Users Workshop, Technical 
+Report No. 379, University of Cambridge Computer Laboratory, 1995.
 
-    "A Comparison of HOL-ST and Isabelle/ZF" by Sten Agerholm
-    Technical Report No. 369, University of Cambridge Computer 
-    Laboratory, 1995.
+This is a condensed version of:
 
-(Proofs converted to Isar and tidied up considerably by lcp)
+"A Comparison of HOL-ST and Isabelle/ZF" by Sten Agerholm
+Technical Report No. 369, University of Cambridge Computer 
+Laboratory, 1995.
 *)
 
 theory Limit  imports  Main begin
@@ -523,24 +521,24 @@
   proof (rule matrix_chain_left [THEN cpo_lub, THEN islub_least], simp_all add: n D DM)
     show "isub(D, M ` n, y)" 
       proof (unfold isub_def, intro conjI ballI y)
-	fix k assume k: "k \<in> nat"
-	show "rel(D, M ` n ` k, y)"
-	  proof (cases "n le k")
-	    case True 
-	    hence yy: "rel(D, M`n`k, M`k`k)" 
-	      by (blast intro: lemma2 n k y DM D chain_rel_gen matrix_chain_right) 
-	    show "?thesis"
-	      by (rule cpo_trans [OF D yy], 
+        fix k assume k: "k \<in> nat"
+        show "rel(D, M ` n ` k, y)"
+          proof (cases "n le k")
+            case True 
+            hence yy: "rel(D, M`n`k, M`k`k)" 
+              by (blast intro: lemma2 n k y DM D chain_rel_gen matrix_chain_right) 
+            show "?thesis"
+              by (rule cpo_trans [OF D yy], 
                   simp_all add: k rel n y DM matrix_in)
-	  next
-	    case False
-	    hence le: "k le n" 
-	      by (blast intro: not_le_iff_lt [THEN iffD1, THEN leI] nat_into_Ord n k) 
-	    show "?thesis"
-	      by (rule cpo_trans [OF D chain_rel_gen [OF le]], 
-		  simp_all add: n y k rel DM D matrix_chain_left)
-	  qed
-	qed
+          next
+            case False
+            hence le: "k le n" 
+              by (blast intro: not_le_iff_lt [THEN iffD1, THEN leI] nat_into_Ord n k) 
+            show "?thesis"
+              by (rule cpo_trans [OF D chain_rel_gen [OF le]], 
+                  simp_all add: n y k rel DM D matrix_chain_left)
+          qed
+        qed
   qed
   moreover
   have "M ` n \<in> nat \<rightarrow> set(D)" by (blast intro: DM n matrix_fun [THEN apply_type])
@@ -952,7 +950,7 @@
     "[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p');
        rel(cf(E,D),p',p)|] ==> rel(cf(D,E),e,e')"
 apply (rule_tac b=e
-	 in projpair_e_cont [THEN cont_fun, THEN comp_id, THEN subst],
+         in projpair_e_cont [THEN cont_fun, THEN comp_id, THEN subst],
        assumption)
 apply (rule_tac e1=e' in projpair_eq [THEN subst], assumption)
 apply (rule cpo_trans)
@@ -1113,7 +1111,7 @@
 apply (rule_tac P = "%t. rel (DD`na,t,lub (DD`na,\<lambda>x \<in> nat. X`x`na))"
             and b1 = "%n. X`n`na" in beta [THEN subst])
 apply (simp del: beta_if
-	    add: chain_iprod [THEN cpo_lub, THEN islub_ub] iprodE
+            add: chain_iprod [THEN cpo_lub, THEN islub_ub] iprodE
                 chain_in)+
 apply (blast intro: iprodI lam_type chain_iprod [THEN cpo_lub, THEN islub_in])
 apply (rule rel_iprodI)
@@ -1414,7 +1412,7 @@
 apply (subst comp_assoc)
 apply (assumption | rule comp_mono_eq refl)+
 apply (simp del: add_succ_right add: add_succ_right [symmetric]
-	    add: e_less_eq add_type nat_succI)
+            add: e_less_eq add_type nat_succI)
 apply (subst id_comp) (* simp cannot unify/inst right, use brr below (?) . *)
 apply (assumption |
        rule emb_e_less_add [THEN emb_cont, THEN cont_fun] refl nat_succI)+
@@ -1487,7 +1485,7 @@
 apply (assumption | rule comp_mono_eq refl)+
 (* New direct subgoal *)
 apply (simp del: add_succ_right add: add_succ_right [symmetric]
-	    add: e_gr_eq)
+            add: e_gr_eq)
 apply (subst comp_id) (* simp cannot unify/inst right, use brr below (?) . *)
 apply (assumption | rule e_gr_fun add_type refl add_le_self nat_succI)+
 done
@@ -1541,7 +1539,7 @@
 apply (blast intro: e_less_cont [THEN cont_fun] add_le_self)
 apply (rule refl)
 apply (simp del: add_succ_right add: add_succ_right [symmetric]
-	    add: e_gr_eq)
+            add: e_gr_eq)
 apply (blast intro: id_comp [symmetric] e_less_cont [THEN cont_fun]
                     add_le_self)
 done
@@ -1568,7 +1566,7 @@
 apply (blast intro!: e_less_cont [THEN cont_fun] add_le_mono nat_le_refl)
 apply (rule refl)
 apply (simp del: add_succ_right add: add_succ_right [symmetric]
-	    add: e_less_eq)
+            add: e_less_eq)
 apply (blast intro: comp_id [symmetric] e_gr_cont [THEN cont_fun] add_le_self)
 done
 
@@ -2112,7 +2110,7 @@
      ==> chain(cf(G,E),\<lambda>n \<in> nat. r(n) O Rp(DD`n,G,f(n)))"
 apply (rule chainI)
 apply (blast intro: lam_type cont_cf comp_pres_cont commute_emb Rp_cont 
-		    emb_cont emb_chain_cpo, simp)
+                    emb_cont emb_chain_cpo, simp)
 apply (rule_tac r1 = r and m1 = n in commute_eq [THEN subst])
 apply (rule_tac [5] r1 = f and m1 = n in commute_eq [THEN subst])
 apply (assumption | rule le_succ nat_succI)+
@@ -2122,11 +2120,11 @@
 apply (rule_tac r1 = "r (succ (n))" in comp_assoc [THEN ssubst])
 apply (rule comp_mono)
 apply (blast intro: comp_pres_cont eps_cont emb_eps commute_emb Rp_cont 
-		    emb_cont emb_chain_cpo le_succ)+
+                    emb_cont emb_chain_cpo le_succ)+
 apply (rule_tac b="r(succ(n))" in comp_id [THEN subst]) (* 1 subst too much *)
 apply (rule_tac [2] comp_mono)
 apply (blast intro: comp_pres_cont eps_cont emb_eps emb_id commute_emb 
-		    Rp_cont emb_cont cont_fun emb_chain_cpo le_succ)+
+                    Rp_cont emb_cont cont_fun emb_chain_cpo le_succ)+
 apply (subst comp_id) (* Undoes "1 subst too much", typing next anyway *)
 apply (blast intro: cont_fun Rp_cont emb_cont commute_emb cont_cf cpo_cf 
                     emb_chain_cpo embRp_rel emb_eps le_succ)+
@@ -2158,9 +2156,9 @@
       commute(DD,ee,E,r); commute(DD,ee,G,f);
       emb_chain(DD,ee); cpo(E); cpo(G) |]
    ==> projpair
-	(E,G,
-	 lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n))),
-	 lub(cf(G,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,G,f(n))))"
+        (E,G,
+         lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n))),
+         lub(cf(G,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,G,f(n))))"
 apply (simp add: projpair_def rho_proj_def, safe)
 apply (rule_tac [3] comp_lubs [THEN ssubst])
 (* The following one line is 15 lines in HOL, and includes existentials. *)
@@ -2185,8 +2183,8 @@
 
 lemma emb_theta:
     "[| lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
-	commute(DD,ee,E,r); commute(DD,ee,G,f);
-	emb_chain(DD,ee); cpo(E); cpo(G) |] 
+        commute(DD,ee,E,r); commute(DD,ee,G,f);
+        emb_chain(DD,ee); cpo(E); cpo(G) |] 
      ==> emb(E,G,lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n))))"
 apply (simp add: emb_def)
 apply (blast intro: theta_projpair)
@@ -2271,9 +2269,9 @@
 
 lemma lub_universal_unique:
     "[| mediating(E,G,r,f,t);
-	lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
-	commute(DD,ee,E,r); commute(DD,ee,G,f);
-	emb_chain(DD,ee); cpo(E); cpo(G) |] 
+        lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
+        commute(DD,ee,E,r); commute(DD,ee,G,f);
+        emb_chain(DD,ee); cpo(E); cpo(G) |] 
      ==> t = lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n)))"
 apply (rule_tac b = t in comp_id [THEN subst])
 apply (erule_tac [2] subst)
--- a/src/ZF/ex/Primes.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/ex/Primes.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/ex/Primes.thy
-    ID:         $Id$
     Author:     Christophe Tabacznyj and Lawrence C Paulson
     Copyright   1996  University of Cambridge
 *)
@@ -20,8 +19,8 @@
 definition
   gcd     :: "[i,i]=>i"       --{*Euclid's algorithm for the gcd*}  where
     "gcd(m,n) == transrec(natify(n),
-			%n f. \<lambda>m \<in> nat.
-			        if n=0 then m else f`(m mod n)`n) ` natify(m)"
+                        %n f. \<lambda>m \<in> nat.
+                                if n=0 then m else f`(m mod n)`n) ` natify(m)"
 
 definition
   coprime :: "[i,i]=>o"       --{*the coprime relation*}  where
--- a/src/ZF/func.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/func.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,8 +1,6 @@
 (*  Title:      ZF/func.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1991  University of Cambridge
-
 *)
 
 header{*Functions, Function Spaces, Lambda-Abstraction*}
@@ -230,7 +228,7 @@
 (*thm by Mark Staples, proof by lcp*)
 lemma fun_subset_eq: "[| f:Pi(A,B); g:Pi(A,C) |] ==> f <= g <-> (f = g)"
 by (blast dest: apply_Pair
-	  intro: fun_extension apply_equality [symmetric])
+          intro: fun_extension apply_equality [symmetric])
 
 
 (*Every element of Pi(A,B) may be expressed as a lambda abstraction!*)
--- a/src/ZF/ind_syntax.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/ind_syntax.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/ind_syntax.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 
@@ -25,7 +24,7 @@
 fun mk_all_imp (A,P) =
     FOLogic.all_const iT $
       Abs("v", iT, FOLogic.imp $ (@{const mem} $ Bound 0 $ A) $
-	           Term.betapply(P, Bound 0));
+                   Term.betapply(P, Bound 0));
 
 fun mk_Collect (a, D, t) = @{const Collect} $ D $ absfree (a, iT, t);
 
@@ -84,9 +83,9 @@
     fun mk_intr ((id,T,syn), name, args, prems) =
       Logic.list_implies
         (map FOLogic.mk_Trueprop prems,
-	 FOLogic.mk_Trueprop
-	    (@{const mem} $ list_comb (Const (Sign.full_bname sg name, T), args)
-	               $ rec_tm))
+         FOLogic.mk_Trueprop
+            (@{const mem} $ list_comb (Const (Sign.full_bname sg name, T), args)
+                       $ rec_tm))
   in  map mk_intr constructs  end;
 
 fun mk_all_intr_tms sg arg = flat (ListPair.map (mk_intr_tms sg) arg);
@@ -115,7 +114,7 @@
   | tryres (th, []) = raise THM("tryres", 0, [th]);
 
 fun gen_make_elim elim_rls rl =
-      standard (tryres (rl, elim_rls @ [revcut_rl]));
+      Drule.standard (tryres (rl, elim_rls @ [revcut_rl]));
 
 (*Turns iff rules into safe elimination rules*)
 fun mk_free_SEs iffs = map (gen_make_elim [@{thm conjE}, @{thm FalseE}]) (iffs RL [@{thm iffD1}]);
--- a/src/ZF/int_arith.ML	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/int_arith.ML	Sat Oct 17 14:51:30 2009 +0200
@@ -112,7 +112,7 @@
 
 (*push the unary minus down: - x * y = x * - y *)
 val int_minus_mult_eq_1_to_2 =
-    [@{thm zmult_zminus}, @{thm zmult_zminus_right} RS sym] MRS trans |> standard;
+    [@{thm zmult_zminus}, @{thm zmult_zminus_right} RS sym] MRS trans |> Drule.standard;
 
 (*to extract again any uncancelled minuses*)
 val int_minus_from_mult_simps =
--- a/src/ZF/upair.thy	Sat Oct 17 13:46:55 2009 +0200
+++ b/src/ZF/upair.thy	Sat Oct 17 14:51:30 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      ZF/upair.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson and Martin D Coen, CU Computer Laboratory
     Copyright   1993  University of Cambridge
 
@@ -226,7 +225,7 @@
 by (case_tac Q, simp_all)
 
 (** Rewrite rules for boolean case-splitting: faster than 
-	addsplits[split_if]
+        addsplits[split_if]
 **)
 
 lemmas split_if_eq1 = split_if [of "%x. x = b", standard]