de-applying (mostly Set_Interval)
authorpaulson <lp15@cam.ac.uk>
Thu, 12 Jul 2018 17:22:39 +0100
changeset 68618 3db8520941a4
parent 68616 cedf3480fdad
child 68619 79abf4201e8d
de-applying (mostly Set_Interval)
src/HOL/Nat.thy
src/HOL/Set_Interval.thy
src/HOL/Transitive_Closure.thy
--- a/src/HOL/Nat.thy	Wed Jul 11 23:24:25 2018 +0100
+++ b/src/HOL/Nat.thy	Thu Jul 12 17:22:39 2018 +0100
@@ -2093,7 +2093,7 @@
 
 lemma inj_on_diff_nat:
   fixes k :: nat
-  assumes "\<forall>n \<in> N. k \<le> n"
+  assumes "\<And>n. n \<in> N \<Longrightarrow> k \<le> n"
   shows "inj_on (\<lambda>n. n - k) N"
 proof (rule inj_onI)
   fix x y
--- a/src/HOL/Set_Interval.thy	Wed Jul 11 23:24:25 2018 +0100
+++ b/src/HOL/Set_Interval.thy	Thu Jul 12 17:22:39 2018 +0100
@@ -91,29 +91,27 @@
 
 lemma Compl_lessThan [simp]:
     "!!k:: 'a::linorder. -lessThan k = atLeast k"
-apply (auto simp add: lessThan_def atLeast_def)
-done
+  by (auto simp add: lessThan_def atLeast_def)
 
 lemma single_Diff_lessThan [simp]: "!!k:: 'a::order. {k} - lessThan k = {k}"
-by auto
+  by auto
 
 lemma (in ord) greaterThan_iff [iff]: "(i \<in> greaterThan k) = (k<i)"
-by (simp add: greaterThan_def)
+  by (simp add: greaterThan_def)
 
 lemma Compl_greaterThan [simp]:
     "!!k:: 'a::linorder. -greaterThan k = atMost k"
   by (auto simp add: greaterThan_def atMost_def)
 
 lemma Compl_atMost [simp]: "!!k:: 'a::linorder. -atMost k = greaterThan k"
-apply (subst Compl_greaterThan [symmetric])
-apply (rule double_complement)
-done
+  apply (subst Compl_greaterThan [symmetric])
+  apply (rule double_complement)
+  done
 
 lemma (in ord) atLeast_iff [iff]: "(i \<in> atLeast k) = (k<=i)"
 by (simp add: atLeast_def)
 
-lemma Compl_atLeast [simp]:
-    "!!k:: 'a::linorder. -atLeast k = lessThan k"
+lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k"
   by (auto simp add: lessThan_def atLeast_def)
 
 lemma (in ord) atMost_iff [iff]: "(i \<in> atMost k) = (i<=k)"
@@ -146,35 +144,25 @@
 
 lemma greaterThan_subset_iff [iff]:
      "(greaterThan x \<subseteq> greaterThan y) = (y \<le> (x::'a::linorder))"
-apply (auto simp add: greaterThan_def)
- apply (subst linorder_not_less [symmetric], blast)
-done
+  unfolding greaterThan_def by (auto simp: linorder_not_less [symmetric])
 
 lemma greaterThan_eq_iff [iff]:
      "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))"
-apply (rule iffI)
- apply (erule equalityE)
- apply simp_all
-done
+  by (auto simp: elim!: equalityE)
 
 lemma atMost_subset_iff [iff]: "(atMost x \<subseteq> atMost y) = (x \<le> (y::'a::order))"
-by (blast intro: order_trans)
+  by (blast intro: order_trans)
 
 lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::linorder))"
-by (blast intro: order_antisym order_trans)
+  by (blast intro: order_antisym order_trans)
 
 lemma lessThan_subset_iff [iff]:
      "(lessThan x \<subseteq> lessThan y) = (x \<le> (y::'a::linorder))"
-apply (auto simp add: lessThan_def)
- apply (subst linorder_not_less [symmetric], blast)
-done
+  unfolding lessThan_def by (auto simp: linorder_not_less [symmetric])
 
 lemma lessThan_eq_iff [iff]:
      "(lessThan x = lessThan y) = (x = (y::'a::linorder))"
-apply (rule iffI)
- apply (erule equalityE)
- apply simp_all
-done
+  by (auto simp: elim!: equalityE)
 
 lemma lessThan_strict_subset_iff:
   fixes m n :: "'a::linorder"
@@ -195,21 +183,17 @@
 context ord
 begin
 
-lemma greaterThanLessThan_iff [simp]:
-  "(i \<in> {l<..<u}) = (l < i \<and> i < u)"
-by (simp add: greaterThanLessThan_def)
-
-lemma atLeastLessThan_iff [simp]:
-  "(i \<in> {l..<u}) = (l \<le> i \<and> i < u)"
-by (simp add: atLeastLessThan_def)
-
-lemma greaterThanAtMost_iff [simp]:
-  "(i \<in> {l<..u}) = (l < i \<and> i \<le> u)"
-by (simp add: greaterThanAtMost_def)
-
-lemma atLeastAtMost_iff [simp]:
-  "(i \<in> {l..u}) = (l \<le> i \<and> i \<le> u)"
-by (simp add: atLeastAtMost_def)
+lemma greaterThanLessThan_iff [simp]: "(i \<in> {l<..<u}) = (l < i \<and> i < u)"
+  by (simp add: greaterThanLessThan_def)
+
+lemma atLeastLessThan_iff [simp]: "(i \<in> {l..<u}) = (l \<le> i \<and> i < u)"
+  by (simp add: atLeastLessThan_def)
+
+lemma greaterThanAtMost_iff [simp]: "(i \<in> {l<..u}) = (l < i \<and> i \<le> u)"
+  by (simp add: greaterThanAtMost_def)
+
+lemma atLeastAtMost_iff [simp]: "(i \<in> {l..u}) = (l \<le> i \<and> i \<le> u)"
+  by (simp add: atLeastAtMost_def)
 
 text \<open>The above four lemmas could be declared as iffs. Unfortunately this
 breaks many proofs. Since it only helps blast, it is better to leave them
@@ -467,11 +451,11 @@
 
 lemma (in linorder) atLeastLessThan_subset_iff:
   "{a..<b} \<subseteq> {c..<d} \<Longrightarrow> b \<le> a \<or> c\<le>a \<and> b\<le>d"
-apply (auto simp:subset_eq Ball_def)
-apply(frule_tac x=a in spec)
-apply(erule_tac x=d in allE)
-apply (simp add: less_imp_le)
-done
+  apply (auto simp:subset_eq Ball_def not_le)
+  apply(frule_tac x=a in spec)
+  apply(erule_tac x=d in allE)
+  apply (auto simp: )
+  done
 
 lemma atLeastLessThan_inj:
   fixes a b c d :: "'a::linorder"
@@ -718,17 +702,15 @@
 subsubsection \<open>The Constant @{term greaterThan}\<close>
 
 lemma greaterThan_0: "greaterThan 0 = range Suc"
-apply (simp add: greaterThan_def)
-apply (blast dest: gr0_conv_Suc [THEN iffD1])
-done
+  unfolding greaterThan_def
+  by (blast dest: gr0_conv_Suc [THEN iffD1])
 
 lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}"
