misc tuning;
authorwenzelm
Tue, 21 Feb 2012 16:48:10 +0100
changeset 46575 f1e387195a56
parent 46574 41701fce8ac7
child 46576 ae9286f64574
misc tuning; more indentation;
src/HOL/Library/Function_Algebras.thy
src/HOL/Library/Ramsey.thy
--- a/src/HOL/Library/Function_Algebras.thy	Tue Feb 21 16:42:57 2012 +0100
+++ b/src/HOL/Library/Function_Algebras.thy	Tue Feb 21 16:48:10 2012 +0100
@@ -13,9 +13,7 @@
 instantiation "fun" :: (type, plus) plus
 begin
 
-definition
-  "f + g = (\<lambda>x. f x + g x)"
-
+definition "f + g = (\<lambda>x. f x + g x)"
 instance ..
 
 end
@@ -23,9 +21,7 @@
 instantiation "fun" :: (type, zero) zero
 begin
 
-definition
-  "0 = (\<lambda>x. 0)"
-
+definition "0 = (\<lambda>x. 0)"
 instance ..
 
 end
@@ -33,9 +29,7 @@
 instantiation "fun" :: (type, times) times
 begin
 
-definition
-  "f * g = (\<lambda>x. f x * g x)"
-
+definition "f * g = (\<lambda>x. f x * g x)"
 instance ..
 
 end
@@ -43,9 +37,7 @@
 instantiation "fun" :: (type, one) one
 begin
 
-definition
-  "1 = (\<lambda>x. 1)"
-
+definition "1 = (\<lambda>x. 1)"
 instance ..
 
 end
@@ -53,69 +45,70 @@
 
 text {* Additive structures *}
 
-instance "fun" :: (type, semigroup_add) semigroup_add proof
-qed (simp add: plus_fun_def add.assoc)
+instance "fun" :: (type, semigroup_add) semigroup_add
+  by default (simp add: plus_fun_def add.assoc)
 
-instance "fun" :: (type, cancel_semigroup_add) cancel_semigroup_add proof
-qed (simp_all add: plus_fun_def fun_eq_iff)
+instance "fun" :: (type, cancel_semigroup_add) cancel_semigroup_add
+  by default (simp_all add: plus_fun_def fun_eq_iff)
 
-instance "fun" :: (type, ab_semigroup_add) ab_semigroup_add proof
-qed (simp add: plus_fun_def add.commute)
+instance "fun" :: (type, ab_semigroup_add) ab_semigroup_add
+  by default (simp add: plus_fun_def add.commute)
 
-instance "fun" :: (type, cancel_ab_semigroup_add) cancel_ab_semigroup_add proof
-qed simp
+instance "fun" :: (type, cancel_ab_semigroup_add) cancel_ab_semigroup_add
+  by default simp
 
-instance "fun" :: (type, monoid_add) monoid_add proof
-qed (simp_all add: plus_fun_def zero_fun_def)
+instance "fun" :: (type, monoid_add) monoid_add
+  by default (simp_all add: plus_fun_def zero_fun_def)
 
-instance "fun" :: (type, comm_monoid_add) comm_monoid_add proof
-qed simp
+instance "fun" :: (type, comm_monoid_add) comm_monoid_add
+  by default simp
 
 instance "fun" :: (type, cancel_comm_monoid_add) cancel_comm_monoid_add ..
 
-instance "fun" :: (type, group_add) group_add proof
-qed (simp_all add: plus_fun_def zero_fun_def fun_Compl_def fun_diff_def diff_minus)
+instance "fun" :: (type, group_add) group_add
+  by default
+    (simp_all add: plus_fun_def zero_fun_def fun_Compl_def fun_diff_def diff_minus)
 
-instance "fun" :: (type, ab_group_add) ab_group_add proof
-qed (simp_all add: diff_minus)
+instance "fun" :: (type, ab_group_add) ab_group_add
+  by default (simp_all add: diff_minus)
 
 
 text {* Multiplicative structures *}
 
-instance "fun" :: (type, semigroup_mult) semigroup_mult proof
-qed (simp add: times_fun_def mult.assoc)
+instance "fun" :: (type, semigroup_mult) semigroup_mult
+  by default (simp add: times_fun_def mult.assoc)
 