-apply (simp add: greaterThan_def)
-apply (auto elim: linorder_neqE)
-done
+  unfolding greaterThan_def
+  by (auto elim: linorder_neqE)
 
 lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}"
-by blast
+  by blast
 
 subsubsection \<open>The Constant @{term atLeast}\<close>
 
@@ -736,29 +718,24 @@
 by (unfold atLeast_def UNIV_def, simp)
 
 lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}"
-apply (simp add: atLeast_def)
-apply (simp add: Suc_le_eq)
-apply (simp add: order_le_less, blast)
-done
+  unfolding atLeast_def by (auto simp: order_le_less Suc_le_eq)
 
 lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k"
   by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le)
 
 lemma UN_atLeast_UNIV: "(UN m::nat. atLeast m) = UNIV"
-by blast
+  by blast
 
 subsubsection \<open>The Constant @{term atMost}\<close>
 
 lemma atMost_0 [simp]: "atMost (0::nat) = {0}"
-by (simp add: atMost_def)
+  by (simp add: atMost_def)
 
 lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)"
-apply (simp add: atMost_def)
-apply (simp add: less_Suc_eq order_le_less, blast)
-done
+  unfolding atMost_def by (auto simp add: less_Suc_eq order_le_less)
 
 lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV"
-by blast
+  by blast
 
 subsubsection \<open>The Constant @{term atLeastLessThan}\<close>
 
@@ -770,28 +747,24 @@
 specific concept to a more general one.\<close>
 
 lemma atLeast0LessThan [code_abbrev]: "{0::nat..<n} = {..<n}"
-by(simp add:lessThan_def atLeastLessThan_def)
+  by(simp add:lessThan_def atLeastLessThan_def)
 
 lemma atLeast0AtMost [code_abbrev]: "{0..n::nat} = {..n}"
-by(simp add:atMost_def atLeastAtMost_def)
-
-lemma lessThan_atLeast0:
-  "{..<n} = {0::nat..<n}"
+  by(simp add:atMost_def atLeastAtMost_def)
+
+lemma lessThan_atLeast0: "{..<n} = {0::nat..<n}"
   by (simp add: atLeast0LessThan)
 
-lemma atMost_atLeast0:
-  "{..n} = {0::nat..n}"
+lemma atMost_atLeast0: "{..n} = {0::nat..n}"
   by (simp add: atLeast0AtMost)
 
 lemma atLeastLessThan0: "{m..<0::nat} = {}"
-by (simp add: atLeastLessThan_def)
-
-lemma atLeast0_lessThan_Suc:
-  "{0..<Suc n} = insert n {0..<n}"
+  by (simp add: atLeastLessThan_def)
+
+lemma atLeast0_lessThan_Suc: "{0..<Suc n} = insert n {0..<n}"
   by (simp add: atLeast0LessThan lessThan_Suc)
 
-lemma atLeast0_lessThan_Suc_eq_insert_0:
-  "{0..<Suc n} = insert 0 (Suc ` {0..<n})"
+lemma atLeast0_lessThan_Suc_eq_insert_0: "{0..<Suc n} = insert 0 (Suc ` {0..<n})"
   by (simp add: atLeast0LessThan lessThan_Suc_eq_insert_0)
 
 
@@ -815,19 +788,13 @@
 
 lemma atLeastLessThan_singleton [simp]: "{m..<Suc m} = {m}"
 by (auto simp add: atLeastLessThan_def)
-(*
-lemma atLeast_sum_LessThan [simp]: "{m + k..<k::nat} = {}"
-by (induct k, simp_all add: atLeastLessThanSuc)
-
-lemma atLeastSucLessThan [simp]: "{Suc n..<n} = {}"
-by (auto simp add: atLeastLessThan_def)
-*)
+
 lemma atLeastLessThanSuc_atLeastAtMost: "{l..<Suc u} = {l..u}"
   by (simp add: lessThan_Suc_atMost atLeastAtMost_def atLeastLessThan_def)
 
 lemma atLeastSucAtMost_greaterThanAtMost: "{Suc l..u} = {l<..u}"
   by (simp add: atLeast_Suc_greaterThan atLeastAtMost_def
-    greaterThanAtMost_def)
+      greaterThanAtMost_def)
 
 lemma atLeastSucLessThan_greaterThanLessThan: "{Suc l..<u} = {l<..<u}"
   by (simp add: atLeast_Suc_greaterThan atLeastLessThan_def
@@ -846,9 +813,7 @@
   by (auto intro: set_eqI)
 
 lemma atLeastLessThan_add_Un: "i \<le> j \<Longrightarrow> {i..<j+k} = {i..<j} \<union> {j..<j+k::nat}"
-  apply (induct k)
-  apply (simp_all add: atLeastLessThanSuc)
-  done
+  by (induct k) (simp_all add: atLeastLessThanSuc)
 
 
 subsubsection \<open>Intervals and numerals\<close>
@@ -1123,7 +1088,7 @@
   by auto
 
 lemma image_add_int_atLeastLessThan:
-    "(%x. x + (l::int)) ` {0..<u-l} = {l..<u}"
+    "(\<lambda>x. x + (l::int)) ` {0..<u-l} = {l..<u}"
   apply (auto simp add: image_def)
   apply (rule_tac x = "x - l" in bexI)
   apply auto
@@ -1165,25 +1130,23 @@
 
 lemma finite_greaterThanLessThan [iff]:
   fixes l :: nat shows "finite {l<..<u}"
-by (simp add: greaterThanLessThan_def)
+  by (simp add: greaterThanLessThan_def)
 
 lemma finite_atLeastLessThan [iff]:
   fixes l :: nat shows "finite {l..<u}"
-by (simp add: atLeastLessThan_def)
+  by (simp add: atLeastLessThan_def)
 
 lemma finite_greaterThanAtMost [iff]:
   fixes l :: nat shows "finite {l<..u}"
-by (simp add: greaterThanAtMost_def)
+  by (simp add: greaterThanAtMost_def)
 
 lemma finite_atLeastAtMost [iff]:
   fixes l :: nat shows "finite {l..u}"
-by (simp add: atLeastAtMost_def)
+  by (simp add: atLeastAtMost_def)
 
 text \<open>A bounded set of natural numbers is finite.\<close>
 lemma bounded_nat_set_is_finite: "(\<forall>i\<in>N. i < (n::nat)) \<Longrightarrow> finite N"
-apply (rule finite_subset)
- apply (rule_tac [2] finite_lessThan, auto)
-done
+  by (rule finite_subset [OF _ finite_lessThan]) auto
 
 text \<open>A set of natural numbers is finite iff it is bounded.\<close>
 lemma finite_nat_set_iff_bounded:
@@ -1195,11 +1158,9 @@
   assume ?B show ?F using \<open>?B\<close> by(blast intro:bounded_nat_set_is_finite)
 qed
 
-lemma finite_nat_set_iff_bounded_le:
-  "finite(N::nat set) = (\<exists>m. \<forall>n\<in>N. n<=m)"
-apply(simp add:finite_nat_set_iff_bounded)
-apply(blast dest:less_imp_le_nat le_imp_less_Suc)
-done
+lemma finite_nat_set_iff_bounded_le: "finite(N::nat set) = (\<exists>m. \<forall>n\<in>N. n\<le>m)"
+  unfolding finite_nat_set_iff_bounded
+  by (blast dest:less_imp_le_nat le_imp_less_Suc)
 
 lemma finite_less_ub:
      "!!f::nat=>nat. (!!n. n \<le> f n) ==> finite {n. f n \<le> u}"
@@ -1298,10 +1259,9 @@
 lemma UN_finite2_eq:
   "(\<And>n::nat. (\<Union>i\<in>{0..<n}. A i) = (\<Union>i\<in>{0..<n + k}. B i)) \<Longrightarrow>
     (\<Union>n. A n) = (\<Union>n. B n)"
-  apply (rule subset_antisym)
-   apply (rule UN_finite2_subset, blast)
-  apply (rule UN_finite2_subset [where k=k])
-  apply (force simp add: atLeastLessThan_add_Un [of 0])
+  apply (rule subset_antisym [OF UN_finite_subset UN_finite2_subset])
+   apply auto
+  apply (force simp add: atLeastLessThan_add_Un [of 0])+
   done
 
 
@@ -1315,7 +1275,7 @@
 
 lemma card_atLeastLessThan [simp]: "card {l..<u} = u - l"
 proof -
-  have "{l..<u} = (%x. x + l) ` {..<u-l}"
+  have "{l..<u} = (\<lambda>x. x + l) ` {..<u-l}"
     apply (auto simp add: image_def atLeastLessThan_def lessThan_def)
     apply (rule_tac x = "x - l" in exI)
     apply arith
@@ -1349,24 +1309,24 @@
 
 lemma ex_bij_betw_nat_finite:
   "finite M \<Longrightarrow> \<exists>h. bij_betw h {0..<card M} M"
-apply(drule finite_imp_nat_seg_image_inj_on)
-apply(auto simp:atLeast0LessThan[symmetric] lessThan_def[symmetric] card_image bij_betw_def)
-done
+  apply(drule finite_imp_nat_seg_image_inj_on)
+  apply(auto simp:atLeast0LessThan[symmetric] lessThan_def[symmetric] card_image bij_betw_def)
+  done
 
 lemma ex_bij_betw_finite_nat:
   "finite M \<Longrightarrow> \<exists>h. bij_betw h M {0..<card M}"
-by (blast dest: ex_bij_betw_nat_finite bij_betw_inv)
+  by (blast dest: ex_bij_betw_nat_finite bij_betw_inv)
 
 lemma finite_same_card_bij:
   "finite A \<Longrightarrow> finite B \<Longrightarrow> card A = card B \<Longrightarrow> \<exists>h. bij_betw h A B"
-apply(drule ex_bij_betw_finite_nat)
-apply(drule ex_bij_betw_nat_finite)
-apply(auto intro!:bij_betw_trans)
-done
+  apply(drule ex_bij_betw_finite_nat)
+  apply(drule ex_bij_betw_nat_finite)
+  apply(auto intro!:bij_betw_trans)
+  done
 
 lemma ex_bij_betw_nat_finite_1:
   "finite M \<Longrightarrow> \<exists>h. bij_betw h {1 .. card M} M"
-by (rule finite_same_card_bij) auto
+  by (rule finite_same_card_bij) auto
 
 lemma bij_betw_iff_card:
   assumes "finite A" "finite B"
@@ -1428,26 +1388,21 @@
 
 lemma image_atLeastZeroLessThan_int: "0 \<le> u ==>
     {(0::int)..<u} = int ` {..<nat u}"
-  apply (unfold image_def lessThan_def)
+  unfolding image_def lessThan_def
   apply auto
   apply (rule_tac x = "nat x" in exI)
   apply (auto simp add: zless_nat_eq_int_zless [THEN sym])
   done
 
 lemma finite_atLeastZeroLessThan_int: "finite {(0::int)..<u}"
-  apply (cases "0 \<le> u")
-  apply (subst image_atLeastZeroLessThan_int, assumption)
-  apply (rule finite_imageI)
-  apply auto
-  done
+proof (cases "0 \<le> u")
+  case True
+  then show ?thesis
+    by (auto simp: image_atLeastZeroLessThan_int)
+qed auto
 
 lemma finite_atLeastLessThan_int [iff]: "finite {l..<u::int}"
-  apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}")
-  apply (erule subst)
-  apply (rule finite_imageI)
-  apply (rule finite_atLeastZeroLessThan_int)
-  apply (rule image_add_int_atLeastLessThan)
-  done
+  by (simp only: image_add_int_atLeastLessThan [symmetric, of l] finite_imageI finite_atLeastZeroLessThan_int)
 
 lemma finite_atLeastAtMost_int [iff]: "finite {l..(u::int)}"
   by (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym], simp)
@@ -1462,32 +1417,33 @@
 subsubsection \<open>Cardinality\<close>
 
 lemma card_atLeastZeroLessThan_int: "card {(0::int)..<u} = nat u"
-  apply (cases "0 \<le> u")
-  apply (subst image_atLeastZeroLessThan_int, assumption)
-  apply (subst card_image)
-  apply (auto simp add: inj_on_def)
-  done
+proof (cases "0 \<le> u")
+  case True
+  then show ?thesis
+    by (auto simp: image_atLeastZeroLessThan_int card_image inj_on_def)    
+qed auto
 
 lemma card_atLeastLessThan_int [simp]: "card {l..<u} = nat (u - l)"
-  apply (subgoal_tac "card {l..<u} = card {0..<u-l}")
-  apply (erule ssubst, rule card_atLeastZeroLessThan_int)
-  apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}")
-  apply (erule subst)
-  apply (rule card_image)
-  apply (simp add: inj_on_def)
-  apply (rule image_add_int_atLeastLessThan)
-  done
+proof -
+  have "card {l..<u} = card {0..<u-l}"
+    apply (subst image_add_int_atLeastLessThan [symmetric])
+    apply (rule card_image)
+    apply (simp add: inj_on_def)
+    done
+  then show ?thesis
+    by (simp add: card_atLeastZeroLessThan_int)
+qed
 
 lemma card_atLeastAtMost_int [simp]: "card {l..u} = nat (u - l + 1)"
-apply (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym])
-apply (auto simp add: algebra_simps)
-done
+  apply (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym])
+  apply (auto simp add: algebra_simps)
+  done
 
 lemma card_greaterThanAtMost_int [simp]: "card {l<..u} = nat (u - l)"
-by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
+  by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
 
 lemma card_greaterThanLessThan_int [simp]: "card {l<..<u} = nat (u - (l + 1))"
-by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
+  by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
 
 lemma finite_M_bounded_by_nat: "finite {k. P k \<and> k < (i::nat)}"
 proof -
@@ -1496,43 +1452,44 @@
 qed
 
 lemma card_less:
-assumes zero_in_M: "0 \<in> M"
-shows "card {k \<in> M. k < Suc i} \<noteq> 0"
+  assumes zero_in_M: "0 \<in> M"
+  shows "card {k \<in> M. k < Suc i} \<noteq> 0"
 proof -
   from zero_in_M have "{k \<in> M. k < Suc i} \<noteq> {}" by auto
   with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff)
 qed
 
-lemma card_less_Suc2: "0 \<notin> M \<Longrightarrow> card {k. Suc k \<in> M \<and> k < i} = card {k \<in> M. k < Suc i}"
-apply (rule card_bij_eq [of Suc _ _ "\<lambda>x. x - Suc 0"])
-apply auto
-apply (rule inj_on_diff_nat)
-apply auto
-apply (case_tac x)
-apply auto
-apply (case_tac xa)
-apply auto
-apply (case_tac xa)
-apply auto
-done
+lemma card_less_Suc2: 
+  assumes "0 \<notin> M" shows "card {k. Suc k \<in> M \<and> k < i} = card {k \<in> M. k < Suc i}"
+proof -
+  have *: "\<lbrakk>j \<in> M; j < Suc i\<rbrakk> \<Longrightarrow> j - Suc 0 < i \<and> Suc (j - Suc 0) \<in> M \<and> Suc 0 \<le> j" for j
+    by (cases j) (use assms in auto)
+  show ?thesis
+  proof (rule card_bij_eq)
+    show "inj_on Suc {k. Suc k \<in> M \<and> k < i}"
+      by force
+    show "inj_on (\<lambda>x. x - Suc 0) {k \<in> M. k < Suc i}"
+      by (rule inj_on_diff_nat) (use * in blast)
+  qed (use * in auto)
+qed
 
 lemma card_less_Suc:
-  assumes zero_in_M: "0 \<in> M"
+  assumes "0 \<in> M"
     shows "Suc (card {k. Suc k \<in> M \<and> k < i}) = card {k \<in> M. k < Suc i}"
 proof -
-  from assms have a: "0 \<in> {k \<in> M. k < Suc i}" by simp
-  hence c: "{k \<in> M. k < Suc i} = insert 0 ({k \<in> M. k < Suc i} - {0})"
-    by (auto simp only: insert_Diff)
-  have b: "{k \<in> M. k < Suc i} - {0} = {k \<in> M - {0}. k < Suc i}"  by auto
-  from finite_M_bounded_by_nat[of "\<lambda>x. x \<in> M" "Suc i"]
-  have "Suc (card {k. Suc k \<in> M \<and> k < i}) = card (insert 0 ({k \<in> M. k < Suc i} - {0}))"
-    apply (subst card_insert)
-    apply simp_all
-    apply (subst b)
-    apply (subst card_less_Suc2[symmetric])
-    apply simp_all
-    done
-  with c show ?thesis by simp
+  have "Suc (card {k. Suc k \<in> M \<and> k < i}) = Suc (card {k. Suc k \<in> M - {0} \<and> k < i})"
+    by simp
+  also have "\<dots> = Suc (card {k \<in> M - {0}. k < Suc i})"
+    apply (subst card_less_Suc2)
+    using assms by auto
+  also have "\<dots> = Suc (card ({k \<in> M. k < Suc i} - {0}))"
+    by (force intro: arg_cong [where f=card])
+  also have "\<dots> = card (insert 0 ({k \<in> M. k < Suc i} - {0}))"
+    by (simp add: card_insert)
+  also have "... = card {k \<in> M. k < Suc i}"
+    using assms
+    by (force simp add: intro: arg_cong [where f=card])
+  finally show ?thesis.
 qed
 
 
@@ -1549,41 +1506,41 @@
   "{..<u} Un {u::'a::linorder} = {..u}"
   "(l::'a::linorder) < u ==> {l} Un {l<..<u} = {l..<u}"
   "(l::'a::linorder) < u ==> {l<..<u} Un {u} = {l<..u}"
-  "(l::'a::linorder) <= u ==> {l} Un {l<..u} = {l..u}"
-  "(l::'a::linorder) <= u ==> {l..<u} Un {u} = {l..u}"
+  "(l::'a::linorder) \<le> u ==> {l} Un {l<..u} = {l..u}"
+  "(l::'a::linorder) \<le> u ==> {l..<u} Un {u} = {l..u}"
 by auto
 
 text \<open>One- and two-sided intervals\<close>
 
 lemma ivl_disj_un_one:
   "(l::'a::linorder) < u ==> {..l} Un {l<..<u} = {..<u}"
-  "(l::'a::linorder) <= u ==> {..<l} Un {l..<u} = {..<u}"
-  "(l::'a::linorder) <= u ==> {..l} Un {l<..u} = {..u}"
-  "(l::'a::linorder) <= u ==> {..<l} Un {l..u} = {..u}"
-  "(l::'a::linorder) <= u ==> {l<..u} Un {u<..} = {l<..}"
+  "(l::'a::linorder) \<le> u ==> {..<l} Un {l..<u} = {..<u}"
+  "(l::'a::linorder) \<le> u ==> {..l} Un {l<..u} = {..u}"
+  "(l::'a::linorder) \<le> u ==> {..<l} Un {l..u} = {..u}"
+  "(l::'a::linorder) \<le> u ==> {l<..u} Un {u<..} = {l<..}"
   "(l::'a::linorder) < u ==> {l<..<u} Un {u..} = {l<..}"
-  "(l::'a::linorder) <= u ==> {l..u} Un {u<..} = {l..}"
-  "(l::'a::linorder) <= u ==> {l..<u} Un {u..} = {l..}"
+  "(l::'a::linorder) \<le> u ==> {l..u} Un {u<..} = {l..}"
+  "(l::'a::linorder) \<le> u ==> {l..<u} Un {u..} = {l..}"
 by auto
 
 text \<open>Two- and two-sided intervals\<close>
 
 lemma ivl_disj_un_two:
-  "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..<u} = {l<..<u}"
-  "[| (l::'a::linorder) <= m; m < u |] ==> {l<..m} Un {m<..<u} = {l<..<u}"
-  "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..<u} = {l..<u}"
-  "[| (l::'a::linorder) <= m; m < u |] ==> {l..m} Un {m<..<u} = {l..<u}"
-  "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..u} = {l<..u}"
-  "[| (l::'a::linorder) <= m; m <= u |] ==> {l<..m} Un {m<..u} = {l<..u}"
-  "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..u} = {l..u}"
-  "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m} Un {m<..u} = {l..u}"
+  "[| (l::'a::linorder) < m; m \<le> u |] ==> {l<..<m} Un {m..<u} = {l<..<u}"
+  "[| (l::'a::linorder) \<le> m; m < u |] ==> {l<..m} Un {m<..<u} = {l<..<u}"
+  "[| (l::'a::linorder) \<le> m; m \<le> u |] ==> {l..<m} Un {m..<u} = {l..<u}"
+  "[| (l::'a::linorder) \<le> m; m < u |] ==> {l..m} Un {m<..<u} = {l..<u}"
+  "[| (l::'a::linorder) < m; m \<le> u |] ==> {l<..<m} Un {m..u} = {l<..u}"
+  "[| (l::'a::linorder) \<le> m; m \<le> u |] ==> {l<..m} Un {m<..u} = {l<..u}"
+  "[| (l::'a::linorder) \<le> m; m \<le> u |] ==> {l..<m} Un {m..u} = {l..u}"
+  "[| (l::'a::linorder) \<le> m; m \<le> u |] ==> {l..m} Un {m<..u} = {l..u}"
 by auto
 
 lemma ivl_disj_un_two_touch:
   "[| (l::'a::linorder) < m; m < u |] ==> {l<..m} Un {m..<u} = {l<..<u}"
-  "[| (l::'a::linorder) <= m; m < u |] ==> {l..m} Un {m..<u} = {l..<u}"
-  "[| (l::'a::linorder) < m; m <= u |] ==> {l<..m} Un {m..u} = {l<..u}"
-  "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m} Un {m..u} = {l..u}"
+  "[| (l::'a::linorder) \<le> m; m < u |] ==> {l..m} Un {m..<u} = {l..<u}"
+  "[| (l::'a::linorder) < m; m \<le> u |] ==> {l<..m} Un {m..u} = {l<..u}"
+  "[| (l::'a::linorder) \<le> m; m \<le> u |] ==> {l..m} Un {m..u} = {l..u}"
 by auto
 
 lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ivl_disj_un_two_touch
@@ -1635,14 +1592,11 @@
 
 subsubsection \<open>Some Subset Conditions\<close>
 
-lemma ivl_subset [simp]:
- "({i..<j} \<subseteq> {m..<n}) = (j \<le> i \<or> m \<le> i \<and> j \<le> (n::'a::linorder))"
-apply(auto simp:linorder_not_le)
-apply(rule ccontr)
-apply(insert linorder_le_less_linear[of i n])
-apply(clarsimp simp:linorder_not_le)
-apply(fastforce)
-done
+lemma ivl_subset [simp]: "({i..<j} \<subseteq> {m..<n}) = (j \<le> i \<or> m \<le> i \<and> j \<le> (n::'a::linorder))"
+  using linorder_class.le_less_linear[of i n]
+  apply (auto simp: linorder_not_le)
+   apply (force intro: leI)+
+  done
 
 
 subsection \<open>Generic big monoid operation over intervals\<close>
@@ -1866,14 +1820,14 @@
   "sum f {m..<Suc n} = (if n < m then 0 else sum f {m..<n} + f(n))"
   by (auto simp: ac_simps atLeastLessThanSuc)
 (*
-lemma sum_cl_ivl_add_one_nat: "(n::nat) <= m + 1 ==>
+lemma sum_cl_ivl_add_one_nat: "(n::nat) \<le> m + 1 ==>
     (\<Sum>i=n..m+1. f i) = (\<Sum>i=n..m. f i) + f(m + 1)"
 by (auto simp:ac_simps atLeastAtMostSuc_conv)
 *)
 
 lemma sum_head:
   fixes n :: nat
-  assumes mn: "m <= n"
+  assumes mn: "m \<le> n"
   shows "(\<Sum>x\<in>{m..n}. P x) = P m + (\<Sum>x\<in>{m<..n}. P x)" (is "?lhs = ?rhs")
 proof -
   from mn
@@ -1904,24 +1858,25 @@
 lemmas sum_add_nat_ivl = sum.atLeastLessThan_concat
 
 lemma sum_diff_nat_ivl:
-fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
-shows "\<lbrakk> m \<le> n; n \<le> p \<rbrakk> \<Longrightarrow>
-  sum f {m..<p} - sum f {m..<n} = sum f {n..<p}"
-using sum_add_nat_ivl [of m n p f,symmetric]
-apply (simp add: ac_simps)
-done
+  fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
+  shows "\<lbrakk> m \<le> n; n \<le> p \<rbrakk> \<Longrightarrow> sum f {m..<p} - sum f {m..<n} = sum f {n..<p}"
+  using sum_add_nat_ivl [of m n p f,symmetric]
+  by (simp add: ac_simps)
 
 lemma sum_natinterval_difff:
   fixes f:: "nat \<Rightarrow> ('a::ab_group_add)"
   shows  "sum (\<lambda>k. f k - f(k + 1)) {(m::nat) .. n} =
-          (if m <= n then f m - f(n + 1) else 0)"
+          (if m \<le> n then f m - f(n + 1) else 0)"
 by (induct n, auto simp add: algebra_simps not_le le_Suc_eq)
 
 lemma sum_nat_group: "(\<Sum>m<n::nat. sum f {m * k ..< m*k + k}) = sum f {..< n * k}"
-  apply (subgoal_tac "k = 0 \<or> 0 < k", auto)
-  apply (induct "n")
-  apply (simp_all add: sum_add_nat_ivl add.commute atLeast0LessThan[symmetric])
-  done
+proof (cases k)
+  case (Suc l)
+  then have "k > 0"
+    by auto
+  then show ?thesis
+    by (induct n) (simp_all add: sum_add_nat_ivl add.commute atLeast0LessThan[symmetric])
+qed auto   
 
 lemma sum_triangle_reindex:
   fixes n :: nat
@@ -1971,19 +1926,19 @@
 subsubsection \<open>Shifting bounds\<close>
 
 lemma sum_shift_bounds_nat_ivl:
-  "sum f {m+k..<n+k} = sum (%i. f(i + k)){m..<n::nat}"
+  "sum f {m+k..<n+k} = sum (\<lambda>i. f(i + k)){m..<n::nat}"
 by (induct "n", auto simp:atLeastLessThanSuc)
 
 lemma sum_shift_bounds_cl_nat_ivl:
-  "sum f {m+k..n+k} = sum (%i. f(i + k)){m..n::nat}"
+  "sum f {m+k..n+k} = sum (\<lambda>i. f(i + k)){m..n::nat}"
   by (rule sum.reindex_bij_witness[where i="\<lambda>i. i + k" and j="\<lambda>i. i - k"]) auto
 
 corollary sum_shift_bounds_cl_Suc_ivl:
-  "sum f {Suc m..Suc n} = sum (%i. f(Suc i)){m..n}"
+  "sum f {Suc m..Suc n} = sum (\<lambda>i. f(Suc i)){m..n}"
 by (simp add:sum_shift_bounds_cl_nat_ivl[where k="Suc 0", simplified])
 
 corollary sum_shift_bounds_Suc_ivl:
-  "sum f {Suc m..<Suc n} = sum (%i. f(Suc i)){m..<n}"
+  "sum f {Suc m..<Suc n} = sum (\<lambda>i. f(Suc i)){m..<n}"
 by (simp add:sum_shift_bounds_nat_ivl[where k="Suc 0", simplified])
 
 context comm_monoid_add
@@ -2008,8 +1963,7 @@
     using \<open>f 0 = 0\<close> by simp
 qed
 
-lemma sum_shift_lb_Suc0_0:
-  "sum f {Suc 0..k} = sum f {0..k}"
+lemma sum_shift_lb_Suc0_0: "sum f {Suc 0..k} = sum f {0..k}"
 proof (cases k)
   case 0
   with \<open>f 0 = 0\<close> show ?thesis
@@ -2054,7 +2008,7 @@
   shows "(\<Sum>i\<le>n. f i) = f 0 + (\<Sum>i<n. f (Suc i))"
 by (metis atLeast0AtMost atLeast0LessThan atLeastLessThanSuc_atLeastAtMost atLeastSucAtMost_greaterThanAtMost le0 sum_head sum_shift_bounds_Suc_ivl)
 
-lemma sum_last_plus: fixes n::nat shows "m <= n \<Longrightarrow> (\<Sum>i = m..n. f i) = f n + (\<Sum>i = m..<n. f i)"
+lemma sum_last_plus: fixes n::nat shows "m \<le> n \<Longrightarrow> (\<Sum>i = m..n. f i) = f n + (\<Sum>i = m..<n. f i)"
   by (cases n) (auto simp: atLeastLessThanSuc_atLeastAtMost add.commute)
 
 lemma sum_Suc_diff:
@@ -2436,19 +2390,19 @@
 subsubsection \<open>Shifting bounds\<close>
 
 lemma prod_shift_bounds_nat_ivl:
-  "prod f {m+k..<n+k} = prod (%i. f(i + k)){m..<n::nat}"
+  "prod f {m+k..<n+k} = prod (\<lambda>i. f(i + k)){m..<n::nat}"
 by (induct "n", auto simp:atLeastLessThanSuc)
 
 lemma prod_shift_bounds_cl_nat_ivl:
-  "prod f {m+k..n+k} = prod (%i. f(i + k)){m..n::nat}"
+  "prod f {m+k..n+k} = prod (\<lambda>i. f(i + k)){m..n::nat}"
   by (rule prod.reindex_bij_witness[where i="\<lambda>i. i + k" and j="\<lambda>i. i - k"]) auto
 
 corollary prod_shift_bounds_cl_Suc_ivl:
-  "prod f {Suc m..Suc n} = prod (%i. f(Suc i)){m..n}"
+  "prod f {Suc m..Suc n} = prod (\<lambda>i. f(Suc i)){m..n}"
 by (simp add:prod_shift_bounds_cl_nat_ivl[where k="Suc 0", simplified])
 
 corollary prod_shift_bounds_Suc_ivl:
-  "prod f {Suc m..<Suc n} = prod (%i. f(Suc i)){m..<n}"
+  "prod f {Suc m..<Suc n} = prod (\<lambda>i. f(Suc i)){m..<n}"
 by (simp add:prod_shift_bounds_nat_ivl[where k="Suc 0", simplified])
 
 lemma prod_lessThan_Suc [simp]: "prod f {..<Suc n} = prod f {..<n} * f n"
--- a/src/HOL/Transitive_Closure.thy	Wed Jul 11 23:24:25 2018 +0100
+++ b/src/HOL/Transitive_Closure.thy	Thu Jul 12 17:22:39 2018 +0100
@@ -83,7 +83,7 @@
 subsection \<open>Reflexive-transitive closure\<close>
 
 lemma reflcl_set_eq [pred_set_conv]: "(sup (\<lambda>x y. (x, y) \<in> r) (=)) = (\<lambda>x y. (x, y) \<in> r \<union> Id)"
-  by (auto simp add: fun_eq_iff)
+  by (auto simp: fun_eq_iff)
 
 lemma r_into_rtrancl [intro]: "\<And>p. p \<in> r \<Longrightarrow> p \<in> r\<^sup>*"
   \<comment> \<open>\<open>rtrancl\<close> of \<open>r\<close> contains \<open>r\<close>\<close>
@@ -156,18 +156,16 @@
     (base) "a = b"
   | (step) y where "(a, y) \<in> r\<^sup>*" and "(y, b) \<in> r"
   \<comment> \<open>elimination of \<open>rtrancl\<close> -- by induction on a special formula\<close>
-  apply (subgoal_tac "a = b \<or> (\<exists>y. (a, y) \<in> r\<^sup>* \<and> (y, b) \<in> r)")
-   apply (rule_tac [2] major [THEN rtrancl_induct])
-    prefer 2 apply blast
-   prefer 2 apply blast
-  apply (erule asm_rl exE disjE conjE base step)+
-  done
+proof -
+  have "a = b \<or> (\<exists>y. (a, y) \<in> r\<^sup>* \<and> (y, b) \<in> r)"
+    by (rule major [THEN rtrancl_induct]) blast+
+  then show ?thesis
+    by (auto intro: base step)
+qed
 
 lemma rtrancl_Int_subset: "Id \<subseteq> s \<Longrightarrow> (r\<^sup>* \<inter> s) O r \<subseteq> s \<Longrightarrow> r\<^sup>* \<subseteq> s"
-  apply (rule subsetI)
-  apply auto
-  apply (erule rtrancl_induct)
-  apply auto
+  apply clarify
+  apply (erule rtrancl_induct, auto)
   done
 
 lemma converse_rtranclp_into_rtranclp: "r a b \<Longrightarrow> r\<^sup>*\<^sup>* b c \<Longrightarrow> r\<^sup>*\<^sup>* a c"
@@ -193,14 +191,11 @@
   done
 
 lemma rtrancl_subset_rtrancl: "r \<subseteq> s\<^sup>* \<Longrightarrow> r\<^sup>* \<subseteq> s\<^sup>*"
-  apply (drule rtrancl_mono)
-  apply simp
-  done
+by (drule rtrancl_mono, simp)
 
 lemma rtranclp_subset: "R \<le> S \<Longrightarrow> S \<le> R\<^sup>*\<^sup>* \<Longrightarrow> S\<^sup>*\<^sup>* = R\<^sup>*\<^sup>*"
   apply (drule rtranclp_mono)
-  apply (drule rtranclp_mono)
-  apply simp
+  apply (drule rtranclp_mono, simp)
   done
 
 lemmas rtrancl_subset = rtranclp_subset [to_set]
@@ -216,21 +211,10 @@
 lemmas rtrancl_reflcl [simp] = rtranclp_reflclp [to_set]
 
 lemma rtrancl_r_diff_Id: "(r - Id)\<^sup>* = r\<^sup>*"
-  apply (rule sym)
-  apply (rule rtrancl_subset)
-   apply blast
-  apply clarify
-  apply (rename_tac a b)
-  apply (case_tac "a = b")
-   apply blast
-  apply blast
-  done
+  by (rule rtrancl_subset [symmetric]) auto
 
 lemma rtranclp_r_diff_Id: "(inf r (\<noteq>))\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*"
-  apply (rule sym)
-  apply (rule rtranclp_subset)
-   apply blast+
-  done
+  by (rule rtranclp_subset [symmetric]) auto
 
 theorem rtranclp_converseD:
   assumes "(r\<inverse>\<inverse>)\<^sup>*\<^sup>* x y"
@@ -272,12 +256,12 @@
   assumes major: "r\<^sup>*\<^sup>* x z"
     and cases: "x = z \<Longrightarrow> P" "\<And>y. r x y \<Longrightarrow> r\<^sup>*\<^sup>* y z \<Longrightarrow> P"
   shows P
-  apply (subgoal_tac "x = z \<or> (\<exists>y. r x y \<and> r\<^sup>*\<^sup>* y z)")
-   apply (rule_tac [2] major [THEN converse_rtranclp_induct])
-    prefer 2 apply iprover
-   prefer 2 apply iprover
-  apply (erule asm_rl exE disjE conjE cases)+
-  done
+proof -
+  have "x = z \<or> (\<exists>y. r x y \<and> r\<^sup>*\<^sup>* y z)"
+    by (rule_tac major [THEN converse_rtranclp_induct]) iprover+
+  then show ?thesis
+    by (auto intro: cases)
+qed
 
 lemmas converse_rtranclE = converse_rtranclpE [to_set]
 
@@ -360,8 +344,7 @@
 
 lemma rtranclp_into_tranclp2: "r a b \<Longrightarrow> r\<^sup>*\<^sup>* b c \<Longrightarrow> r\<^sup>+\<^sup>+ a c"
   \<comment> \<open>intro rule from \<open>r\<close> and \<open>rtrancl\<close>\<close>
-  apply (erule rtranclp.cases)
-   apply iprover
+  apply (erule rtranclp.cases, iprover)
   apply (rule rtranclp_trans [THEN rtranclp_into_tranclp1])
     apply (simp | rule r_into_rtranclp)+
   done
@@ -401,10 +384,8 @@
   using assms by cases simp_all
 
 lemma trancl_Int_subset: "r \<subseteq> s \<Longrightarrow> (r\<^sup>+ \<inter> s) O r \<subseteq> s \<Longrightarrow> r\<^sup>+ \<subseteq> s"
-  apply (rule subsetI)
-  apply auto
-  apply (erule trancl_induct)
-   apply auto
+  apply clarify
+  apply (erule trancl_induct, auto)
   done
 
 lemma trancl_unfold: "r\<^sup>+ = r \<union> r\<^sup>+ O r"
@@ -438,10 +419,8 @@
 
 lemma trancl_id [simp]: "trans r \<Longrightarrow> r\<^sup>+ = r"
   apply auto
-  apply (erule trancl_induct)
-   apply assumption
-  apply (unfold trans_def)
-  apply blast
+  apply (erule trancl_induct, assumption)
+  apply (unfold trans_def, blast)
   done
 
 lemma rtranclp_tranclp_tranclp:
@@ -485,16 +464,14 @@
     and cases: "\<And>y. r y b \<Longrightarrow> P y" "\<And>y z. r y z \<Longrightarrow> r\<^sup>+\<^sup>+ z b \<Longrightarrow> P z \<Longrightarrow> P y"
   shows "P a"
   apply (rule tranclp_induct [OF tranclp_converseI, OF conversepI, OF major])
-   apply (rule cases)
-   apply (erule conversepD)
+   apply (blast intro: cases)
   apply (blast intro: assms dest!: tranclp_converseD)
   done
 
 lemmas converse_trancl_induct = converse_tranclp_induct [to_set]
 
 lemma tranclpD: "R\<^sup>+\<^sup>+ x y \<Longrightarrow> \<exists>z. R x z \<and> R\<^sup>*\<^sup>* z y"
-  apply (erule converse_tranclp_induct)
-   apply auto
+  apply (erule converse_tranclp_induct, auto)
   apply (blast intro: rtranclp_trans)
   done
 
@@ -537,8 +514,7 @@
   by (induct rule: rtrancl_induct) auto
 
 lemma trancl_subset_Sigma: "r \<subseteq> A \<times> A \<Longrightarrow> r\<^sup>+ \<subseteq> A \<times> A"
-  apply (rule subsetI)
-  apply (simp only: split_tupled_all)
+  apply (clarsimp simp:)
   apply (erule tranclE)
    apply (blast dest!: trancl_into_rtrancl trancl_subset_Sigma_aux)+
   done
@@ -552,13 +528,19 @@
 lemmas reflcl_trancl [simp] = reflclp_tranclp [to_set]
 
 lemma trancl_reflcl [simp]: "(r\<^sup>=)\<^sup>+ = r\<^sup>*"
-  apply safe
-   apply (drule trancl_into_rtrancl, simp)
-  apply (erule rtranclE, safe)
-   apply (rule r_into_trancl, simp)
-  apply (rule rtrancl_into_trancl1)
-   apply (erule rtrancl_reflcl [THEN equalityD2, THEN subsetD], fast)
-  done
+proof -
+  have "(a, b) \<in> (r\<^sup>=)\<^sup>+ \<Longrightarrow> (a, b) \<in> r\<^sup>*" for a b
+    by (force dest: trancl_into_rtrancl)
+  moreover have "(a, b) \<in> (r\<^sup>=)\<^sup>+" if "(a, b) \<in> r\<^sup>*" for a b
+    using that
+  proof (cases a b rule: rtranclE)
+    case step
+    show ?thesis
+      by (rule rtrancl_into_trancl1) (use step in auto)
+  qed auto
+  ultimately show ?thesis 
+    by auto
+qed
 
 lemma rtrancl_trancl_reflcl [code]: "r\<^sup>* = (r\<^sup>+)\<^sup>="
   by simp
@@ -570,7 +552,7 @@
   by (rule subst [OF reflcl_trancl]) simp
 
 lemma rtranclpD: "R\<^sup>*\<^sup>* a b \<Longrightarrow> a = b \<or> a \<noteq> b \<and> R\<^sup>+\<^sup>+ a b"
-  by (force simp add: reflclp_tranclp [symmetric] simp del: reflclp_tranclp)
+  by (force simp: reflclp_tranclp [symmetric] simp del: reflclp_tranclp)
 
 lemmas rtranclD = rtranclpD [to_set]
 
@@ -585,19 +567,20 @@
 
 lemma trancl_insert: "(insert (y, x) r)\<^sup>+ = r\<^sup>+ \<union> {(a, b). (a, y) \<in> r\<^sup>* \<and> (x, b) \<in> r\<^sup>*}"
   \<comment> \<open>primitive recursion for \<open>trancl\<close> over finite relations\<close>
-  apply (rule equalityI)
-   apply (rule subsetI)
-   apply (simp only: split_tupled_all)
-   apply (erule trancl_induct, blast)
-   apply (blast intro: rtrancl_into_trancl1 trancl_into_rtrancl trancl_trans)
-  apply (rule subsetI)
-  apply (blast intro: trancl_mono rtrancl_mono
-      [THEN [2] rev_subsetD] rtrancl_trancl_trancl rtrancl_into_trancl2)
-  done
+proof -
+  have "\<And>a b. (a, b) \<in> (insert (y, x) r)\<^sup>+ \<Longrightarrow>
+           (a, b) \<in> r\<^sup>+ \<union> {(a, b). (a, y) \<in> r\<^sup>* \<and> (x, b) \<in> r\<^sup>*}"
+    by (erule trancl_induct) (blast intro: rtrancl_into_trancl1 trancl_into_rtrancl trancl_trans)+
+  moreover have "r\<^sup>+ \<union> {(a, b). (a, y) \<in> r\<^sup>* \<and> (x, b) \<in> r\<^sup>*}  \<subseteq> (insert (y, x) r)\<^sup>+"
+    by (blast intro: trancl_mono rtrancl_mono [THEN [2] rev_subsetD]
+                     rtrancl_trancl_trancl rtrancl_into_trancl2)
+  ultimately show ?thesis
+    by auto
+qed
 
 lemma trancl_insert2:
   "(insert (a, b) r)\<^sup>+ = r\<^sup>+ \<union> {(x, y). ((x, a) \<in> r\<^sup>+ \<or> x = a) \<and> ((b, y) \<in> r\<^sup>+ \<or> y = b)}"
-  by (auto simp add: trancl_insert rtrancl_eq_or_trancl)
+  by (auto simp: trancl_insert rtrancl_eq_or_trancl)
 
 lemma rtrancl_insert: "(insert (a,b) r)\<^sup>* = r\<^sup>* \<union> {(x, y). (x, a) \<in> r\<^sup>* \<and> (b, y) \<in> r\<^sup>*}"
   using trancl_insert[of a b r]
@@ -636,28 +619,30 @@
 lemma trancl_range [simp]: "Range (r\<^sup>+) = Range r"
   unfolding Domain_converse [symmetric] by (simp add: trancl_converse [symmetric])
 
-lemma Not_Domain_rtrancl: "x \<notin> Domain R \<Longrightarrow> (x, y) \<in> R\<^sup>* \<longleftrightarrow> x = y"
-  apply auto
-  apply (erule rev_mp)
-  apply (erule rtrancl_induct)
-   apply auto
-  done
+lemma Not_Domain_rtrancl: 
+  assumes "x \<notin> Domain R" shows "(x, y) \<in> R\<^sup>* \<longleftrightarrow> x = y"
+proof -
+have "(x, y) \<in> R\<^sup>* \<Longrightarrow> x = y"
+  by (erule rtrancl_induct) (use assms in auto)
+  then show ?thesis
+    by auto
+qed
 
 lemma trancl_subset_Field2: "r\<^sup>+ \<subseteq> Field r \<times> Field r"
   apply clarify
   apply (erule trancl_induct)
-   apply (auto simp add: Field_def)
+   apply (auto simp: Field_def)
   done
 
 lemma finite_trancl[simp]: "finite (r\<^sup>+) = finite r"
-  apply auto
-   prefer 2
+proof
+  show "finite (r\<^sup>+) \<Longrightarrow> finite r"
+    by (blast intro: r_into_trancl' finite_subset)
+  show "finite r \<Longrightarrow> finite (r\<^sup>+)"
    apply (rule trancl_subset_Field2 [THEN finite_subset])
-   apply (rule finite_SigmaI)
-    prefer 3
-    apply (blast intro: r_into_trancl' finite_subset)
-   apply (auto simp add: finite_Field)
+   apply (auto simp: finite_Field)
   done
+qed
 
 lemma finite_rtrancl_Image[simp]: assumes "finite R" "finite A" shows "finite (R\<^sup>* `` A)"
 proof (rule ccontr)
@@ -670,13 +655,20 @@
   be merged with main body.\<close>
 
 lemma single_valued_confluent:
-  "single_valued r \<Longrightarrow> (x, y) \<in> r\<^sup>* \<Longrightarrow> (x, z) \<in> r\<^sup>* \<Longrightarrow> (y, z) \<in> r\<^sup>* \<or> (z, y) \<in> r\<^sup>*"
-  apply (erule rtrancl_induct)
-   apply simp
-  apply (erule disjE)
-   apply (blast elim:converse_rtranclE dest:single_valuedD)
-  apply (blast intro:rtrancl_trans)
-  done
+  assumes "single_valued r" and xy: "(x, y) \<in> r\<^sup>*" and xz: "(x, z) \<in> r\<^sup>*"
+  shows "(y, z) \<in> r\<^sup>* \<or> (z, y) \<in> r\<^sup>*"
+  using xy
+proof (induction rule: rtrancl_induct)
+  case base
+  show ?case
+    by (simp add: assms)   
+next
+  case (step y z)
+  with xz \<open>single_valued r\<close> show ?case
+    apply (auto simp: elim: converse_rtranclE dest: single_valuedD)
+    apply (blast intro: rtrancl_trans)
+    done
+qed
 
 lemma r_r_into_trancl: "(a, b) \<in> R \<Longrightarrow> (b, c) \<in> R \<Longrightarrow> (a, c) \<in> R\<^sup>+"
   by (fast intro: trancl_trans)
@@ -824,16 +816,18 @@
   by (fact relpow_Suc_D2' [to_pred])
 
 lemma relpow_E2:
-  "(x, z) \<in> R ^^ n \<Longrightarrow>
-    (n = 0 \<Longrightarrow> x = z \<Longrightarrow> P) \<Longrightarrow>
-    (\<And>y m. n = Suc m \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> (y, z) \<in> R ^^ m \<Longrightarrow> P) \<Longrightarrow> P"
-  apply (cases n)
-   apply simp
-  apply (rename_tac nat)
-  apply (cut_tac n=nat and R=R in relpow_Suc_D2')
-  apply simp
-  apply blast
-  done
+  assumes "(x, z) \<in> R ^^ n" "n = 0 \<Longrightarrow> x = z \<Longrightarrow> P"
+          "\<And>y m. n = Suc m \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> (y, z) \<in> R ^^ m \<Longrightarrow> P"
+      shows "P"
+proof (cases n)
+  case 0
+  with assms show ?thesis
+    by simp
+next
+  case (Suc m)
+  with assms relpow_Suc_D2' [of m R] show ?thesis
+    by force
+qed
 
 lemma relpowp_E2:
   "(P ^^ n) x z \<Longrightarrow>
@@ -915,22 +909,23 @@
   by (simp add: rtranclp_is_Sup_relpowp)
 
 lemma trancl_power: "p \<in> R\<^sup>+ \<longleftrightarrow> (\<exists>n > 0. p \<in> R ^^ n)"
-  apply (cases p)
-  apply simp
-  apply (rule iffI)
-   apply (drule tranclD2)
-   apply (clarsimp simp: rtrancl_is_UN_relpow)
-   apply (rule_tac x="Suc x" in exI)
-   apply (clarsimp simp: relcomp_unfold)
-   apply fastforce
-  apply clarsimp
-  apply (case_tac n)
-   apply simp
-  apply clarsimp
-  apply (drule relpow_imp_rtrancl)
-  apply (drule rtrancl_into_trancl1)
-   apply auto
-  done
+proof -
+  have "((a, b) \<in> R\<^sup>+) = (\<exists>n>0. (a, b) \<in> R ^^ n)" for a b
+  proof safe
+    show "(a, b) \<in> R\<^sup>+ \<Longrightarrow> \<exists>n>0. (a, b) \<in> R ^^ n"
+      apply (drule tranclD2)
+      apply (fastforce simp: rtrancl_is_UN_relpow relcomp_unfold)
+      done
+    show "(a, b) \<in> R\<^sup>+" if "n > 0" "(a, b) \<in> R ^^ n" for n
+    proof (cases n)
+      case (Suc m)
+      with that show ?thesis
+        by (auto simp: dest: relpow_imp_rtrancl rtrancl_into_trancl1)
+    qed (use that in auto)
+  qed
+  then show ?thesis
+    by (cases p) auto
+qed
 
 lemma tranclp_power: "(P\<^sup>+\<^sup>+) x y \<longleftrightarrow> (\<exists>n > 0. (P ^^ n) x y)"
   using trancl_power [to_pred, of P "(x, y)"] by simp
@@ -1070,8 +1065,7 @@
   fixes R :: "('a \<times> 'a) set"
   assumes "finite R"
   shows "R^^k \<subseteq> (UN n:{n. n \<le> card R}. R^^n)"
-  apply (cases k)
-   apply force
+  apply (cases k, force)
   apply (use relpow_finite_bounded1[OF assms, of k] in auto)
   done
 
@@ -1088,7 +1082,7 @@
   shows "finite (R O S)"
 proof-
   have "R O S = (\<Union>(x, y)\<in>R. \<Union>(u, v)\<in>S. if u = y then {(x, v)} else {})"
-    by (force simp add: split_def image_constant_conv split: if_splits)
+    by (force simp: split_def image_constant_conv split: if_splits)
   then show ?thesis
     using assms by clarsimp
 qed
@@ -1166,7 +1160,7 @@
   by (auto simp: Let_def)
 
 lemma finite_trancl_ntranl: "finite R \<Longrightarrow> trancl R = ntrancl (card R - 1) R"
-  by (cases "card R") (auto simp add: trancl_finite_eq_relpow relpow_empty ntrancl_def)
+  by (cases "card R") (auto simp: trancl_finite_eq_relpow relpow_empty ntrancl_def)
 
 
 subsection \<open>Acyclic relations\<close>