-instance "fun" :: (type, ab_semigroup_mult) ab_semigroup_mult proof
-qed (simp add: times_fun_def mult.commute)
+instance "fun" :: (type, ab_semigroup_mult) ab_semigroup_mult
+  by default (simp add: times_fun_def mult.commute)
 
-instance "fun" :: (type, ab_semigroup_idem_mult) ab_semigroup_idem_mult proof
-qed (simp add: times_fun_def)
+instance "fun" :: (type, ab_semigroup_idem_mult) ab_semigroup_idem_mult
+  by default (simp add: times_fun_def)
 
-instance "fun" :: (type, monoid_mult) monoid_mult proof
-qed (simp_all add: times_fun_def one_fun_def)
+instance "fun" :: (type, monoid_mult) monoid_mult
+  by default (simp_all add: times_fun_def one_fun_def)
 
-instance "fun" :: (type, comm_monoid_mult) comm_monoid_mult proof
-qed simp
+instance "fun" :: (type, comm_monoid_mult) comm_monoid_mult
+  by default simp
 
 
 text {* Misc *}
 
 instance "fun" :: (type, "Rings.dvd") "Rings.dvd" ..
 
-instance "fun" :: (type, mult_zero) mult_zero proof
-qed (simp_all add: zero_fun_def times_fun_def)
+instance "fun" :: (type, mult_zero) mult_zero
+  by default (simp_all add: zero_fun_def times_fun_def)
 
-instance "fun" :: (type, zero_neq_one) zero_neq_one proof
-qed (simp add: zero_fun_def one_fun_def fun_eq_iff)
+instance "fun" :: (type, zero_neq_one) zero_neq_one
+  by default (simp add: zero_fun_def one_fun_def fun_eq_iff)
 
 
 text {* Ring structures *}
 
-instance "fun" :: (type, semiring) semiring proof
-qed (simp_all add: plus_fun_def times_fun_def algebra_simps)
+instance "fun" :: (type, semiring) semiring
+  by default (simp_all add: plus_fun_def times_fun_def algebra_simps)
 
-instance "fun" :: (type, comm_semiring) comm_semiring proof
-qed (simp add: plus_fun_def times_fun_def algebra_simps)
+instance "fun" :: (type, comm_semiring) comm_semiring
+  by default (simp add: plus_fun_def times_fun_def algebra_simps)
 
 instance "fun" :: (type, semiring_0) semiring_0 ..
 
@@ -127,8 +120,7 @@
 
 instance "fun" :: (type, semiring_1) semiring_1 ..
 
-lemma of_nat_fun:
-  shows "of_nat n = (\<lambda>x::'a. of_nat n)"
+lemma of_nat_fun: "of_nat n = (\<lambda>x::'a. of_nat n)"
 proof -
   have comp: "comp = (\<lambda>f g x. f (g x))"
     by (rule ext)+ simp
@@ -147,7 +139,8 @@
 
 instance "fun" :: (type, comm_semiring_1_cancel) comm_semiring_1_cancel ..
 
-instance "fun" :: (type, semiring_char_0) semiring_char_0 proof
+instance "fun" :: (type, semiring_char_0) semiring_char_0
+proof
   from inj_of_nat have "inj (\<lambda>n (x::'a). of_nat n :: 'b)"
     by (rule inj_fun)
   then have "inj (\<lambda>n. of_nat n :: 'a \<Rightarrow> 'b)"
@@ -168,23 +161,24 @@
 
 text {* Ordereded structures *}
 
-instance "fun" :: (type, ordered_ab_semigroup_add) ordered_ab_semigroup_add proof
-qed (auto simp add: plus_fun_def le_fun_def intro: add_left_mono)
+instance "fun" :: (type, ordered_ab_semigroup_add) ordered_ab_semigroup_add
+  by default (auto simp add: plus_fun_def le_fun_def intro: add_left_mono)
 
 instance "fun" :: (type, ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add ..
 
-instance "fun" :: (type, ordered_ab_semigroup_add_imp_le) ordered_ab_semigroup_add_imp_le proof
-qed (simp add: plus_fun_def le_fun_def)
+instance "fun" :: (type, ordered_ab_semigroup_add_imp_le) ordered_ab_semigroup_add_imp_le
+  by default (simp add: plus_fun_def le_fun_def)
 
 instance "fun" :: (type, ordered_comm_monoid_add) ordered_comm_monoid_add ..
 
 instance "fun" :: (type, ordered_ab_group_add) ordered_ab_group_add ..
 
-instance "fun" :: (type, ordered_semiring) ordered_semiring proof
-qed (auto simp add: zero_fun_def times_fun_def le_fun_def intro: mult_left_mono mult_right_mono)
+instance "fun" :: (type, ordered_semiring) ordered_semiring
+  by default
+    (auto simp add: zero_fun_def times_fun_def le_fun_def intro: mult_left_mono mult_right_mono)
 
-instance "fun" :: (type, ordered_comm_semiring) ordered_comm_semiring proof
-qed (fact mult_left_mono)
+instance "fun" :: (type, ordered_comm_semiring) ordered_comm_semiring
+  by default (fact mult_left_mono)
 
 instance "fun" :: (type, ordered_cancel_semiring) ordered_cancel_semiring ..
 
--- a/src/HOL/Library/Ramsey.thy	Tue Feb 21 16:42:57 2012 +0100
+++ b/src/HOL/Library/Ramsey.thy	Tue Feb 21 16:48:10 2012 +0100
@@ -47,11 +47,11 @@
     qed
   } moreover
   { assume "m\<noteq>0" "n\<noteq>0"
-    hence "k = (m - 1) + n" "k = m + (n - 1)" using `Suc k = m+n` by auto
-    from Suc(1)[OF this(1)] Suc(1)[OF this(2)] 
+    then have "k = (m - 1) + n" "k = m + (n - 1)" using `Suc k = m+n` by auto
+    from Suc(1)[OF this(1)] Suc(1)[OF this(2)]
     obtain r1 r2 where "r1\<ge>1" "r2\<ge>1" "?R (m - 1) n r1" "?R m (n - 1) r2"
       by auto
-    hence "r1+r2 \<ge> 1" by arith
+    then have "r1+r2 \<ge> 1" by arith
     moreover
     have "?R m n (r1+r2)" (is "ALL V E. _ \<longrightarrow> ?EX V E m n")
     proof clarify
@@ -62,12 +62,12 @@
       let ?M = "{w : V. w\<noteq>v & {v,w} : E}"
       let ?N = "{w : V. w\<noteq>v & {v,w} ~: E}"
       have "V = insert v (?M \<union> ?N)" using `v : V` by auto
-      hence "card V = card(insert v (?M \<union> ?N))" by metis
+      then have "card V = card(insert v (?M \<union> ?N))" by metis
       also have "\<dots> = card ?M + card ?N + 1" using `finite V`
         by(fastforce intro: card_Un_disjoint)
       finally have "card V = card ?M + card ?N + 1" .
-      hence "r1+r2 \<le> card ?M + card ?N + 1" using `r1+r2 \<le> card V` by simp
-      hence "r1 \<le> card ?M \<or> r2 \<le> card ?N" by arith
+      then have "r1+r2 \<le> card ?M + card ?N + 1" using `r1+r2 \<le> card V` by simp
+      then have "r1 \<le> card ?M \<or> r2 \<le> card ?N" by arith
       moreover
       { assume "r1 \<le> card ?M"
         moreover have "finite ?M" using `finite V` by auto
@@ -82,7 +82,7 @@
           with `R <= V` have "?EX V E m n" by blast
         } moreover
         { assume "?C"
-          hence "clique (insert v R) E" using `R <= ?M`
+          then have "clique (insert v R) E" using `R <= ?M`
            by(auto simp:clique_def insert_commute)
           moreover have "card(insert v R) = m"
             using `?C` `finite R` `v ~: R` `m\<noteq>0` by simp
@@ -102,7 +102,7 @@
           with `R <= V` have "?EX V E m n" by blast
         } moreover
         { assume "?I"
-          hence "indep (insert v R) E" using `R <= ?N`
+          then have "indep (insert v R) E" using `R <= ?N`
             by(auto simp:indep_def insert_commute)
           moreover have "card(insert v R) = n"
             using `?I` `finite R` `v ~: R` `n\<noteq>0` by simp
@@ -124,17 +124,17 @@
     choice_0:   "choice P r 0 = (SOME x. P x)"
   | choice_Suc: "choice P r (Suc n) = (SOME y. P y & (choice P r n, y) \<in> r)"
 
-lemma choice_n: 
+lemma choice_n:
   assumes P0: "P x0"
       and Pstep: "!!x. P x ==> \<exists>y. P y & (x,y) \<in> r"
   shows "P (choice P r n)"
 proof (induct n)
-  case 0 show ?case by (force intro: someI P0) 
+  case 0 show ?case by (force intro: someI P0)
 next
-  case Suc thus ?case by (auto intro: someI2_ex [OF Pstep]) 
+  case Suc then show ?case by (auto intro: someI2_ex [OF Pstep])
 qed
 
-lemma dependent_choice: 
+lemma dependent_choice:
   assumes trans: "trans r"
       and P0: "P x0"
       and Pstep: "!!x. P x ==> \<exists>y. P y & (x,y) \<in> r"
@@ -144,7 +144,7 @@
   fix n
   show "P (choice P r n)" by (blast intro: choice_n [OF P0 Pstep])
 next
-  have PSuc: "\<forall>n. (choice P r n, choice P r (Suc n)) \<in> r" 
+  have PSuc: "\<forall>n. (choice P r n, choice P r (Suc n)) \<in> r"
     using Pstep [OF choice_n [OF P0 Pstep]]
     by (auto intro: someI2_ex)
   fix n m :: nat
@@ -156,8 +156,7 @@
 
 subsubsection {* Partitions of a Set *}
 
-definition
-  part :: "nat => nat => 'a set => ('a set => nat) => bool"
+definition part :: "nat => nat => 'a set => ('a set => nat) => bool"
   --{*the function @{term f} partitions the @{term r}-subsets of the typically
        infinite set @{term Y} into @{term s} distinct categories.*}
 where
@@ -165,52 +164,52 @@
 
 text{*For induction, we decrease the value of @{term r} in partitions.*}
 lemma part_Suc_imp_part:
-     "[| infinite Y; part (Suc r) s Y f; y \<in> Y |] 
+     "[| infinite Y; part (Suc r) s Y f; y \<in> Y |]
       ==> part r s (Y - {y}) (%u. f (insert y u))"
   apply(simp add: part_def, clarify)
   apply(drule_tac x="insert y X" in spec)
   apply(force)
   done
 
-lemma part_subset: "part r s YY f ==> Y \<subseteq> YY ==> part r s Y f" 
+lemma part_subset: "part r s YY f ==> Y \<subseteq> YY ==> part r s Y f"
   unfolding part_def by blast
-  
+
 
 subsection {* Ramsey's Theorem: Infinitary Version *}
 
-lemma Ramsey_induction: 
+lemma Ramsey_induction:
   fixes s and r::nat
   shows
-  "!!(YY::'a set) (f::'a set => nat). 
+  "!!(YY::'a set) (f::'a set => nat).
       [|infinite YY; part r s YY f|]
-      ==> \<exists>Y' t'. Y' \<subseteq> YY & infinite Y' & t' < s & 
+      ==> \<exists>Y' t'. Y' \<subseteq> YY & infinite Y' & t' < s &
                   (\<forall>X. X \<subseteq> Y' & finite X & card X = r --> f X = t')"
 proof (induct r)
   case 0
-  thus ?case by (auto simp add: part_def card_eq_0_iff cong: conj_cong)
+  then show ?case by (auto simp add: part_def card_eq_0_iff cong: conj_cong)
 next
-  case (Suc r) 
+  case (Suc r)
   show ?case
   proof -
     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).     
+    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)"
     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)
-    have transr: "trans ?ramr" by (force simp add: trans_def) 
+    have transr: "trans ?ramr" by (force simp add: trans_def)
     from Suc.hyps [OF infYY' partf']
     obtain Y0 and t0
     where "Y0 \<subseteq> YY - {yy}"  "infinite Y0"  "t0 < s"
           "\<forall>X. X\<subseteq>Y0 \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yy) X = t0"
-        by blast 
+        by blast
     with yy have propr0: "?propr(yy,Y0,t0)" by blast
-    have proprstep: "\<And>x. ?propr x \<Longrightarrow> \<exists>y. ?propr y \<and> (x, y) \<in> ?ramr" 
+    have proprstep: "\<And>x. ?propr x \<Longrightarrow> \<exists>y. ?propr y \<and> (x, y) \<in> ?ramr"
     proof -
       fix x
-      assume px: "?propr x" thus "?thesis x"
+      assume px: "?propr x" then show "?thesis x"
       proof (cases x)
         case (fields yx Yx tx)
         then obtain yx' where yx': "yx' \<in> Yx" using px
@@ -223,7 +222,7 @@
         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 
+            by blast
         show ?thesis
         proof
           show "?propr (yx',Y',t') & (x, (yx',Y',t')) \<in> ?ramr"
@@ -258,51 +257,51 @@
     show ?thesis
     proof (intro exI conjI)
       show "?gy ` {n. ?gt n = s'} \<subseteq> YY" using pg
-        by (auto simp add: Let_def split_beta) 
+        by (auto simp add: Let_def split_beta)
       show "infinite (?gy ` {n. ?gt n = s'})" using infeqs'
-        by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD) 
+        by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD)
       show "s' < s" by (rule less')
-      show "\<forall>X. X \<subseteq> ?gy ` {n. ?gt n = s'} & finite X & card X = Suc r 
+      show "\<forall>X. X \<subseteq> ?gy ` {n. ?gt n = s'} & finite X & card X = Suc r
           --> f X = s'"
       proof -
-        {fix X 
+        {fix X
          assume "X \<subseteq> ?gy ` {n. ?gt n = s'}"
             and cardX: "finite X" "card X = Suc r"
-         then obtain AA where AA: "AA \<subseteq> {n. ?gt n = s'}" and Xeq: "X = ?gy`AA" 
-             by (auto simp add: subset_image_iff) 
+         then obtain AA where AA: "AA \<subseteq> {n. ?gt n = s'}" and Xeq: "X = ?gy`AA"
+             by (auto simp add: subset_image_iff)
          with cardX have "AA\<noteq>{}" by auto
-         hence AAleast: "(LEAST x. x \<in> AA) \<in> AA" by (auto intro: LeastI_ex) 
+         then have AAleast: "(LEAST x. x \<in> AA) \<in> AA" by (auto intro: LeastI_ex)
          have "f X = s'"
-         proof (cases "g (LEAST x. x \<in> AA)") 
+         proof (cases "g (LEAST x. x \<in> AA)")
            case (fields ya Ya ta)
-           with AAleast Xeq 
-           have ya: "ya \<in> X" by (force intro!: rev_image_eqI) 
-           hence "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb)
-           also have "... = ta" 
+           with AAleast Xeq
+           have ya: "ya \<in> X" by (force intro!: rev_image_eqI)
+           then have "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb)
+           also have "... = ta"
            proof -
              have "X - {ya} \<subseteq> Ya"
-             proof 
+             proof
                fix x assume x: "x \<in> X - {ya}"
-               then obtain a' where xeq: "x = ?gy a'" and a': "a' \<in> AA" 
-                 by (auto simp add: Xeq) 
-               hence "a' \<noteq> (LEAST x. x \<in> AA)" using x fields by auto
-               hence lessa': "(LEAST x. x \<in> AA) < a'"
+               then obtain a' where xeq: "x = ?gy a'" and a': "a' \<in> AA"
+                 by (auto simp add: Xeq)
+               then have "a' \<noteq> (LEAST x. x \<in> AA)" using x fields by auto
+               then have lessa': "(LEAST x. x \<in> AA) < a'"
                  using Least_le [of "%x. x \<in> AA", OF a'] by arith
                show "x \<in> Ya" using xeq fields rg [OF lessa'] by auto
              qed
              moreover
              have "card (X - {ya}) = r"
                by (simp add: cardX ya)
-             ultimately show ?thesis 
+             ultimately show ?thesis
                using pg [of "LEAST x. x \<in> AA"] fields cardX
                by (clarsimp simp del:insert_Diff_single)
            qed
            also have "... = s'" using AA AAleast fields by auto
            finally show ?thesis .
          qed}
-        thus ?thesis by blast
-      qed 
-    qed 
+        then show ?thesis by blast
+      qed
+    qed
   qed
 qed
 
@@ -312,7 +311,7 @@
   shows
    "[|infinite Z;
       \<forall>X. X \<subseteq> Z & finite X & card X = r --> f X < s|]
-  ==> \<exists>Y t. Y \<subseteq> Z & infinite Y & t < s 
+  ==> \<exists>Y t. Y \<subseteq> Z & infinite Y & t < s
             & (\<forall>X. X \<subseteq> Y & finite X & card X = r --> f X = t)"
 by (blast intro: Ramsey_induction [unfolded part_def])
 
@@ -326,7 +325,7 @@
 proof -
   have part2: "\<forall>X. X \<subseteq> Z & finite X & card X = 2 --> f X < s"
     using part by (fastforce simp add: eval_nat_numeral card_Suc_eq)
-  obtain Y t 
+  obtain Y t
     where "Y \<subseteq> Z" "infinite Y" "t < s"
           "(\<forall>X. X \<subseteq> Y & finite X & card X = 2 --> f X = t)"
     by (insert Ramsey [OF infZ part2]) auto
@@ -342,39 +341,36 @@
   \cite{Podelski-Rybalchenko}.
 *}
 
-definition
-  disj_wf         :: "('a * 'a)set => bool"
-where
-  "disj_wf r = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r = (\<Union>i<n. T i))"
+definition disj_wf :: "('a * 'a)set => bool"
+  where "disj_wf r = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r = (\<Union>i<n. T i))"
 
-definition
-  transition_idx :: "[nat => 'a, nat => ('a*'a)set, nat set] => nat"
-where
-  "transition_idx s T A =
-    (LEAST k. \<exists>i j. A = {i,j} & i<j & (s j, s i) \<in> T k)"
+definition transition_idx :: "[nat => 'a, nat => ('a*'a)set, nat set] => nat"
+  where
+    "transition_idx s T A =
+      (LEAST k. \<exists>i j. A = {i,j} & i<j & (s j, s i) \<in> T k)"
 
 
 lemma transition_idx_less:
     "[|i<j; (s j, s i) \<in> T k; k<n|] ==> transition_idx s T {i,j} < n"
-apply (subgoal_tac "transition_idx s T {i, j} \<le> k", simp) 
-apply (simp add: transition_idx_def, blast intro: Least_le) 
+apply (subgoal_tac "transition_idx s T {i, j} \<le> k", simp)
+apply (simp add: transition_idx_def, blast intro: Least_le)
 done
 
 lemma transition_idx_in:
     "[|i<j; (s j, s i) \<in> T k|] ==> (s j, s i) \<in> T (transition_idx s T {i,j})"
-apply (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR 
-            cong: conj_cong) 
-apply (erule LeastI) 
+apply (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR
+            cong: conj_cong)
+apply (erule LeastI)
 done
 
 text{*To be equal to the union of some well-founded relations is equivalent
 to being the subset of such a union.*}
 lemma disj_wf:
      "disj_wf(r) = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r \<subseteq> (\<Union>i<n. T i))"
-apply (auto simp add: disj_wf_def) 
-apply (rule_tac x="%i. T i Int r" in exI) 
-apply (rule_tac x=n in exI) 
-apply (force simp add: wf_Int1) 
+apply (auto simp add: disj_wf_def)
+apply (rule_tac x="%i. T i Int r" in exI)
+apply (rule_tac x=n in exI)
+apply (force simp add: wf_Int1)
 done
 
 theorem trans_disj_wf_implies_wf:
@@ -388,13 +384,13 @@
   proof -
     fix i and j::nat
     assume less: "i<j"
-    thus "(s j, s i) \<in> r"
+    then show "(s j, s i) \<in> r"
     proof (rule less_Suc_induct)
-      show "\<And>i. (s (Suc i), s i) \<in> r" by (simp add: sSuc) 
+      show "\<And>i. (s (Suc i), s i) \<in> r" by (simp add: sSuc)
       show "\<And>i j k. \<lbrakk>(s j, s i) \<in> r; (s k, s j) \<in> r\<rbrakk> \<Longrightarrow> (s k, s i) \<in> r"
-        using transr by (unfold trans_def, blast) 
+        using transr by (unfold trans_def, blast)
     qed
-  qed    
+  qed
   from dwf
   obtain T and n::nat where wfT: "\<forall>k<n. wf(T k)" and r: "r = (\<Union>k<n. T k)"
     by (auto simp add: disj_wf_def)
@@ -402,20 +398,20 @@
   proof -
     fix i and j::nat
     assume less: "i<j"
-    hence "(s j, s i) \<in> r" by (rule s [of i j]) 
-    thus "\<exists>k. (s j, s i) \<in> T k & k<n" by (auto simp add: r)
-  qed    
+    then have "(s j, s i) \<in> r" by (rule s [of i j])
+    then show "\<exists>k. (s j, s i) \<in> T k & k<n" by (auto simp add: r)
+  qed
   have trless: "!!i j. i\<noteq>j ==> transition_idx s T {i,j} < n"
     apply (auto simp add: linorder_neq_iff)
-    apply (blast dest: s_in_T transition_idx_less) 
-    apply (subst insert_commute)   
-    apply (blast dest: s_in_T transition_idx_less) 
+    apply (blast dest: s_in_T transition_idx_less)
+    apply (subst insert_commute)
+    apply (blast dest: s_in_T transition_idx_less)
     done
   have
-   "\<exists>K k. K \<subseteq> UNIV & infinite K & k < n & 
+   "\<exists>K k. K \<subseteq> UNIV & infinite K & k < n &
           (\<forall>i\<in>K. \<forall>j\<in>K. i\<noteq>j --> transition_idx s T {i,j} = k)"
-    by (rule Ramsey2) (auto intro: trless nat_infinite) 
-  then obtain K and k 
+    by (rule Ramsey2) (auto intro: trless nat_infinite)
+  then obtain K and k
     where infK: "infinite K" and less: "k < n" and
           allk: "\<forall>i\<in>K. \<forall>j\<in>K. i\<noteq>j --> transition_idx s T {i,j} = k"
     by auto
@@ -424,18 +420,18 @@
     fix m::nat
     let ?j = "enumerate K (Suc m)"
     let ?i = "enumerate K m"
-    have jK: "?j \<in> K" by (simp add: enumerate_in_set infK) 
-    have iK: "?i \<in> K" by (simp add: enumerate_in_set infK) 
-    have ij: "?i < ?j" by (simp add: enumerate_step infK) 
-    have ijk: "transition_idx s T {?i,?j} = k" using iK jK ij 
+    have jK: "?j \<in> K" by (simp add: enumerate_in_set infK)
+    have iK: "?i \<in> K" by (simp add: enumerate_in_set infK)
+    have ij: "?i < ?j" by (simp add: enumerate_step infK)
+    have ijk: "transition_idx s T {?i,?j} = k" using iK jK ij
       by (simp add: allk)
-    obtain k' where "(s ?j, s ?i) \<in> T k'" "k'<n" 
+    obtain k' where "(s ?j, s ?i) \<in> T k'" "k'<n"
       using s_in_T [OF ij] by blast
-    thus "(s ?j, s ?i) \<in> T k" 
-      by (simp add: ijk [symmetric] transition_idx_in ij) 
+    then show "(s ?j, s ?i) \<in> T k"
+      by (simp add: ijk [symmetric] transition_idx_in ij)
   qed
-  hence "~ wf(T k)" by (force simp add: wf_iff_no_infinite_down_chain) 
-  thus False using wfT less by blast
+  then have "~ wf(T k)" by (force simp add: wf_iff_no_infinite_down_chain)
+  then show False using wfT less by blast
 qed
 
 end