Merge
authorpaulson <lp15@cam.ac.uk>
Mon, 23 May 2016 15:46:30 +0100
changeset 63115 39dca4891601
parent 63114 27afe7af7379 (current diff)
parent 63113 fe31996e3898 (diff)
child 63116 32492105b015
child 63119 547460dc5c1e
child 63125 f2d3f8427bc2
Merge
--- a/NEWS	Mon May 23 15:33:24 2016 +0100
+++ b/NEWS	Mon May 23 15:46:30 2016 +0100
@@ -117,6 +117,9 @@
     INCOMPATIBILITY.
   - The "size" plugin has been made compatible again with locales.
 
+* Removed obsolete theorem nat_less_cases. INCOMPATIBILITY, use
+linorder_cases instead.
+
 * Renamed split_if -> if_split and split_if_asm -> if_split_asm to
 resemble the f.split naming convention, INCOMPATIBILITY.
 
--- a/src/HOL/Library/Omega_Words_Fun.thy	Mon May 23 15:33:24 2016 +0100
+++ b/src/HOL/Library/Omega_Words_Fun.thy	Mon May 23 15:46:30 2016 +0100
@@ -800,8 +800,8 @@
   assumes iseq: "idx_sequence idx"
     and eq: "idx m = idx n"
   shows "m = n"
-proof (rule nat_less_cases)
-  assume "n<m"
+proof (cases m n rule: linorder_cases)
+  case greater
   then obtain k where "m = Suc(n+k)"
     by (auto simp add: less_iff_Suc_add)
   with iseq have "idx n < idx m"
@@ -809,14 +809,14 @@
   with eq show ?thesis
     by simp
 next
-  assume "m<n"
+  case less
   then obtain k where "n = Suc(m+k)"
     by (auto simp add: less_iff_Suc_add)
   with iseq have "idx m < idx n"
     by (simp add: idx_sequence_less)
   with eq show ?thesis
     by simp
-qed (simp)
+qed
 
 lemma idx_sequence_mono:
   assumes iseq: "idx_sequence idx"
@@ -883,8 +883,8 @@
     and k: "n \<in> {idx k ..< idx (Suc k)}"
     and m: "n \<in> {idx m ..< idx (Suc m)}"
   shows "k = m"
-proof (rule nat_less_cases)
-  assume "k < m"
+proof (cases k m rule: linorder_cases)
+  case less
   hence "Suc k \<le> m" by simp
   with iseq have "idx (Suc k) \<le> idx m"
     by (rule idx_sequence_mono)
@@ -894,7 +894,7 @@
     by simp
   thus ?thesis ..
 next
-  assume "m < k"
+  case greater
   hence "Suc m \<le> k" by simp
   with iseq have "idx (Suc m) \<le> idx k"
     by (rule idx_sequence_mono)
@@ -903,7 +903,7 @@
   with m have "False"
     by simp
   thus ?thesis ..
-qed (simp)
+qed
 
 lemma idx_sequence_unique_interval:
   assumes iseq: "idx_sequence idx"
--- a/src/HOL/Nat.thy	Mon May 23 15:33:24 2016 +0100
+++ b/src/HOL/Nat.thy	Mon May 23 15:46:30 2016 +0100
@@ -21,10 +21,10 @@
 
 typedecl ind
 
-axiomatization Zero_Rep :: ind and Suc_Rep :: "ind => ind" where
-  \<comment> \<open>the axiom of infinity in 2 parts\<close>
-  Suc_Rep_inject:       "Suc_Rep x = Suc_Rep y ==> x = y" and
-  Suc_Rep_not_Zero_Rep: "Suc_Rep x \<noteq> Zero_Rep"
+axiomatization Zero_Rep :: ind and Suc_Rep :: "ind \<Rightarrow> ind"
+  \<comment> \<open>The axiom of infinity in 2 parts:\<close>
+where Suc_Rep_inject: "Suc_Rep x = Suc_Rep y \<Longrightarrow> x = y"
+  and Suc_Rep_not_Zero_Rep: "Suc_Rep x \<noteq> Zero_Rep"
 
 subsection \<open>Type nat\<close>
 
@@ -120,8 +120,8 @@
 
 setup \<open>Sign.parent_path\<close>
 
-abbreviation rec_nat :: "'a \<Rightarrow> (nat \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a" where
-  "rec_nat \<equiv> old.rec_nat"
+abbreviation rec_nat :: "'a \<Rightarrow> (nat \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
+  where "rec_nat \<equiv> old.rec_nat"
 
 declare nat.sel[code del]
 
@@ -199,8 +199,12 @@
 text \<open>A special form of induction for reasoning
   about @{term "m < n"} and @{term "m - n"}\<close>
 
-lemma diff_induct: "(!!x. P x 0) ==> (!!y. P 0 (Suc y)) ==>
-    (!!x y. P x y ==> P (Suc x) (Suc y)) ==> P m n"
+lemma diff_induct:
+  assumes "\<And>x. P x 0"
+    and "\<And>y. P 0 (Suc y)"
+    and "\<And>x y. P x y \<Longrightarrow> P (Suc x) (Suc y)"
+  shows "P m n"
+  using assms
   apply (rule_tac x = m in spec)
   apply (induct n)
   prefer 2
@@ -215,10 +219,10 @@
 begin
 
 primrec plus_nat where
-  add_0:      "0 + n = (n::nat)"
-| add_Suc:  "Suc m + n = Suc (m + n)"
-
-lemma add_0_right [simp]: "m + 0 = (m::nat)"
+  add_0: "0 + n = (n::nat)"
+| add_Suc: "Suc m + n = Suc (m + n)"
+
+lemma add_0_right [simp]: "m + 0 = m" for m :: nat
   by (induct m) simp_all
 
 lemma add_Suc_right [simp]: "m + Suc n = Suc (m + n)"
@@ -231,17 +235,18 @@
 
 primrec minus_nat where
   diff_0 [code]: "m - 0 = (m::nat)"
-| diff_Suc: "m - Suc n = (case m - n of 0 => 0 | Suc k => k)"
+| diff_Suc: "m - Suc n = (case m - n of 0 \<Rightarrow> 0 | Suc k \<Rightarrow> k)"
 
 declare diff_Suc [simp del]
 
-lemma diff_0_eq_0 [simp, code]: "0 - n = (0::nat)"
+lemma diff_0_eq_0 [simp, code]: "0 - n = 0" for n :: nat
   by (induct n) (simp_all add: diff_Suc)
 
 lemma diff_Suc_Suc [simp, code]: "Suc m - Suc n = m - n"
   by (induct n) (simp_all add: diff_Suc)
 
-instance proof
+instance
+proof
   fix n m q :: nat
   show "(n + m) + q = n + (m + q)" by (induct n) simp_all
   show "n + m = m + n" by (induct n) simp_all
@@ -265,25 +270,24 @@
   mult_0: "0 * n = (0::nat)"
 | mult_Suc: "Suc m * n = n + (m * n)"
 
-lemma mult_0_right [simp]: "(m::nat) * 0 = 0"
+lemma mult_0_right [simp]: "m * 0 = 0" for m :: nat
   by (induct m) simp_all
 
 lemma mult_Suc_right [simp]: "m * Suc n = m + (m * n)"
   by (induct m) (simp_all add: add.left_commute)
 
-lemma add_mult_distrib: "(m + n) * k = (m * k) + ((n * k)::nat)"
+lemma add_mult_distrib: "(m + n) * k = (m * k) + (n * k)" for m n k :: nat
   by (induct m) (simp_all add: add.assoc)
 
-instance proof
-  fix n m q :: nat
+instance
+proof
+  fix k n m q :: nat
   show "0 \<noteq> (1::nat)" unfolding One_nat_def by simp
   show "1 * n = n" unfolding One_nat_def by simp
   show "n * m = m * n" by (induct n) simp_all
   show "(n * m) * q = n * (m * q)" by (induct n) (simp_all add: add_mult_distrib)
   show "(n + m) * q = n * q + m * q" by (rule add_mult_distrib)
-next
-  fix k m n :: nat
-  show "k * ((m::nat) - n) = (k * m) - (k * n)"
+  show "k * (m - n) = (k * m) - (k * n)"
     by (induct m n rule: diff_induct) simp_all
 qed
 
@@ -294,29 +298,23 @@
 
 text \<open>Reasoning about \<open>m + 0 = 0\<close>, etc.\<close>
 
-lemma add_is_0 [iff]:
-  fixes m n :: nat
-  shows "(m + n = 0) = (m = 0 & n = 0)"
+lemma add_is_0 [iff]: "(m + n = 0) = (m = 0 \<and> n = 0)" for m n :: nat
   by (cases m) simp_all
 
-lemma add_is_1:
-  "(m+n= Suc 0) = (m= Suc 0 & n=0 | m=0 & n= Suc 0)"
+lemma add_is_1: "m + n = Suc 0 \<longleftrightarrow> m = Suc 0 \<and> n = 0 | m = 0 \<and> n = Suc 0"
   by (cases m) simp_all
 
-lemma one_is_add:
-  "(Suc 0 = m + n) = (m = Suc 0 & n = 0 | m = 0 & n = Suc 0)"
+lemma one_is_add: "Suc 0 = m + n \<longleftrightarrow> m = Suc 0 \<and> n = 0 | m = 0 \<and> n = Suc 0"
   by (rule trans, rule eq_commute, rule add_is_1)
 
-lemma add_eq_self_zero:
-  fixes m n :: nat
-  shows "m + n = m \<Longrightarrow> n = 0"
+lemma add_eq_self_zero: "m + n = m \<Longrightarrow> n = 0" for m n :: nat
   by (induct m) simp_all
 
-lemma inj_on_add_nat[simp]: "inj_on (%n::nat. n+k) N"
+lemma inj_on_add_nat[simp]: "inj_on (\<lambda>n. n + k) N" for k :: nat
   apply (induct k)
    apply simp
-  apply(drule comp_inj_on[OF _ inj_Suc])
-  apply (simp add:o_def)
+  apply (drule comp_inj_on[OF _ inj_Suc])
+  apply (simp add: o_def)
   done
 
 lemma Suc_eq_plus1: "Suc n = n + 1"
@@ -336,43 +334,45 @@
 
 subsubsection \<open>Multiplication\<close>
 
-lemma mult_is_0 [simp]: "((m::nat) * n = 0) = (m=0 | n=0)"
+lemma mult_is_0 [simp]: "m * n = 0 \<longleftrightarrow> m = 0 \<or> n = 0" for m n :: nat
   by (induct m) auto
 
-lemma mult_eq_1_iff [simp]: "(m * n = Suc 0) = (m = Suc 0 & n = Suc 0)"
+lemma mult_eq_1_iff [simp]: "m * n = Suc 0 \<longleftrightarrow> m = Suc 0 \<and> n = Suc 0"
   apply (induct m)
    apply simp
   apply (induct n)
    apply auto
   done
 
-lemma one_eq_mult_iff [simp]: "(Suc 0 = m * n) = (m = Suc 0 & n = Suc 0)"
+lemma one_eq_mult_iff [simp]: "Suc 0 = m * n \<longleftrightarrow> m = Suc 0 \<and> n = Suc 0"
   apply (rule trans)
   apply (rule_tac [2] mult_eq_1_iff, fastforce)
   done
 
-lemma nat_mult_eq_1_iff [simp]: "m * n = (1::nat) \<longleftrightarrow> m = 1 \<and> n = 1"
+lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \<longleftrightarrow> m = 1 \<and> n = 1" for m n :: nat
   unfolding One_nat_def by (rule mult_eq_1_iff)
 
-lemma nat_1_eq_mult_iff [simp]: "(1::nat) = m * n \<longleftrightarrow> m = 1 \<and> n = 1"
+lemma nat_1_eq_mult_iff [simp]: "1 = m * n \<longleftrightarrow> m = 1 \<and> n = 1" for m n :: nat
   unfolding One_nat_def by (rule one_eq_mult_iff)
 
-lemma mult_cancel1 [simp]: "(k * m = k * n) = (m = n | (k = (0::nat)))"
+lemma mult_cancel1 [simp]: "k * m = k * n \<longleftrightarrow> m = n \<or> k = 0" for k m n :: nat
 proof -
   have "k \<noteq> 0 \<Longrightarrow> k * m = k * n \<Longrightarrow> m = n"
   proof (induct n arbitrary: m)
-    case 0 then show "m = 0" by simp
+    case 0
+    then show "m = 0" by simp
   next
-    case (Suc n) then show "m = Suc n"
-      by (cases m) (simp_all add: eq_commute [of "0"])
+    case (Suc n)
+    then show "m = Suc n"
+      by (cases m) (simp_all add: eq_commute [of 0])
   qed
   then show ?thesis by auto
 qed
 
-lemma mult_cancel2 [simp]: "(m * k = n * k) = (m = n | (k = (0::nat)))"
+lemma mult_cancel2 [simp]: "m * k = n * k \<longleftrightarrow> m = n \<or> k = 0" for k m n :: nat
   by (simp add: mult.commute)
 
-lemma Suc_mult_cancel1: "(Suc k * m = Suc k * n) = (m = n)"
+lemma Suc_mult_cancel1: "Suc k * m = Suc k * n \<longleftrightarrow> m = n"
   by (subst mult_cancel1) simp
 
 
@@ -388,8 +388,12 @@
 | "Suc m \<le> n \<longleftrightarrow> (case n of 0 \<Rightarrow> False | Suc n \<Rightarrow> m \<le> n)"
 
 declare less_eq_nat.simps [simp del]
-lemma le0 [iff]: "0 \<le> (n::nat)" by (simp add: less_eq_nat.simps)
-lemma [code]: "(0::nat) \<le> n \<longleftrightarrow> True" by simp
+
+lemma le0 [iff]: "0 \<le> n" for n :: nat
+  by (simp add: less_eq_nat.simps)
+
+lemma [code]: "0 \<le> n \<longleftrightarrow> True" for n :: nat
+  by simp
 
 definition less_nat where
   less_eq_Suc_le: "n < m \<longleftrightarrow> Suc n \<le> m"
@@ -400,13 +404,13 @@
 lemma Suc_le_eq [code]: "Suc m \<le> n \<longleftrightarrow> m < n"
   unfolding less_eq_Suc_le ..
 
-lemma le_0_eq [iff]: "(n::nat) \<le> 0 \<longleftrightarrow> n = 0"
+lemma le_0_eq [iff]: "n \<le> 0 \<longleftrightarrow> n = 0" for n :: nat
   by (induct n) (simp_all add: less_eq_nat.simps(2))
 
-lemma not_less0 [iff]: "\<not> n < (0::nat)"
+lemma not_less0 [iff]: "\<not> n < 0" for n :: nat
   by (simp add: less_eq_Suc_le)
 
-lemma less_nat_zero_code [code]: "n < (0::nat) \<longleftrightarrow> False"
+lemma less_nat_zero_code [code]: "n < 0 \<longleftrightarrow> False" for n :: nat
   by simp
 
 lemma Suc_less_eq [iff]: "Suc m < Suc n \<longleftrightarrow> m < n"
@@ -419,8 +423,7 @@
   by (cases m) auto
 
 lemma le_SucI: "m \<le> n \<Longrightarrow> m \<le> Suc n"
-  by (induct m arbitrary: n)
-    (simp_all add: less_eq_nat.simps(2) split: nat.splits)
+  by (induct m arbitrary: n) (simp_all add: less_eq_nat.simps(2) split: nat.splits)
 
 lemma Suc_leD: "Suc m \<le> n \<Longrightarrow> m \<le> n"
   by (cases n) (auto intro: le_SucI)
@@ -433,33 +436,32 @@
 
 instance
 proof
-  fix n m :: nat
+  fix n m q :: nat
   show "n < m \<longleftrightarrow> n \<le> m \<and> \<not> m \<le> n"
   proof (induct n arbitrary: m)
-    case 0 then show ?case by (cases m) (simp_all add: less_eq_Suc_le)
+    case 0
+    then show ?case by (cases m) (simp_all add: less_eq_Suc_le)
   next
-    case (Suc n) then show ?case by (cases m) (simp_all add: less_eq_Suc_le)
+    case (Suc n)
+    then show ?case by (cases m) (simp_all add: less_eq_Suc_le)
   qed
-next
-  fix n :: nat show "n \<le> n" by (induct n) simp_all
-next
-  fix n m :: nat assume "n \<le> m" and "m \<le> n"
-  then show "n = m"
-    by (induct n arbitrary: m)
+  show "n \<le> n" by (induct n) simp_all
+  then show "n = m" if "n \<le> m" and "m \<le> n"
+    using that by (induct n arbitrary: m)
       (simp_all add: less_eq_nat.simps(2) split: nat.splits)
-next
-  fix n m q :: nat assume "n \<le> m" and "m \<le> q"
-  then show "n \<le> q"
+  show "n \<le> q" if "n \<le> m" and "m \<le> q"
+    using that
   proof (induct n arbitrary: m q)
-    case 0 show ?case by simp
+    case 0
+    show ?case by simp
   next
-    case (Suc n) then show ?case
+    case (Suc n)
+    then show ?case
       by (simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify,
         simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify,
         simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits)
   qed
-next
-  fix n m :: nat show "n \<le> m \<or> m \<le> n"
+  show "n \<le> m \<or> m \<le> n"
     by (induct n arbitrary: m)
       (simp_all add: less_eq_nat.simps(2) split: nat.splits)
 qed
@@ -469,11 +471,9 @@
 instantiation nat :: order_bot
 begin
 
-definition bot_nat :: nat where
-  "bot_nat = 0"
-
-instance proof
-qed (simp add: bot_nat_def)
+definition bot_nat :: nat where "bot_nat = 0"
+
+instance by standard (simp add: bot_nat_def)
 
 end
 
@@ -492,94 +492,91 @@
 
 subsubsection \<open>Elimination properties\<close>
 
-lemma less_not_refl: "~ n < (n::nat)"
+lemma less_not_refl: "\<not> n < n" for n :: nat
   by (rule order_less_irrefl)
 
-lemma less_not_refl2: "n < m ==> m \<noteq> (n::nat)"
+lemma less_not_refl2: "n < m \<Longrightarrow> m \<noteq> n" for m n :: nat
   by (rule not_sym) (rule less_imp_neq)
 
-lemma less_not_refl3: "(s::nat) < t ==> s \<noteq> t"
+lemma less_not_refl3: "s < t \<Longrightarrow> s \<noteq> t" for s t :: nat
   by (rule less_imp_neq)
 
-lemma less_irrefl_nat: "(n::nat) < n ==> R"
+lemma less_irrefl_nat: "n < n \<Longrightarrow> R" for n :: nat
   by (rule notE, rule less_not_refl)
 
-lemma less_zeroE: "(n::nat) < 0 ==> R"
+lemma less_zeroE: "n < 0 \<Longrightarrow> R" for n :: nat
   by (rule notE) (rule not_less0)
 
-lemma less_Suc_eq: "(m < Suc n) = (m < n | m = n)"
+lemma less_Suc_eq: "m < Suc n \<longleftrightarrow> m < n \<or> m = n"
   unfolding less_Suc_eq_le le_less ..
 
 lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)"
   by (simp add: less_Suc_eq)
 
-lemma less_one [iff]: "(n < (1::nat)) = (n = 0)"
+lemma less_one [iff]: "n < 1 \<longleftrightarrow> n = 0" for n :: nat
   unfolding One_nat_def by (rule less_Suc0)
 
-lemma Suc_mono: "m < n ==> Suc m < Suc n"
+lemma Suc_mono: "m < n \<Longrightarrow> Suc m < Suc n"
   by simp
 
 text \<open>"Less than" is antisymmetric, sort of\<close>
 lemma less_antisym: "\<lbrakk> \<not> n < m; n < Suc m \<rbrakk> \<Longrightarrow> m = n"
   unfolding not_less less_Suc_eq_le by (rule antisym)
 
-lemma nat_neq_iff: "((m::nat) \<noteq> n) = (m < n | n < m)"
+lemma nat_neq_iff: "m \<noteq> n \<longleftrightarrow> m < n \<or> n < m" for m n :: nat
   by (rule linorder_neq_iff)
 
-lemma nat_less_cases: assumes major: "(m::nat) < n ==> P n m"
-  and eqCase: "m = n ==> P n m" and lessCase: "n<m ==> P n m"
-  shows "P n m"
-  apply (rule less_linear [THEN disjE])
-  apply (erule_tac [2] disjE)
-  apply (erule lessCase)
-  apply (erule sym [THEN eqCase])
-  apply (erule major)
-  done
-
 
 subsubsection \<open>Inductive (?) properties\<close>
 
-lemma Suc_lessI: "m < n ==> Suc m \<noteq> n ==> Suc m < n"
+lemma Suc_lessI: "m < n \<Longrightarrow> Suc m \<noteq> n \<Longrightarrow> Suc m < n"
   unfolding less_eq_Suc_le [of m] le_less by simp
 
 lemma lessE:
   assumes major: "i < k"
-  and p1: "k = Suc i ==> P" and p2: "!!j. i < j ==> k = Suc j ==> P"
+    and 1: "k = Suc i \<Longrightarrow> P"
+    and 2: "\<And>j. i < j \<Longrightarrow> k = Suc j \<Longrightarrow> P"
   shows P
 proof -
   from major have "\<exists>j. i \<le> j \<and> k = Suc j"
     unfolding less_eq_Suc_le by (induct k) simp_all
   then have "(\<exists>j. i < j \<and> k = Suc j) \<or> k = Suc i"
-    by (clarsimp simp add: less_le)
-  with p1 p2 show P by auto
+    by (auto simp add: less_le)
+  with 1 2 show P by auto
 qed
 
-lemma less_SucE: assumes major: "m < Suc n"
-  and less: "m < n ==> P" and eq: "m = n ==> P" shows P
+lemma less_SucE:
+  assumes major: "m < Suc n"
+    and less: "m < n \<Longrightarrow> P"
+    and eq: "m = n \<Longrightarrow> P"
+  shows P
   apply (rule major [THEN lessE])
   apply (rule eq, blast)
   apply (rule less, blast)
   done
 
-lemma Suc_lessE: assumes major: "Suc i < k"
-  and minor: "!!j. i < j ==> k = Suc j ==> P" shows P
+lemma Suc_lessE:
+  assumes major: "Suc i < k"
+    and minor: "\<And>j. i < j \<Longrightarrow> k = Suc j \<Longrightarrow> P"
+  shows P
   apply (rule major [THEN lessE])
   apply (erule lessI [THEN minor])
   apply (erule Suc_lessD [THEN minor], assumption)
   done
 
-lemma Suc_less_SucD: "Suc m < Suc n ==> m < n"
+lemma Suc_less_SucD: "Suc m < Suc n \<Longrightarrow> m < n"
   by simp
 
 lemma less_trans_Suc:
-  assumes le: "i < j" shows "j < k ==> Suc i < k"
+  assumes le: "i < j"
+  shows "j < k \<Longrightarrow> Suc i < k"
   apply (induct k, simp_all)
-  apply (insert le)
+  using le
   apply (simp add: less_Suc_eq)
   apply (blast dest: Suc_lessD)
   done
 
-text \<open>Can be used with \<open>less_Suc_eq\<close> to get @{term "n = m | n < m"}\<close>
+text \<open>Can be used with \<open>less_Suc_eq\<close> to get @{prop "n = m \<or> n < m"}\<close>
 lemma not_less_eq: "\<not> m < n \<longleftrightarrow> n < Suc m"
   unfolding not_less less_Suc_eq_le ..
 
@@ -588,138 +585,138 @@
 
 text \<open>Properties of "less than or equal"\<close>
 
-lemma le_imp_less_Suc: "m \<le> n ==> m < Suc n"
+lemma le_imp_less_Suc: "m \<le> n \<Longrightarrow> m < Suc n"
   unfolding less_Suc_eq_le .
 
-lemma Suc_n_not_le_n: "~ Suc n \<le> n"
+lemma Suc_n_not_le_n: "\<not> Suc n \<le> n"
   unfolding not_le less_Suc_eq_le ..
 
 lemma le_Suc_eq: "(m \<le> Suc n) = (m \<le> n | m = Suc n)"
   by (simp add: less_Suc_eq_le [symmetric] less_Suc_eq)
 
-lemma le_SucE: "m \<le> Suc n ==> (m \<le> n ==> R) ==> (m = Suc n ==> R) ==> R"
+lemma le_SucE: "m \<le> Suc n \<Longrightarrow> (m \<le> n \<Longrightarrow> R) \<Longrightarrow> (m = Suc n \<Longrightarrow> R) \<Longrightarrow> R"
   by (drule le_Suc_eq [THEN iffD1], iprover+)
 
-lemma Suc_leI: "m < n ==> Suc(m) \<le> n"
+lemma Suc_leI: "m < n \<Longrightarrow> Suc(m) \<le> n"
   unfolding Suc_le_eq .
 
 text \<open>Stronger version of \<open>Suc_leD\<close>\<close>
-lemma Suc_le_lessD: "Suc m \<le> n ==> m < n"
+lemma Suc_le_lessD: "Suc m \<le> n \<Longrightarrow> m < n"
   unfolding Suc_le_eq .
 
-lemma less_imp_le_nat: "m < n ==> m \<le> (n::nat)"
+lemma less_imp_le_nat: "m < n \<Longrightarrow> m \<le> n" for m n :: nat
   unfolding less_eq_Suc_le by (rule Suc_leD)
 
 text \<open>For instance, \<open>(Suc m < Suc n) = (Suc m \<le> n) = (m < n)\<close>\<close>
 lemmas le_simps = less_imp_le_nat less_Suc_eq_le Suc_le_eq
 
 
-text \<open>Equivalence of @{term "m \<le> n"} and @{term "m < n | m = n"}\<close>
-
-lemma less_or_eq_imp_le: "m < n | m = n ==> m \<le> (n::nat)"
+text \<open>Equivalence of \<open>m \<le> n\<close> and \<open>m < n \<or> m = n\<close>\<close>
+
+lemma less_or_eq_imp_le: "m < n \<or> m = n \<Longrightarrow> m \<le> n" for m n :: nat
   unfolding le_less .
 
-lemma le_eq_less_or_eq: "(m \<le> (n::nat)) = (m < n | m=n)"
+lemma le_eq_less_or_eq: "m \<le> n \<longleftrightarrow> m < n \<or> m = n" for m n :: nat
   by (rule le_less)
 
 text \<open>Useful with \<open>blast\<close>.\<close>
-lemma eq_imp_le: "(m::nat) = n ==> m \<le> n"
+lemma eq_imp_le: "m = n \<Longrightarrow> m \<le> n" for m n :: nat
   by auto
 
-lemma le_refl: "n \<le> (n::nat)"
+lemma le_refl: "n \<le> n" for n :: nat
   by simp
 
-lemma le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::nat)"
+lemma le_trans: "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> i \<le> k" for i j k :: nat
   by (rule order_trans)
 
-lemma le_antisym: "[| m \<le> n; n \<le> m |] ==> m = (n::nat)"
+lemma le_antisym: "m \<le> n \<Longrightarrow> n \<le> m \<Longrightarrow> m = n" for m n :: nat
   by (rule antisym)
 
-lemma nat_less_le: "((m::nat) < n) = (m \<le> n & m \<noteq> n)"
+lemma nat_less_le: "m < n \<longleftrightarrow> m \<le> n \<and> m \<noteq> n" for m n :: nat
   by (rule less_le)
 
-lemma le_neq_implies_less: "(m::nat) \<le> n ==> m \<noteq> n ==> m < n"
+lemma le_neq_implies_less: "m \<le> n \<Longrightarrow> m \<noteq> n \<Longrightarrow> m < n" for m n :: nat
   unfolding less_le ..
 
-lemma nat_le_linear: "(m::nat) \<le> n | n \<le> m"
+lemma nat_le_linear: "m \<le> n | n \<le> m" for m n :: nat
   by (rule linear)
 
 lemmas linorder_neqE_nat = linorder_neqE [where 'a = nat]
 
-lemma le_less_Suc_eq: "m \<le> n ==> (n < Suc m) = (n = m)"
+lemma le_less_Suc_eq: "m \<le> n \<Longrightarrow> n < Suc m \<longleftrightarrow> n = m"
   unfolding less_Suc_eq_le by auto
 
-lemma not_less_less_Suc_eq: "~ n < m ==> (n < Suc m) = (n = m)"
+lemma not_less_less_Suc_eq: "\<not> n < m \<Longrightarrow> n < Suc m \<longleftrightarrow> n = m"
   unfolding not_less by (rule le_less_Suc_eq)
 
 lemmas not_less_simps = not_less_less_Suc_eq le_less_Suc_eq
 
-lemma not0_implies_Suc: "n \<noteq> 0 ==> \<exists>m. n = Suc m"
-by (cases n) simp_all
-
-lemma gr0_implies_Suc: "n > 0 ==> \<exists>m. n = Suc m"
-by (cases n) simp_all
-
-lemma gr_implies_not0: fixes n :: nat shows "m<n ==> n \<noteq> 0"
-by (cases n) simp_all
-
-lemma neq0_conv[iff]: fixes n :: nat shows "(n \<noteq> 0) = (0 < n)"
-by (cases n) simp_all
+lemma not0_implies_Suc: "n \<noteq> 0 \<Longrightarrow> \<exists>m. n = Suc m"
+  by (cases n) simp_all
+
+lemma gr0_implies_Suc: "n > 0 \<Longrightarrow> \<exists>m. n = Suc m"
+  by (cases n) simp_all
+
+lemma gr_implies_not0: "m < n \<Longrightarrow> n \<noteq> 0" for m n :: nat
+  by (cases n) simp_all
+
+lemma neq0_conv[iff]: "n \<noteq> 0 \<longleftrightarrow> 0 < n" for n :: nat
+  by (cases n) simp_all
 
 text \<open>This theorem is useful with \<open>blast\<close>\<close>
-lemma gr0I: "((n::nat) = 0 ==> False) ==> 0 < n"
-by (rule neq0_conv[THEN iffD1], iprover)
-
-lemma gr0_conv_Suc: "(0 < n) = (\<exists>m. n = Suc m)"
-by (fast intro: not0_implies_Suc)
-
-lemma not_gr0 [iff]: "!!n::nat. (~ (0 < n)) = (n = 0)"
-using neq0_conv by blast
-
-lemma Suc_le_D: "(Suc n \<le> m') ==> (? m. m' = Suc m)"
-by (induct m') simp_all
+lemma gr0I: "(n = 0 \<Longrightarrow> False) \<Longrightarrow> 0 < n" for n :: nat
+  by (rule neq0_conv[THEN iffD1], iprover)
+
+lemma gr0_conv_Suc: "0 < n \<longleftrightarrow> (\<exists>m. n = Suc m)"
+  by (fast intro: not0_implies_Suc)
+
+lemma not_gr0 [iff]: "\<not> 0 < n \<longleftrightarrow> n = 0" for n :: nat
+  using neq0_conv by blast
+
+lemma Suc_le_D: "Suc n \<le> m' \<Longrightarrow> \<exists>m. m' = Suc m"
+  by (induct m') simp_all
 
 text \<open>Useful in certain inductive arguments\<close>
-lemma less_Suc_eq_0_disj: "(m < Suc n) = (m = 0 | (\<exists>j. m = Suc j & j < n))"
-by (cases m) simp_all
+lemma less_Suc_eq_0_disj: "m < Suc n \<longleftrightarrow> m = 0 \<or> (\<exists>j. m = Suc j \<and> j < n)"
+  by (cases m) simp_all
 
 
 subsubsection \<open>Monotonicity of Addition\<close>
 
-lemma Suc_pred [simp]: "n>0 ==> Suc (n - Suc 0) = n"
-by (simp add: diff_Suc split: nat.split)
-
-lemma Suc_diff_1 [simp]: "0 < n ==> Suc (n - 1) = n"
-unfolding One_nat_def by (rule Suc_pred)
-
-lemma nat_add_left_cancel_le [simp]: "(k + m \<le> k + n) = (m\<le>(n::nat))"
-by (induct k) simp_all
-
-lemma nat_add_left_cancel_less [simp]: "(k + m < k + n) = (m<(n::nat))"
-by (induct k) simp_all
-
-lemma add_gr_0 [iff]: "!!m::nat. (m + n > 0) = (m>0 | n>0)"
-by(auto dest:gr0_implies_Suc)
+lemma Suc_pred [simp]: "n > 0 \<Longrightarrow> Suc (n - Suc 0) = n"
+  by (simp add: diff_Suc split: nat.split)
+
+lemma Suc_diff_1 [simp]: "0 < n \<Longrightarrow> Suc (n - 1) = n"
+  unfolding One_nat_def by (rule Suc_pred)
+
+lemma nat_add_left_cancel_le [simp]: "k + m \<le> k + n \<longleftrightarrow> m \<le> n" for k m n :: nat
+  by (induct k) simp_all
+
+lemma nat_add_left_cancel_less [simp]: "k + m < k + n \<longleftrightarrow> m < n" for k m n :: nat
+  by (induct k) simp_all
+
+lemma add_gr_0 [iff]: "m + n > 0 \<longleftrightarrow> m > 0 \<or> n > 0" for m n :: nat
+  by (auto dest: gr0_implies_Suc)
 
 text \<open>strict, in 1st argument\<close>
-lemma add_less_mono1: "i < j ==> i + k < j + (k::nat)"
-by (induct k) simp_all
+lemma add_less_mono1: "i < j \<Longrightarrow> i + k < j + k" for i j k :: nat
+  by (induct k) simp_all
 
 text \<open>strict, in both arguments\<close>
-lemma add_less_mono: "[|i < j; k < l|] ==> i + k < j + (l::nat)"
+lemma add_less_mono: "i < j \<Longrightarrow> k < l \<Longrightarrow> i + k < j + l" for i j k l :: nat
   apply (rule add_less_mono1 [THEN less_trans], assumption+)
   apply (induct j, simp_all)
   done
 
 text \<open>Deleted \<open>less_natE\<close>; use \<open>less_imp_Suc_add RS exE\<close>\<close>
-lemma less_imp_Suc_add: "m < n ==> (\<exists>k. n = Suc (m + k))"
+lemma less_imp_Suc_add: "m < n \<Longrightarrow> \<exists>k. n = Suc (m + k)"
   apply (induct n)
   apply (simp_all add: order_le_less)
   apply (blast elim!: less_SucE
                intro!: Nat.add_0_right [symmetric] add_Suc_right [symmetric])
   done
 
-lemma le_Suc_ex: "(k::nat) \<le> l \<Longrightarrow> (\<exists>n. l = k + n)"
+lemma le_Suc_ex: "k \<le> l \<Longrightarrow> (\<exists>n. l = k + n)" for k l :: nat
   by (auto simp: less_Suc_eq_le[symmetric] dest: less_imp_Suc_add)
 
 text \<open>strict, in 1st argument; proof is by induction on \<open>k > 0\<close>\<close>
@@ -727,173 +724,162 @@
   fixes i j :: nat
   assumes "i < j" and "0 < k"
   shows "k * i < k * j"
-using \<open>0 < k\<close> proof (induct k)
-  case 0 then show ?case by simp
+  using \<open>0 < k\<close>
+proof (induct k)
+  case 0
+  then show ?case by simp
 next
-  case (Suc k) with \<open>i < j\<close> show ?case
+  case (Suc k)
+  with \<open>i < j\<close> show ?case
     by (cases k) (simp_all add: add_less_mono)
 qed
 
 text \<open>Addition is the inverse of subtraction:
   if @{term "n \<le> m"} then @{term "n + (m - n) = m"}.\<close>
-lemma add_diff_inverse_nat: "~  m < n ==> n + (m - n) = (m::nat)"
-by (induct m n rule: diff_induct) simp_all
-
-lemma nat_le_iff_add: "(m::nat) \<le> n = (\<exists>k. n = m + k)"
-using nat_add_left_cancel_le[of m 0] by (auto dest: le_Suc_ex)
+lemma add_diff_inverse_nat: "\<not> m < n \<Longrightarrow> n + (m - n) = m" for m n :: nat
+  by (induct m n rule: diff_induct) simp_all
+
+lemma nat_le_iff_add: "m \<le> n \<longleftrightarrow> (\<exists>k. n = m + k)" for m n :: nat
+  using nat_add_left_cancel_le[of m 0] by (auto dest: le_Suc_ex)
 
 text\<open>The naturals form an ordered \<open>semidom\<close> and a \<open>dioid\<close>\<close>
 
 instance nat :: linordered_semidom
 proof
+  fix m n q :: nat
   show "0 < (1::nat)" by simp
-  show "\<And>m n q :: nat. m \<le> n \<Longrightarrow> q + m \<le> q + n" by simp
-  show "\<And>m n q :: nat. m < n \<Longrightarrow> 0 < q \<Longrightarrow> q * m < q * n" by (simp add: mult_less_mono2)
-  show "\<And>m n :: nat. m \<noteq> 0 \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> m * n \<noteq> 0" by simp
-  show "\<And>m n :: nat. n \<le> m ==> (m - n) + n = (m::nat)"
+  show "m \<le> n \<Longrightarrow> q + m \<le> q + n" by simp
+  show "m < n \<Longrightarrow> 0 < q \<Longrightarrow> q * m < q * n" by (simp add: mult_less_mono2)
+  show "m \<noteq> 0 \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> m * n \<noteq> 0" by simp
+  show "n \<le> m \<Longrightarrow> (m - n) + n = m"
     by (simp add: add_diff_inverse_nat add.commute linorder_not_less)
 qed
 
 instance nat :: dioid
-  proof qed (rule nat_le_iff_add)
+  by standard (rule nat_le_iff_add)
 declare le0[simp del] -- \<open>This is now @{thm zero_le}\<close>
 declare le_0_eq[simp del] -- \<open>This is now @{thm le_zero_eq}\<close>
 declare not_less0[simp del] -- \<open>This is now @{thm not_less_zero}\<close>
 declare not_gr0[simp del] -- \<open>This is now @{thm not_gr_zero}\<close>
 
-instance nat :: ordered_cancel_comm_monoid_add
-  proof qed
-
-instance nat :: ordered_cancel_comm_monoid_diff
-  proof qed
+instance nat :: ordered_cancel_comm_monoid_add ..
+instance nat :: ordered_cancel_comm_monoid_diff ..
+
 
 subsubsection \<open>@{term min} and @{term max}\<close>
 
 lemma mono_Suc: "mono Suc"
-by (rule monoI) simp
-
-lemma min_0L [simp]: "min 0 n = (0::nat)"
-by (rule min_absorb1) simp
-
-lemma min_0R [simp]: "min n 0 = (0::nat)"
-by (rule min_absorb2) simp
+  by (rule monoI) simp
+
+lemma min_0L [simp]: "min 0 n = 0" for n :: nat
+  by (rule min_absorb1) simp
+
+lemma min_0R [simp]: "min n 0 = 0" for n :: nat
+  by (rule min_absorb2) simp
 
 lemma min_Suc_Suc [simp]: "min (Suc m) (Suc n) = Suc (min m n)"
-by (simp add: mono_Suc min_of_mono)
-
-lemma min_Suc1:
-   "min (Suc n) m = (case m of 0 => 0 | Suc m' => Suc(min n m'))"
-by (simp split: nat.split)
-
-lemma min_Suc2:
-   "min m (Suc n) = (case m of 0 => 0 | Suc m' => Suc(min m' n))"
-by (simp split: nat.split)
-
-lemma max_0L [simp]: "max 0 n = (n::nat)"
-by (rule max_absorb2) simp
-
-lemma max_0R [simp]: "max n 0 = (n::nat)"
-by (rule max_absorb1) simp
-
-lemma max_Suc_Suc [simp]: "max (Suc m) (Suc n) = Suc(max m n)"
-by (simp add: mono_Suc max_of_mono)
-
-lemma max_Suc1:
-   "max (Suc n) m = (case m of 0 => Suc n | Suc m' => Suc(max n m'))"
-by (simp split: nat.split)
-
-lemma max_Suc2:
-   "max m (Suc n) = (case m of 0 => Suc n | Suc m' => Suc(max m' n))"
-by (simp split: nat.split)
-
-lemma nat_mult_min_left:
-  fixes m n q :: nat
-  shows "min m n * q = min (m * q) (n * q)"
-  by (simp add: min_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans)
-
-lemma nat_mult_min_right:
-  fixes m n q :: nat
-  shows "m * min n q = min (m * n) (m * q)"
-  by (simp add: min_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans)
-
-lemma nat_add_max_left:
-  fixes m n q :: nat
-  shows "max m n + q = max (m + q) (n + q)"
+  by (simp add: mono_Suc min_of_mono)
+
+lemma min_Suc1: "min (Suc n) m = (case m of 0 \<Rightarrow> 0 | Suc m' \<Rightarrow> Suc(min n m'))"
+  by (simp split: nat.split)
+
+lemma min_Suc2: "min m (Suc n) = (case m of 0 \<Rightarrow> 0 | Suc m' \<Rightarrow> Suc(min m' n))"
+  by (simp split: nat.split)
+
+lemma max_0L [simp]: "max 0 n = n" for n :: nat
+  by (rule max_absorb2) simp
+
+lemma max_0R [simp]: "max n 0 = n" for n :: nat
+  by (rule max_absorb1) simp
+
+lemma max_Suc_Suc [simp]: "max (Suc m) (Suc n) = Suc (max m n)"
+  by (simp add: mono_Suc max_of_mono)
+
+lemma max_Suc1: "max (Suc n) m = (case m of 0 \<Rightarrow> Suc n | Suc m' \<Rightarrow> Suc (max n m'))"
+  by (simp split: nat.split)
+
+lemma max_Suc2: "max m (Suc n) = (case m of 0 \<Rightarrow> Suc n | Suc m' \<Rightarrow> Suc (max m' n))"
+  by (simp split: nat.split)
+
+lemma nat_mult_min_left: "min m n * q = min (m * q) (n * q)" for m n q :: nat
+  by (simp add: min_def not_le)
+    (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans)
+
+lemma nat_mult_min_right: "m * min n q = min (m * n) (m * q)" for m n q :: nat
+  by (simp add: min_def not_le)
+    (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans)
+
+lemma nat_add_max_left: "max m n + q = max (m + q) (n + q)" for m n q :: nat
   by (simp add: max_def)
 
-lemma nat_add_max_right:
-  fixes m n q :: nat
-  shows "m + max n q = max (m + n) (m + q)"
+lemma nat_add_max_right: "m + max n q = max (m + n) (m + q)" for m n q :: nat
   by (simp add: max_def)
 
-lemma nat_mult_max_left:
-  fixes m n q :: nat
-  shows "max m n * q = max (m * q) (n * q)"
-  by (simp add: max_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans)
-
-lemma nat_mult_max_right:
-  fixes m n q :: nat
-  shows "m * max n q = max (m * n) (m * q)"
-  by (simp add: max_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans)
+lemma nat_mult_max_left: "max m n * q = max (m * q) (n * q)" for m n q :: nat
+  by (simp add: max_def not_le)
+    (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans)
+
+lemma nat_mult_max_right: "m * max n q = max (m * n) (m * q)" for m n q :: nat
+  by (simp add: max_def not_le)
+    (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans)
 
 
 subsubsection \<open>Additional theorems about @{term "op \<le>"}\<close>
 
 text \<open>Complete induction, aka course-of-values induction\<close>
 
-instance nat :: wellorder proof
+instance nat :: wellorder
+proof
   fix P and n :: nat
-  assume step: "\<And>n::nat. (\<And>m. m < n \<Longrightarrow> P m) \<Longrightarrow> P n"
+  assume step: "(\<And>m. m < n \<Longrightarrow> P m) \<Longrightarrow> P n" for n :: nat
   have "\<And>q. q \<le> n \<Longrightarrow> P q"
   proof (induct n)
     case (0 n)
     have "P 0" by (rule step) auto
-    thus ?case using 0 by auto
+    then show ?case using 0 by auto
   next
     case (Suc m n)
     then have "n \<le> m \<or> n = Suc m" by (simp add: le_Suc_eq)
-    thus ?case
+    then show ?case
     proof
-      assume "n \<le> m" thus "P n" by (rule Suc(1))
+      assume "n \<le> m"
+      then show "P n" by (rule Suc(1))
     next
       assume n: "n = Suc m"
-      show "P n"
-        by (rule step) (rule Suc(1), simp add: n le_simps)
+      show "P n" by (rule step) (rule Suc(1), simp add: n le_simps)
     qed
   qed
   then show "P n" by auto
 qed
 
 
-lemma Least_eq_0[simp]: "P(0::nat) \<Longrightarrow> Least P = 0"
-by (rule Least_equality[OF _ le0])
-
-lemma Least_Suc:
-     "[| P n; ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P(Suc m))"
+lemma Least_eq_0[simp]: "P 0 \<Longrightarrow> Least P = 0" for P :: "nat \<Rightarrow> bool"
+  by (rule Least_equality[OF _ le0])
+
+lemma Least_Suc: "P n \<Longrightarrow> \<not> P 0 \<Longrightarrow> (LEAST n. P n) = Suc (LEAST m. P (Suc m))"
   apply (cases n, auto)
   apply (frule LeastI)
-  apply (drule_tac P = "%x. P (Suc x) " in LeastI)
+  apply (drule_tac P = "\<lambda>x. P (Suc x) " in LeastI)
   apply (subgoal_tac " (LEAST x. P x) \<le> Suc (LEAST x. P (Suc x))")
   apply (erule_tac [2] Least_le)
   apply (cases "LEAST x. P x", auto)
-  apply (drule_tac P = "%x. P (Suc x) " in Least_le)
+  apply (drule_tac P = "\<lambda>x. P (Suc x) " in Least_le)
   apply (blast intro: order_antisym)
   done
 
-lemma Least_Suc2:
-   "[|P n; Q m; ~P 0; !k. P (Suc k) = Q k|] ==> Least P = Suc (Least Q)"
+lemma Least_Suc2: "P n \<Longrightarrow> Q m \<Longrightarrow> \<not> P 0 \<Longrightarrow> \<forall>k. P (Suc k) = Q k \<Longrightarrow> Least P = Suc (Least Q)"
   apply (erule (1) Least_Suc [THEN ssubst])
   apply simp
   done
 
-lemma ex_least_nat_le: "\<not>P(0) \<Longrightarrow> P(n::nat) \<Longrightarrow> \<exists>k\<le>n. (\<forall>i<k. \<not>P i) & P(k)"
+lemma ex_least_nat_le: "\<not> P 0 \<Longrightarrow> P n \<Longrightarrow> \<exists>k\<le>n. (\<forall>i<k. \<not> P i) \<and> P k" for P :: "nat \<Rightarrow> bool"
   apply (cases n)
    apply blast
-  apply (rule_tac x="LEAST k. P(k)" in exI)
+  apply (rule_tac x="LEAST k. P k" in exI)
   apply (blast intro: Least_le dest: not_less_Least intro: LeastI_ex)
   done
 
-lemma ex_least_nat_less: "\<not>P(0) \<Longrightarrow> P(n::nat) \<Longrightarrow> \<exists>k<n. (\<forall>i\<le>k. \<not>P i) & P(k+1)"
+lemma ex_least_nat_less: "\<not> P 0 \<Longrightarrow> P n \<Longrightarrow> \<exists>k<n. (\<forall>i\<le>k. \<not> P i) \<and> P (k + 1)" for P :: "nat \<Rightarrow> bool"
   unfolding One_nat_def
   apply (cases n)
    apply blast
@@ -907,14 +893,16 @@
   done
 
 lemma nat_less_induct:
-  assumes "!!n. \<forall>m::nat. m < n --> P m ==> P n" shows "P n"
+  fixes P :: "nat \<Rightarrow> bool"
+  assumes "\<And>n. \<forall>m. m < n \<longrightarrow> P m \<Longrightarrow> P n"
+  shows "P n"
   using assms less_induct by blast
 
 lemma measure_induct_rule [case_names less]:
   fixes f :: "'a \<Rightarrow> nat"
   assumes step: "\<And>x. (\<And>y. f y < f x \<Longrightarrow> P y) \<Longrightarrow> P x"
   shows "P a"
-by (induct m\<equiv>"f a" arbitrary: a rule: less_induct) (auto intro: step)
+  by (induct m \<equiv> "f a" arbitrary: a rule: less_induct) (auto intro: step)
 
 text \<open>old style induction rules:\<close>
 lemma measure_induct:
@@ -923,18 +911,19 @@
   by (rule measure_induct_rule [of f P a]) iprover
 
 lemma full_nat_induct:
-  assumes step: "(!!n. (ALL m. Suc m <= n --> P m) ==> P n)"
+  assumes step: "\<And>n. (\<forall>m. Suc m \<le> n \<longrightarrow> P m) \<Longrightarrow> P n"
   shows "P n"
   by (rule less_induct) (auto intro: step simp:le_simps)
 
-text\<open>An induction rule for estabilishing binary relations\<close>
+text\<open>An induction rule for establishing binary relations\<close>
 lemma less_Suc_induct [consumes 1]:
-  assumes less:  "i < j"
-     and  step:  "!!i. P i (Suc i)"
-     and  trans: "!!i j k. i < j ==> j < k ==>  P i j ==> P j k ==> P i k"
+  assumes less: "i < j"
+    and step: "\<And>i. P i (Suc i)"
+    and trans: "\<And>i j k. i < j \<Longrightarrow> j < k \<Longrightarrow> P i j \<Longrightarrow> P j k \<Longrightarrow> P i k"
   shows "P i j"
 proof -
-  from less obtain k where j: "j = Suc (i + k)" by (auto dest: less_imp_Suc_add)
+  from less obtain k where j: "j = Suc (i + k)"
+    by (auto dest: less_imp_Suc_add)
   have "P i (Suc (i + k))"
   proof (induct k)
     case 0
@@ -942,55 +931,60 @@
   next
     case (Suc k)
     have "0 + i < Suc k + i" by (rule add_less_mono1) simp
-    hence "i < Suc (i + k)" by (simp add: add.commute)
+    then have "i < Suc (i + k)" by (simp add: add.commute)
     from trans[OF this lessI Suc step]
     show ?case by simp
   qed
-  thus "P i j" by (simp add: j)
+  then show "P i j" by (simp add: j)
 qed
 
-text \<open>The method of infinite descent, frequently used in number theory.
-Provided by Roelof Oosterhuis.
-$P(n)$ is true for all $n\in\mathbb{N}$ if
-\begin{itemize}
-  \item case ``0'': given $n=0$ prove $P(n)$,
-  \item case ``smaller'': given $n>0$ and $\neg P(n)$ prove there exists
-        a smaller integer $m$ such that $\neg P(m)$.
-\end{itemize}\<close>
-
-text\<open>A compact version without explicit base case:\<close>
-lemma infinite_descent:
-  "\<lbrakk> !!n::nat. \<not> P n \<Longrightarrow>  \<exists>m<n. \<not>  P m \<rbrakk> \<Longrightarrow>  P n"
-by (induct n rule: less_induct) auto
-
-lemma infinite_descent0[case_names 0 smaller]:
-  "\<lbrakk> P 0; !!n. n>0 \<Longrightarrow> \<not> P n \<Longrightarrow> (\<exists>m::nat. m < n \<and> \<not>P m) \<rbrakk> \<Longrightarrow> P n"
-by (rule infinite_descent) (case_tac "n>0", auto)
-
 text \<open>
-Infinite descent using a mapping to $\mathbb{N}$:
-$P(x)$ is true for all $x\in D$ if there exists a $V: D \to \mathbb{N}$ and
-\begin{itemize}
-\item case ``0'': given $V(x)=0$ prove $P(x)$,
-\item case ``smaller'': given $V(x)>0$ and $\neg P(x)$ prove there exists a $y \in D$ such that $V(y)<V(x)$ and $~\neg P(y)$.
-\end{itemize}
-NB: the proof also shows how to use the previous lemma.\<close>
-
+  The method of infinite descent, frequently used in number theory.
+  Provided by Roelof Oosterhuis.
+  \<open>P n\<close> is true for all natural numbers if
+  \<^item> case ``0'': given \<open>n = 0\<close> prove \<open>P n\<close>
+  \<^item> case ``smaller'': given \<open>n > 0\<close> and \<open>\<not> P n\<close> prove there exists
+    a smaller natural number \<open>m\<close> such that \<open>\<not> P m\<close>.
+\<close>
+
+lemma infinite_descent: "(\<And>n. \<not> P n \<Longrightarrow> \<exists>m<n. \<not> P m) \<Longrightarrow> P n" for P :: "nat \<Rightarrow> bool"
+  \<comment> \<open>compact version without explicit base case\<close>
+  by (induct n rule: less_induct) auto
+
+lemma infinite_descent0 [case_names 0 smaller]:
+  fixes P :: "nat \<Rightarrow> bool"
+  assumes "P 0"
+    and "\<And>n. n > 0 \<Longrightarrow> \<not> P n \<Longrightarrow> \<exists>m. m < n \<and> \<not> P m"
+  shows "P n"
+  apply (rule infinite_descent)
+  using assms
+  apply (case_tac "n > 0")
+  apply auto
+  done
+
+text \<open>
+  Infinite descent using a mapping to \<open>nat\<close>:
+  \<open>P x\<close> is true for all \<open>x \<in> D\<close> if there exists a \<open>V \<in> D \<Rightarrow> nat\<close> and
+  \<^item> case ``0'': given \<open>V x = 0\<close> prove \<open>P x\<close>
+  \<^item> ``smaller'': given \<open>V x > 0\<close> and \<open>\<not> P x\<close> prove
+  there exists a \<open>y \<in> D\<close> such that \<open>V y < V x\<close> and \<open>\<not> P y\<close>.
+\<close>
 corollary infinite_descent0_measure [case_names 0 smaller]:
-  assumes A0: "!!x. V x = (0::nat) \<Longrightarrow> P x"
-    and   A1: "!!x. V x > 0 \<Longrightarrow> \<not>P x \<Longrightarrow> (\<exists>y. V y < V x \<and> \<not>P y)"
+  fixes V :: "'a \<Rightarrow> nat"
+  assumes 1: "\<And>x. V x = 0 \<Longrightarrow> P x"
+    and 2: "\<And>x. V x > 0 \<Longrightarrow> \<not> P x \<Longrightarrow> \<exists>y. V y < V x \<and> \<not> P y"
   shows "P x"
 proof -
   obtain n where "n = V x" by auto
   moreover have "\<And>x. V x = n \<Longrightarrow> P x"
   proof (induct n rule: infinite_descent0)
-    case 0 \<comment> "i.e. $V(x) = 0$"
-    with A0 show "P x" by auto
-  next \<comment> "now $n>0$ and $P(x)$ does not hold for some $x$ with $V(x)=n$"
+    case 0
+    with 1 show "P x" by auto
+  next
     case (smaller n)
-    then obtain x where vxn: "V x = n " and "V x > 0 \<and> \<not> P x" by auto
-    with A1 obtain y where "V y < V x \<and> \<not> P y" by auto
-    with vxn obtain m where "m = V y \<and> m<n \<and> \<not> P y" by auto
+    then obtain x where *: "V x = n " and "V x > 0 \<and> \<not> P x" by auto
+    with 2 obtain y where "V y < V x \<and> \<not> P y" by auto
+    with * obtain m where "m = V y \<and> m < n \<and> \<not> P y" by auto
     then show ?case by auto
   qed
   ultimately show "P x" by auto
@@ -998,96 +992,96 @@
 
 text\<open>Again, without explicit base case:\<close>
 lemma infinite_descent_measure:
-assumes "!!x. \<not> P x \<Longrightarrow> \<exists>y. (V::'a\<Rightarrow>nat) y < V x \<and> \<not> P y" shows "P x"
+  fixes V :: "'a \<Rightarrow> nat"
+  assumes "\<And>x. \<not> P x \<Longrightarrow> \<exists>y. V y < V x \<and> \<not> P y"
+  shows "P x"
 proof -
   from assms obtain n where "n = V x" by auto
-  moreover have "!!x. V x = n \<Longrightarrow> P x"
+  moreover have "\<And>x. V x = n \<Longrightarrow> P x"
   proof (induct n rule: infinite_descent, auto)
-    fix x assume "\<not> P x"
-    with assms show "\<exists>m < V x. \<exists>y. V y = m \<and> \<not> P y" by auto
+    show "\<exists>m < V x. \<exists>y. V y = m \<and> \<not> P y" if "\<not> P x" for x
+      using assms and that by auto
   qed
   ultimately show "P x" by auto
 qed
 
-text \<open>A [clumsy] way of lifting \<open><\<close>
-  monotonicity to \<open>\<le>\<close> monotonicity\<close>
+text \<open>A (clumsy) way of lifting \<open><\<close> monotonicity to \<open>\<le>\<close> monotonicity\<close>
 lemma less_mono_imp_le_mono:
-  "\<lbrakk> !!i j::nat. i < j \<Longrightarrow> f i < f j; i \<le> j \<rbrakk> \<Longrightarrow> f i \<le> ((f j)::nat)"
-by (simp add: order_le_less) (blast)
+  fixes f :: "nat \<Rightarrow> nat"
+    and i j :: nat
+  assumes "\<And>i j::nat. i < j \<Longrightarrow> f i < f j"
+    and "i \<le> j"
+  shows "f i \<le> f j"
+  using assms by (auto simp add: order_le_less)
 
 
 text \<open>non-strict, in 1st argument\<close>
-lemma add_le_mono1: "i \<le> j ==> i + k \<le> j + (k::nat)"
-by (rule add_right_mono)
+lemma add_le_mono1: "i \<le> j \<Longrightarrow> i + k \<le> j + k" for i j k :: nat
+  by (rule add_right_mono)
 
 text \<open>non-strict, in both arguments\<close>
-lemma add_le_mono: "[| i \<le> j;  k \<le> l |] ==> i + k \<le> j + (l::nat)"
-by (rule add_mono)
-
-lemma le_add2: "n \<le> ((m + n)::nat)"
+lemma add_le_mono: "i \<le> j \<Longrightarrow> k \<le> l \<Longrightarrow> i + k \<le> j + l" for i j k l :: nat
+  by (rule add_mono)
+
+lemma le_add2: "n \<le> m + n" for m n :: nat
   by simp
 
-lemma le_add1: "n \<le> ((n + m)::nat)"
+lemma le_add1: "n \<le> n + m" for m n :: nat
   by simp
 
 lemma less_add_Suc1: "i < Suc (i + m)"
-by (rule le_less_trans, rule le_add1, rule lessI)
+  by (rule le_less_trans, rule le_add1, rule lessI)
 
 lemma less_add_Suc2: "i < Suc (m + i)"
-by (rule le_less_trans, rule le_add2, rule lessI)
-
-lemma less_iff_Suc_add: "(m < n) = (\<exists>k. n = Suc (m + k))"
-by (iprover intro!: less_add_Suc1 less_imp_Suc_add)
-
-lemma trans_le_add1: "(i::nat) \<le> j ==> i \<le> j + m"
-by (rule le_trans, assumption, rule le_add1)
-
-lemma trans_le_add2: "(i::nat) \<le> j ==> i \<le> m + j"
-by (rule le_trans, assumption, rule le_add2)
-
-lemma trans_less_add1: "(i::nat) < j ==> i < j + m"
-by (rule less_le_trans, assumption, rule le_add1)
-
-lemma trans_less_add2: "(i::nat) < j ==> i < m + j"
-by (rule less_le_trans, assumption, rule le_add2)
-
-lemma add_lessD1: "i + j < (k::nat) ==> i < k"
-apply (rule le_less_trans [of _ "i+j"])
-apply (simp_all add: le_add1)
-done
-
-lemma not_add_less1 [iff]: "~ (i + j < (i::nat))"
-apply (rule notI)
-apply (drule add_lessD1)
-apply (erule less_irrefl [THEN notE])
-done
-
-lemma not_add_less2 [iff]: "~ (j + i < (i::nat))"
-by (simp add: add.commute)
-
-lemma add_leD1: "m + k \<le> n ==> m \<le> (n::nat)"
-apply (rule order_trans [of _ "m+k"])
-apply (simp_all add: le_add1)
-done
-
-lemma add_leD2: "m + k \<le> n ==> k \<le> (n::nat)"
-apply (simp add: add.commute)
-apply (erule add_leD1)
-done
-
-lemma add_leE: "(m::nat) + k \<le> n ==> (m \<le> n ==> k \<le> n ==> R) ==> R"
-by (blast dest: add_leD1 add_leD2)
-
-text \<open>needs \<open>!!k\<close> for \<open>ac_simps\<close> to work\<close>
-lemma less_add_eq_less: "!!k::nat. k < l ==> m + l = k + n ==> m < n"
-by (force simp del: add_Suc_right
-    simp add: less_iff_Suc_add add_Suc_right [symmetric] ac_simps)
+  by (rule le_less_trans, rule le_add2, rule lessI)
+
+lemma less_iff_Suc_add: "m < n \<longleftrightarrow> (\<exists>k. n = Suc (m + k))"
+  by (iprover intro!: less_add_Suc1 less_imp_Suc_add)
+
+lemma trans_le_add1: "i \<le> j \<Longrightarrow> i \<le> j + m" for i j m :: nat
+  by (rule le_trans, assumption, rule le_add1)
+
+lemma trans_le_add2: "i \<le> j \<Longrightarrow> i \<le> m + j" for i j m :: nat
+  by (rule le_trans, assumption, rule le_add2)
+
+lemma trans_less_add1: "i < j \<Longrightarrow> i < j + m" for i j m :: nat
+  by (rule less_le_trans, assumption, rule le_add1)
+
+lemma trans_less_add2: "i < j \<Longrightarrow> i < m + j" for i j m :: nat
+  by (rule less_le_trans, assumption, rule le_add2)
+
+lemma add_lessD1: "i + j < k \<Longrightarrow> i < k" for i j k :: nat
+  by (rule le_less_trans [of _ "i+j"]) (simp_all add: le_add1)
+
+lemma not_add_less1 [iff]: "\<not> i + j < i" for i j :: nat
+  apply (rule notI)
+  apply (drule add_lessD1)
+  apply (erule less_irrefl [THEN notE])
+  done
+
+lemma not_add_less2 [iff]: "\<not> j + i < i" for i j :: nat
+  by (simp add: add.commute)
+
+lemma add_leD1: "m + k \<le> n \<Longrightarrow> m \<le> n" for k m n :: nat
+  by (rule order_trans [of _ "m+k"]) (simp_all add: le_add1)
+
+lemma add_leD2: "m + k \<le> n \<Longrightarrow> k \<le> n" for k m n :: nat
+  apply (simp add: add.commute)
+  apply (erule add_leD1)
+  done
+
+lemma add_leE: "m + k \<le> n \<Longrightarrow> (m \<le> n \<Longrightarrow> k \<le> n \<Longrightarrow> R) \<Longrightarrow> R" for k m n :: nat
+  by (blast dest: add_leD1 add_leD2)
+
+text \<open>needs \<open>\<And>k\<close> for \<open>ac_simps\<close> to work\<close>
+lemma less_add_eq_less: "\<And>k::nat. k < l \<Longrightarrow> m + l = k + n \<Longrightarrow> m < n"
+  by (force simp del: add_Suc_right simp add: less_iff_Suc_add add_Suc_right [symmetric] ac_simps)
 
 
 subsubsection \<open>More results about difference\<close>
 
-lemma Suc_diff_le: "n \<le> m ==> Suc m - n = Suc (m - n)"
-by (induct m n rule: diff_induct) simp_all
+lemma Suc_diff_le: "n \<le> m \<Longrightarrow> Suc m - n = Suc (m - n)"
+  by (induct m n rule: diff_induct) simp_all
 
 lemma diff_less_Suc: "m - n < Suc m"
 apply (induct m n rule: diff_induct)
@@ -1095,80 +1089,72 @@
 apply (simp_all add: less_Suc_eq)
 done
 
-lemma diff_le_self [simp]: "m - n \<le> (m::nat)"
-by (induct m n rule: diff_induct) (simp_all add: le_SucI)
-
-lemma less_imp_diff_less: "(j::nat) < k ==> j - n < k"
-by (rule le_less_trans, rule diff_le_self)
-
-lemma diff_Suc_less [simp]: "0<n ==> n - Suc i < n"
-by (cases n) (auto simp add: le_simps)
-
-lemma diff_add_assoc: "k \<le> (j::nat) ==> (i + j) - k = i + (j - k)"
-by (induct j k rule: diff_induct) simp_all
-
-lemma add_diff_assoc [simp]:
-  fixes i j k :: nat
-  shows "k \<le> j \<Longrightarrow> i + (j - k) = i + j - k"
+lemma diff_le_self [simp]: "m - n \<le> m" for m n :: nat
+  by (induct m n rule: diff_induct) (simp_all add: le_SucI)
+
+lemma less_imp_diff_less: "j < k \<Longrightarrow> j - n < k" for j k n :: nat
+  by (rule le_less_trans, rule diff_le_self)
+
+lemma diff_Suc_less [simp]: "0 < n \<Longrightarrow> n - Suc i < n"
+  by (cases n) (auto simp add: le_simps)
+
+lemma diff_add_assoc: "k \<le> j \<Longrightarrow> (i + j) - k = i + (j - k)" for i j k :: nat
+  by (induct j k rule: diff_induct) simp_all
+
+lemma add_diff_assoc [simp]: "k \<le> j \<Longrightarrow> i + (j - k) = i + j - k" for i j k :: nat
   by (fact diff_add_assoc [symmetric])
 
-lemma diff_add_assoc2: "k \<le> (j::nat) ==> (j + i) - k = (j - k) + i"
+lemma diff_add_assoc2: "k \<le> j \<Longrightarrow> (j + i) - k = (j - k) + i" for i j k :: nat
   by (simp add: ac_simps)
 
-lemma add_diff_assoc2 [simp]:
-  fixes i j k :: nat
-  shows "k \<le> j \<Longrightarrow> j - k + i = j + i - k"
+lemma add_diff_assoc2 [simp]: "k \<le> j \<Longrightarrow> j - k + i = j + i - k" for i j k :: nat
   by (fact diff_add_assoc2 [symmetric])
 
-lemma le_imp_diff_is_add: "i \<le> (j::nat) ==> (j - i = k) = (j = k + i)"
-by auto
-
-lemma diff_is_0_eq [simp]: "((m::nat) - n = 0) = (m \<le> n)"
-by (induct m n rule: diff_induct) simp_all
-
-lemma diff_is_0_eq' [simp]: "m \<le> n ==> (m::nat) - n = 0"
-by (rule iffD2, rule diff_is_0_eq)
-
-lemma zero_less_diff [simp]: "(0 < n - (m::nat)) = (m < n)"
-by (induct m n rule: diff_induct) simp_all
+lemma le_imp_diff_is_add: "i \<le> j \<Longrightarrow> (j - i = k) = (j = k + i)" for i j k :: nat
+  by auto
+
+lemma diff_is_0_eq [simp]: "m - n = 0 \<longleftrightarrow> m \<le> n" for m n :: nat
+  by (induct m n rule: diff_induct) simp_all
+
+lemma diff_is_0_eq' [simp]: "m \<le> n \<Longrightarrow> m - n = 0" for m n :: nat
+  by (rule iffD2, rule diff_is_0_eq)
+
+lemma zero_less_diff [simp]: "0 < n - m \<longleftrightarrow> m < n" for m n :: nat
+  by (induct m n rule: diff_induct) simp_all
 
 lemma less_imp_add_positive:
   assumes "i < j"
-  shows "\<exists>k::nat. 0 < k & i + k = j"
+  shows "\<exists>k::nat. 0 < k \<and> i + k = j"
 proof
-  from assms show "0 < j - i & i + (j - i) = j"
+  from assms show "0 < j - i \<and> i + (j - i) = j"
     by (simp add: order_less_imp_le)
 qed
 
 text \<open>a nice rewrite for bounded subtraction\<close>
-lemma nat_minus_add_max:
-  fixes n m :: nat
-  shows "n - m + m = max n m"
+lemma nat_minus_add_max: "n - m + m = max n m" for m n :: nat
     by (simp add: max_def not_le order_less_imp_le)
 
-lemma nat_diff_split:
-  fixes a b :: nat
-  shows "P (a - b) \<longleftrightarrow> (a < b \<longrightarrow> P 0) \<and> (\<forall>d. a = b + d \<longrightarrow> P d)"
+lemma nat_diff_split: "P (a - b) \<longleftrightarrow> (a < b \<longrightarrow> P 0) \<and> (\<forall>d. a = b + d \<longrightarrow> P d)"
+  for a b :: nat
     \<comment> \<open>elimination of \<open>-\<close> on \<open>nat\<close>\<close>
   by (cases "a < b")
     (auto simp add: not_less le_less dest!: add_eq_self_zero [OF sym])
 
-lemma nat_diff_split_asm:
-  fixes a b :: nat
-  shows "P (a - b) \<longleftrightarrow> \<not> (a < b \<and> \<not> P 0 \<or> (\<exists>d. a = b + d \<and> \<not> P d))"
+lemma nat_diff_split_asm: "P (a - b) \<longleftrightarrow> \<not> (a < b \<and> \<not> P 0 \<or> (\<exists>d. a = b + d \<and> \<not> P d))"
+  for a b :: nat
     \<comment> \<open>elimination of \<open>-\<close> on \<open>nat\<close> in assumptions\<close>
   by (auto split: nat_diff_split)
 
-lemma Suc_pred': "0 < n ==> n = Suc(n - 1)"
+lemma Suc_pred': "0 < n \<Longrightarrow> n = Suc(n - 1)"
   by simp
 
-lemma add_eq_if: "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))"
+lemma add_eq_if: "m + n = (if m = 0 then n else Suc ((m - 1) + n))"
   unfolding One_nat_def by (cases m) simp_all
 
-lemma mult_eq_if: "(m::nat) * n = (if m=0 then 0 else n + ((m - 1) * n))"
+lemma mult_eq_if: "m * n = (if m = 0 then 0 else n + ((m - 1) * n))" for m n :: nat
   unfolding One_nat_def by (cases m) simp_all
 
-lemma Suc_diff_eq_diff_pred: "0 < n ==> Suc m - n = m - (n - 1)"
+lemma Suc_diff_eq_diff_pred: "0 < n \<Longrightarrow> Suc m - n = m - (n - 1)"
   unfolding One_nat_def by (cases n) simp_all
 
 lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
@@ -1180,68 +1166,68 @@
 
 subsubsection \<open>Monotonicity of multiplication\<close>
 
-lemma mult_le_mono1: "i \<le> (j::nat) ==> i * k \<le> j * k"
-by (simp add: mult_right_mono)
-
-lemma mult_le_mono2: "i \<le> (j::nat) ==> k * i \<le> k * j"
-by (simp add: mult_left_mono)
+lemma mult_le_mono1: "i \<le> j \<Longrightarrow> i * k \<le> j * k" for i j k :: nat
+  by (simp add: mult_right_mono)
+
+lemma mult_le_mono2: "i \<le> j \<Longrightarrow> k * i \<le> k * j" for i j k :: nat
+  by (simp add: mult_left_mono)
 
 text \<open>\<open>\<le>\<close> monotonicity, BOTH arguments\<close>
-lemma mult_le_mono: "i \<le> (j::nat) ==> k \<le> l ==> i * k \<le> j * l"
-by (simp add: mult_mono)
-
-lemma mult_less_mono1: "(i::nat) < j ==> 0 < k ==> i * k < j * k"
-by (simp add: mult_strict_right_mono)
+lemma mult_le_mono: "i \<le> j \<Longrightarrow> k \<le> l \<Longrightarrow> i * k \<le> j * l" for i j k l :: nat
+  by (simp add: mult_mono)
+
+lemma mult_less_mono1: "i < j \<Longrightarrow> 0 < k \<Longrightarrow> i * k < j * k" for i j k :: nat
+  by (simp add: mult_strict_right_mono)
 
 text\<open>Differs from the standard \<open>zero_less_mult_iff\<close> in that
       there are no negative numbers.\<close>
-lemma nat_0_less_mult_iff [simp]: "(0 < (m::nat) * n) = (0 < m & 0 < n)"
+lemma nat_0_less_mult_iff [simp]: "0 < m * n \<longleftrightarrow> 0 < m \<and> 0 < n" for m n :: nat
   apply (induct m)
    apply simp
   apply (case_tac n)
    apply simp_all
   done
 
-lemma one_le_mult_iff [simp]: "(Suc 0 \<le> m * n) = (Suc 0 \<le> m & Suc 0 \<le> n)"
+lemma one_le_mult_iff [simp]: "Suc 0 \<le> m * n \<longleftrightarrow> Suc 0 \<le> m \<and> Suc 0 \<le> n"
   apply (induct m)
    apply simp
   apply (case_tac n)
    apply simp_all
   done
 
-lemma mult_less_cancel2 [simp]: "((m::nat) * k < n * k) = (0 < k & m < n)"
+lemma mult_less_cancel2 [simp]: "m * k < n * k \<longleftrightarrow> 0 < k \<and> m < n" for k m n :: nat
   apply (safe intro!: mult_less_mono1)
   apply (cases k, auto)
-  apply (simp del: le_0_eq add: linorder_not_le [symmetric])
+  apply (simp add: linorder_not_le [symmetric])
   apply (blast intro: mult_le_mono1)
   done
 
-lemma mult_less_cancel1 [simp]: "(k * (m::nat) < k * n) = (0 < k & m < n)"
-by (simp add: mult.commute [of k])
-
-lemma mult_le_cancel1 [simp]: "(k * (m::nat) \<le> k * n) = (0 < k --> m \<le> n)"
-by (simp add: linorder_not_less [symmetric], auto)
-
-lemma mult_le_cancel2 [simp]: "((m::nat) * k \<le> n * k) = (0 < k --> m \<le> n)"
-by (simp add: linorder_not_less [symmetric], auto)
-
-lemma Suc_mult_less_cancel1: "(Suc k * m < Suc k * n) = (m < n)"
-by (subst mult_less_cancel1) simp
-
-lemma Suc_mult_le_cancel1: "(Suc k * m \<le> Suc k * n) = (m \<le> n)"
-by (subst mult_le_cancel1) simp
-
-lemma le_square: "m \<le> m * (m::nat)"
+lemma mult_less_cancel1 [simp]: "k * m < k * n \<longleftrightarrow> 0 < k \<and> m < n" for k m n :: nat
+  by (simp add: mult.commute [of k])
+
+lemma mult_le_cancel1 [simp]: "k * m \<le> k * n \<longleftrightarrow> (0 < k \<longrightarrow> m \<le> n)" for k m n :: nat
+  by (simp add: linorder_not_less [symmetric], auto)
+
+lemma mult_le_cancel2 [simp]: "m * k \<le> n * k \<longleftrightarrow> (0 < k \<longrightarrow> m \<le> n)" for k m n :: nat
+  by (simp add: linorder_not_less [symmetric], auto)
+
+lemma Suc_mult_less_cancel1: "Suc k * m < Suc k * n \<longleftrightarrow> m < n"
+  by (subst mult_less_cancel1) simp
+
+lemma Suc_mult_le_cancel1: "Suc k * m \<le> Suc k * n \<longleftrightarrow> m \<le> n"
+  by (subst mult_le_cancel1) simp
+
+lemma le_square: "m \<le> m * m" for m :: nat
   by (cases m) (auto intro: le_add1)
 
-lemma le_cube: "(m::nat) \<le> m * (m * m)"
+lemma le_cube: "m \<le> m * (m * m)" for m :: nat
   by (cases m) (auto intro: le_add1)
 
 text \<open>Lemma for \<open>gcd\<close>\<close>
-lemma mult_eq_self_implies_10: "(m::nat) = m * n ==> n = 1 | m = 0"
+lemma mult_eq_self_implies_10: "m = m * n \<Longrightarrow> n = 1 \<or> m = 0" for m n :: nat
   apply (drule sym)
   apply (rule disjCI)
-  apply (rule nat_less_cases, erule_tac [2] _)
+  apply (rule linorder_cases, erule_tac [2] _)
    apply (drule_tac [2] mult_less_mono2)
     apply (auto)
   done
@@ -1261,15 +1247,14 @@
 instantiation nat :: distrib_lattice
 begin
 
-definition
-  "(inf :: nat \<Rightarrow> nat \<Rightarrow> nat) = min"
-
-definition
-  "(sup :: nat \<Rightarrow> nat \<Rightarrow> nat) = max"
-
-instance by intro_classes
-  (auto simp add: inf_nat_def sup_nat_def max_def not_le min_def
-    intro: order_less_imp_le antisym elim!: order_trans order_less_trans)
+definition "(inf :: nat \<Rightarrow> nat \<Rightarrow> nat) = min"
+
+definition "(sup :: nat \<Rightarrow> nat \<Rightarrow> nat) = max"
+
+instance
+  by intro_classes
+    (auto simp add: inf_nat_def sup_nat_def max_def not_le min_def
+      intro: order_less_imp_le antisym elim!: order_trans order_less_trans)
 
 end
 
@@ -1283,8 +1268,8 @@
 
 consts compow :: "nat \<Rightarrow> 'a \<Rightarrow> 'a"
 
-abbreviation compower :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixr "^^" 80) where
-  "f ^^ n \<equiv> compow n f"
+abbreviation compower :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixr "^^" 80)
+  where "f ^^ n \<equiv> compow n f"
 
 notation (latex output)
   compower ("(_\<^bsup>_\<^esup>)" [1000] 1000)
@@ -1292,7 +1277,7 @@
 text \<open>\<open>f ^^ n = f o ... o f\<close>, the n-fold composition of \<open>f\<close>\<close>
 
 overloading
-  funpow == "compow :: nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a)"
+  funpow \<equiv> "compow :: nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a)"
 begin
 
 primrec funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
@@ -1304,10 +1289,10 @@
 lemma funpow_0 [simp]: "(f ^^ 0) x = x"
   by simp
 
-lemma funpow_Suc_right:
-  "f ^^ Suc n = f ^^ n \<circ> f"
+lemma funpow_Suc_right: "f ^^ Suc n = f ^^ n \<circ> f"
 proof (induct n)
-  case 0 then show ?case by simp
+  case 0
+  then show ?case by simp
 next
   fix n
   assume "f ^^ Suc n = f ^^ n \<circ> f"
@@ -1319,27 +1304,23 @@
 
 text \<open>for code generation\<close>
 
-definition funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
-  funpow_code_def [code_abbrev]: "funpow = compow"
+definition funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
+  where funpow_code_def [code_abbrev]: "funpow = compow"
 
 lemma [code]:
-  "funpow (Suc n) f = f o funpow n f"
+  "funpow (Suc n) f = f \<circ> funpow n f"
   "funpow 0 f = id"
   by (simp_all add: funpow_code_def)
 
 hide_const (open) funpow
 
-lemma funpow_add:
-  "f ^^ (m + n) = f ^^ m \<circ> f ^^ n"
+lemma funpow_add: "f ^^ (m + n) = f ^^ m \<circ> f ^^ n"
   by (induct m) simp_all
 
-lemma funpow_mult:
-  fixes f :: "'a \<Rightarrow> 'a"
-  shows "(f ^^ m) ^^ n = f ^^ (m * n)"
+lemma funpow_mult: "(f ^^ m) ^^ n = f ^^ (m * n)" for f :: "'a \<Rightarrow> 'a"
   by (induct n) (simp_all add: funpow_add)
 
-lemma funpow_swap1:
-  "f ((f ^^ n) x) = (f ^^ n) (f x)"
+lemma funpow_swap1: "f ((f ^^ n) x) = (f ^^ n) (f x)"
 proof -
   have "f ((f ^^ n) x) = (f ^^ (n + 1)) x" by simp
   also have "\<dots>  = (f ^^ n o f ^^ 1) x" by (simp only: funpow_add)
@@ -1347,9 +1328,7 @@
   finally show ?thesis .
 qed
 
-lemma comp_funpow:
-  fixes f :: "'a \<Rightarrow> 'a"
-  shows "comp f ^^ n = comp (f ^^ n)"
+lemma comp_funpow: "comp f ^^ n = comp (f ^^ n)" for f :: "'a \<Rightarrow> 'a"
   by (induct n) simp_all
 
 lemma Suc_funpow[simp]: "Suc ^^ n = (op + n)"
@@ -1358,43 +1337,46 @@
 lemma id_funpow[simp]: "id ^^ n = id"
   by (induct n) simp_all
 
-lemma funpow_mono:
-  fixes f :: "'a \<Rightarrow> ('a::lattice)"
-  shows "mono f \<Longrightarrow> A \<le> B \<Longrightarrow> (f ^^ n) A \<le> (f ^^ n) B"
+lemma funpow_mono: "mono f \<Longrightarrow> A \<le> B \<Longrightarrow> (f ^^ n) A \<le> (f ^^ n) B"
+  for f :: "'a \<Rightarrow> ('a::lattice)"
   by (induct n arbitrary: A B)
      (auto simp del: funpow.simps(2) simp add: funpow_Suc_right mono_def)
 
+
 subsection \<open>Kleene iteration\<close>
 
 lemma Kleene_iter_lpfp:
-assumes "mono f" and "f p \<le> p" shows "(f^^k) (bot::'a::order_bot) \<le> p"
+  assumes "mono f"
+    and "f p \<le> p"
+  shows "(f^^k) (bot::'a::order_bot) \<le> p"
 proof(induction k)
-  case 0 show ?case by simp
+  case 0
+  show ?case by simp
 next
   case Suc
-  from monoD[OF assms(1) Suc] assms(2)
-  show ?case by simp
+  from monoD[OF assms(1) Suc] assms(2) show ?case by simp
 qed
 
-lemma lfp_Kleene_iter: assumes "mono f" and "(f^^Suc k) bot = (f^^k) bot"
-shows "lfp f = (f^^k) bot"
-proof(rule antisym)
+lemma lfp_Kleene_iter:
+  assumes "mono f"
+    and "(f^^Suc k) bot = (f^^k) bot"
+  shows "lfp f = (f^^k) bot"
+proof (rule antisym)
   show "lfp f \<le> (f^^k) bot"
-  proof(rule lfp_lowerbound)
-    show "f ((f^^k) bot) \<le> (f^^k) bot" using assms(2) by simp
+  proof (rule lfp_lowerbound)
+    show "f ((f^^k) bot) \<le> (f^^k) bot"
+      using assms(2) by simp
   qed
-next
   show "(f^^k) bot \<le> lfp f"
     using Kleene_iter_lpfp[OF assms(1)] lfp_unfold[OF assms(1)] by simp
 qed
 
-lemma mono_pow:
-  fixes f :: "'a \<Rightarrow> 'a::complete_lattice"
-  shows "mono f \<Longrightarrow> mono (f ^^ n)"
-  by (induction n) (auto simp: mono_def)
+lemma mono_pow: "mono f \<Longrightarrow> mono (f ^^ n)" for f :: "'a \<Rightarrow> 'a::complete_lattice"
+  by (induct n) (auto simp: mono_def)
 
 lemma lfp_funpow:
-  assumes f: "mono f" shows "lfp (f ^^ Suc n) = lfp f"
+  assumes f: "mono f"
+  shows "lfp (f ^^ Suc n) = lfp f"
 proof (rule antisym)
   show "lfp f \<le> lfp (f ^^ Suc n)"
   proof (rule lfp_lowerbound)
@@ -1404,13 +1386,14 @@
       by (simp add: comp_def)
   qed
   have "(f^^n) (lfp f) = lfp f" for n
-    by (induction n) (auto intro: f lfp_unfold[symmetric])
+    by (induct n) (auto intro: f lfp_unfold[symmetric])
   then show "lfp (f^^Suc n) \<le> lfp f"
     by (intro lfp_lowerbound) (simp del: funpow.simps)
 qed
 
 lemma gfp_funpow:
-  assumes f: "mono f" shows "gfp (f ^^ Suc n) = gfp f"
+  assumes f: "mono f"
+  shows "gfp (f ^^ Suc n) = gfp f"
 proof (rule antisym)
   show "gfp f \<ge> gfp (f ^^ Suc n)"
   proof (rule gfp_upperbound)
@@ -1420,18 +1403,19 @@
       by (simp add: comp_def)
   qed
   have "(f^^n) (gfp f) = gfp f" for n
-    by (induction n) (auto intro: f gfp_unfold[symmetric])
+    by (induct n) (auto intro: f gfp_unfold[symmetric])
   then show "gfp (f^^Suc n) \<ge> gfp f"
     by (intro gfp_upperbound) (simp del: funpow.simps)
 qed
 
+
 subsection \<open>Embedding of the naturals into any \<open>semiring_1\<close>: @{term of_nat}\<close>
 
 context semiring_1
 begin
 
-definition of_nat :: "nat \<Rightarrow> 'a" where
-  "of_nat n = (plus 1 ^^ n) 0"
+definition of_nat :: "nat \<Rightarrow> 'a"
+  where "of_nat n = (plus 1 ^^ n) 0"
 
 lemma of_nat_simps [simp]:
   shows of_nat_0: "of_nat 0 = 0"
@@ -1448,16 +1432,16 @@
   by (induct m) (simp_all add: ac_simps distrib_right)
 
 lemma mult_of_nat_commute: "of_nat x * y = y * of_nat x"
-  by (induction x) (simp_all add: algebra_simps)
+  by (induct x) (simp_all add: algebra_simps)
 
 primrec of_nat_aux :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" where
   "of_nat_aux inc 0 i = i"
 | "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" \<comment> \<open>tail recursive\<close>
 
-lemma of_nat_code:
-  "of_nat n = of_nat_aux (\<lambda>i. i + 1) n 0"
+lemma of_nat_code: "of_nat n = of_nat_aux (\<lambda>i. i + 1) n 0"
 proof (induct n)
-  case 0 then show ?case by simp
+  case 0
+  then show ?case by simp
 next
   case (Suc n)
   have "\<And>i. of_nat_aux (\<lambda>i. i + 1) n (i + 1) = of_nat_aux (\<lambda>i. i + 1) n i + 1"
@@ -1475,11 +1459,11 @@
 begin
 
 lemma of_nat_diff: "n \<le> m \<Longrightarrow> of_nat (m - n) = of_nat m - of_nat n"
-by (simp add: algebra_simps of_nat_add [symmetric])
+  by (simp add: algebra_simps of_nat_add [symmetric])
 
 end
 
-text\<open>Class for unital semirings with characteristic zero.
+text \<open>Class for unital semirings with characteristic zero.
  Includes non-ordered rings like the complex numbers.\<close>
 
 class semiring_char_0 = semiring_1 +
@@ -1489,7 +1473,7 @@
 lemma of_nat_eq_iff [simp]: "of_nat m = of_nat n \<longleftrightarrow> m = n"
   by (auto intro: inj_of_nat injD)
 
-text\<open>Special cases where either operand is zero\<close>
+text \<open>Special cases where either operand is zero\<close>
 
 lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \<longleftrightarrow> 0 = n"
   by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0])
@@ -1530,12 +1514,12 @@
 lemma of_nat_less_imp_less: "of_nat m < of_nat n \<Longrightarrow> m < n"
   by simp
 
-text\<open>Every \<open>linordered_semidom\<close> has characteristic zero.\<close>
-
-subclass semiring_char_0 proof
-qed (auto intro!: injI simp add: eq_iff)
-
-text\<open>Special cases where either operand is zero\<close>
+text \<open>Every \<open>linordered_semidom\<close> has characteristic zero.\<close>
+
+subclass semiring_char_0
+  by standard (auto intro!: injI simp add: eq_iff)
+
+text \<open>Special cases where either operand is zero\<close>
 
 lemma of_nat_le_0_iff [simp]: "of_nat m \<le> 0 \<longleftrightarrow> m = 0"
   by (rule of_nat_le_iff [of _ 0, simplified])
@@ -1615,9 +1599,9 @@
 subsection \<open>Further arithmetic facts concerning the natural numbers\<close>
 
 lemma subst_equals:
-  assumes 1: "t = s" and 2: "u = t"
+  assumes "t = s" and "u = t"
   shows "u = s"
-  using 2 1 by (rule trans)
+  using assms(2,1) by (rule trans)
 
 ML_file "Tools/nat_arith.ML"
 
@@ -1647,7 +1631,10 @@
   case True
   then show ?thesis
     by (induct n n' rule: less_Suc_induct) (auto intro: mono)
-qed (insert \<open>n \<le> n'\<close>, auto) \<comment> \<open>trivial for @{prop "n = n'"}\<close>
+next
+  case False
+  with \<open>n \<le> n'\<close> show ?thesis by auto
+qed
 
 lemma lift_Suc_antimono_le:
   assumes mono: "\<And>n. f n \<ge> f (Suc n)" and "n \<le> n'"
@@ -1656,27 +1643,26 @@
   case True
   then show ?thesis
     by (induct n n' rule: less_Suc_induct) (auto intro: mono)
-qed (insert \<open>n \<le> n'\<close>, auto) \<comment> \<open>trivial for @{prop "n = n'"}\<close>
+next
+  case False
+  with \<open>n \<le> n'\<close> show ?thesis by auto
+qed
 
 lemma lift_Suc_mono_less:
   assumes mono: "\<And>n. f n < f (Suc n)" and "n < n'"
   shows "f n < f n'"
-using \<open>n < n'\<close>
-by (induct n n' rule: less_Suc_induct) (auto intro: mono)
-
-lemma lift_Suc_mono_less_iff:
-  "(\<And>n. f n < f (Suc n)) \<Longrightarrow> f n < f m \<longleftrightarrow> n < m"
+  using \<open>n < n'\<close> by (induct n n' rule: less_Suc_induct) (auto intro: mono)
+
+lemma lift_Suc_mono_less_iff: "(\<And>n. f n < f (Suc n)) \<Longrightarrow> f n < f m \<longleftrightarrow> n < m"
   by (blast intro: less_asym' lift_Suc_mono_less [of f]
     dest: linorder_not_less[THEN iffD1] le_eq_less_or_eq [THEN iffD1])
 
 end
 
-lemma mono_iff_le_Suc:
-  "mono f \<longleftrightarrow> (\<forall>n. f n \<le> f (Suc n))"
+lemma mono_iff_le_Suc: "mono f \<longleftrightarrow> (\<forall>n. f n \<le> f (Suc n))"
   unfolding mono_def by (auto intro: lift_Suc_mono_le [of f])
 
-lemma antimono_iff_le_Suc:
-  "antimono f \<longleftrightarrow> (\<forall>n. f (Suc n) \<le> f n)"
+lemma antimono_iff_le_Suc: "antimono f \<longleftrightarrow> (\<forall>n. f (Suc n) \<le> f n)"
   unfolding antimono_def by (auto intro: lift_Suc_antimono_le [of f])
 
 lemma mono_nat_linear_lb:
@@ -1684,7 +1670,8 @@
   assumes "\<And>m n. m < n \<Longrightarrow> f m < f n"
   shows "f m + k \<le> f (m + k)"
 proof (induct k)
-  case 0 then show ?case by simp
+  case 0
+  then show ?case by simp
 next
   case (Suc k)
   then have "Suc (f m + k) \<le> Suc (f (m + k))" by simp
@@ -1694,7 +1681,7 @@
 qed
 
 
-text\<open>Subtraction laws, mostly by Clemens Ballarin\<close>
+text \<open>Subtraction laws, mostly by Clemens Ballarin\<close>
 
 lemma diff_less_mono:
   fixes a b c :: nat
@@ -1706,105 +1693,73 @@
   then show ?thesis by simp
 qed
 
-lemma less_diff_conv:
-  fixes i j k :: nat
-  shows "i < j - k \<longleftrightarrow> i + k < j" (is "?P \<longleftrightarrow> ?Q")
-  by (cases "k \<le> j")
-    (auto simp add: not_le dest: less_imp_Suc_add le_Suc_ex)
-
-lemma less_diff_conv2:
-  fixes j k i :: nat
-  assumes "k \<le> j"
-  shows "j - k < i \<longleftrightarrow> j < i + k" (is "?P \<longleftrightarrow> ?Q")
-  using assms by (auto dest: le_Suc_ex)
-
-lemma le_diff_conv:
-  fixes j k i :: nat
-  shows "j - k \<le> i \<longleftrightarrow> j \<le> i + k"
-  by (cases "k \<le> j")
-    (auto simp add: not_le dest!: less_imp_Suc_add le_Suc_ex)
-
-lemma diff_diff_cancel [simp]:
-  fixes i n :: nat
-  shows "i \<le> n \<Longrightarrow> n - (n - i) = i"
+lemma less_diff_conv: "i < j - k \<longleftrightarrow> i + k < j" for i j k :: nat
+  by (cases "k \<le> j") (auto simp add: not_le dest: less_imp_Suc_add le_Suc_ex)
+
+lemma less_diff_conv2: "k \<le> j \<Longrightarrow> j - k < i \<longleftrightarrow> j < i + k" for j k i :: nat
   by (auto dest: le_Suc_ex)
 
-lemma diff_less [simp]:
-  fixes i n :: nat
-  shows "0 < n \<Longrightarrow> 0 < m \<Longrightarrow> m - n < m"
+lemma le_diff_conv: "j - k \<le> i \<longleftrightarrow> j \<le> i + k" for j k i :: nat
+  by (cases "k \<le> j") (auto simp add: not_le dest!: less_imp_Suc_add le_Suc_ex)
+
+lemma diff_diff_cancel [simp]: "i \<le> n \<Longrightarrow> n - (n - i) = i" for i n :: nat
+  by (auto dest: le_Suc_ex)
+
+lemma diff_less [simp]: "0 < n \<Longrightarrow> 0 < m \<Longrightarrow> m - n < m" for i n :: nat
   by (auto dest: less_imp_Suc_add)
 
 text \<open>Simplification of relational expressions involving subtraction\<close>
 
-lemma diff_diff_eq:
-  fixes m n k :: nat
-  shows "k \<le> m \<Longrightarrow> k \<le> n \<Longrightarrow> m - k - (n - k) = m - n"
+lemma diff_diff_eq: "k \<le> m \<Longrightarrow> k \<le> n \<Longrightarrow> m - k - (n - k) = m - n" for m n k :: nat
   by (auto dest!: le_Suc_ex)
 
 hide_fact (open) diff_diff_eq
 
-lemma eq_diff_iff:
-  fixes m n k :: nat
-  shows "k \<le> m \<Longrightarrow> k \<le> n \<Longrightarrow> m - k = n - k \<longleftrightarrow> m = n"
+lemma eq_diff_iff: "k \<le> m \<Longrightarrow> k \<le> n \<Longrightarrow> m - k = n - k \<longleftrightarrow> m = n" for m n k :: nat
   by (auto dest: le_Suc_ex)
 
-lemma less_diff_iff:
-  fixes m n k :: nat
-  shows "k \<le> m \<Longrightarrow> k \<le> n \<Longrightarrow> m - k < n - k \<longleftrightarrow> m < n"
+lemma less_diff_iff: "k \<le> m \<Longrightarrow> k \<le> n \<Longrightarrow> m - k < n - k \<longleftrightarrow> m < n" for m n k :: nat
   by (auto dest!: le_Suc_ex)
 
-lemma le_diff_iff:
-  fixes m n k :: nat
-  shows "k \<le> m \<Longrightarrow> k \<le> n \<Longrightarrow> m - k \<le> n - k \<longleftrightarrow> m \<le> n"
+lemma le_diff_iff: "k \<le> m \<Longrightarrow> k \<le> n \<Longrightarrow> m - k \<le> n - k \<longleftrightarrow> m \<le> n" for m n k :: nat
   by (auto dest!: le_Suc_ex)
 
-lemma le_diff_iff': "a \<le> c \<Longrightarrow> b \<le> c \<Longrightarrow> c - a \<le> c - b \<longleftrightarrow> b \<le> (a::nat)"
+lemma le_diff_iff': "a \<le> c \<Longrightarrow> b \<le> c \<Longrightarrow> c - a \<le> c - b \<longleftrightarrow> b \<le> a" for a b c :: nat
   by (force dest: le_Suc_ex)
-  
-
-text\<open>(Anti)Monotonicity of subtraction -- by Stephan Merz\<close>
-
-lemma diff_le_mono:
-  fixes m n l :: nat
-  shows "m \<le> n \<Longrightarrow> m - l \<le> n - l"
+
+
+text \<open>(Anti)Monotonicity of subtraction -- by Stephan Merz\<close>
+
+lemma diff_le_mono: "m \<le> n \<Longrightarrow> m - l \<le> n - l" for m n l :: nat
   by (auto dest: less_imp_le less_imp_Suc_add split add: nat_diff_split)
 
-lemma diff_le_mono2:
-  fixes m n l :: nat
-  shows "m \<le> n \<Longrightarrow> l - n \<le> l - m"
+lemma diff_le_mono2: "m \<le> n \<Longrightarrow> l - n \<le> l - m" for m n l :: nat
   by (auto dest: less_imp_le le_Suc_ex less_imp_Suc_add less_le_trans split add: nat_diff_split)
 
-lemma diff_less_mono2:
-  fixes m n l :: nat
-  shows "m < n \<Longrightarrow> m < l \<Longrightarrow> l - n < l - m"
+lemma diff_less_mono2: "m < n \<Longrightarrow> m < l \<Longrightarrow> l - n < l - m" for m n l :: nat
   by (auto dest: less_imp_Suc_add split add: nat_diff_split)
 
-lemma diffs0_imp_equal:
-  fixes m n :: nat
-  shows "m - n = 0 \<Longrightarrow> n - m = 0 \<Longrightarrow> m = n"
+lemma diffs0_imp_equal: "m - n = 0 \<Longrightarrow> n - m = 0 \<Longrightarrow> m = n" for m n :: nat
   by (simp split add: nat_diff_split)
 
-lemma min_diff:
-  fixes m n i :: nat
-  shows "min (m - i) (n - i) = min m n - i" (is "?lhs = ?rhs")
+lemma min_diff: "min (m - i) (n - i) = min m n - i" for m n i :: nat
   by (cases m n rule: le_cases)
     (auto simp add: not_le min.absorb1 min.absorb2 min.absorb_iff1 [symmetric] diff_le_mono)
 
 lemma inj_on_diff_nat:
-  assumes k_le_n: "\<forall>n \<in> N. k \<le> (n::nat)"
+  fixes k :: nat
+  assumes "\<forall>n \<in> N. k \<le> n"
   shows "inj_on (\<lambda>n. n - k) N"
 proof (rule inj_onI)
   fix x y
   assume a: "x \<in> N" "y \<in> N" "x - k = y - k"
-  with k_le_n have "x - k + k = y - k + k" by auto
-  with a k_le_n show "x = y" by (auto simp add: eq_diff_iff)
+  with assms have "x - k + k = y - k + k" by auto
+  with a assms show "x = y" by (auto simp add: eq_diff_iff)
 qed
 
-text\<open>Rewriting to pull differences out\<close>
-
-lemma diff_diff_right [simp]:
-  fixes i j k :: nat
-  shows "k \<le> j \<Longrightarrow> i - (j - k) = i + k - j"
+text \<open>Rewriting to pull differences out\<close>
+
+lemma diff_diff_right [simp]: "k \<le> j \<Longrightarrow> i - (j - k) = i + k - j" for i j k :: nat
   by (fact diff_diff_right)
 
 lemma diff_Suc_diff_eq1 [simp]:
@@ -1842,24 +1797,22 @@
   then show ?thesis by simp
 qed
 
-lemma one_less_mult:
-  "Suc 0 < n \<Longrightarrow> Suc 0 < m \<Longrightarrow> Suc 0 < m * n"
+lemma one_less_mult: "Suc 0 < n \<Longrightarrow> Suc 0 < m \<Longrightarrow> Suc 0 < m * n"
   using less_1_mult [of n m] by (simp add: ac_simps)
 
-lemma n_less_m_mult_n:
-  "0 < n \<Longrightarrow> Suc 0 < m \<Longrightarrow> n < m * n"
+lemma n_less_m_mult_n: "0 < n \<Longrightarrow> Suc 0 < m \<Longrightarrow> n < m * n"
   using mult_strict_right_mono [of 1 m n] by simp
 
-lemma n_less_n_mult_m:
-  "0 < n \<Longrightarrow> Suc 0 < m \<Longrightarrow> n < n * m"
+lemma n_less_n_mult_m: "0 < n \<Longrightarrow> Suc 0 < m \<Longrightarrow> n < n * m"
   using mult_strict_left_mono [of 1 m n] by simp
 
+
 text \<open>Specialized induction principles that work "backwards":\<close>
 
 lemma inc_induct [consumes 1, case_names base step]:
   assumes less: "i \<le> j"
-  assumes base: "P j"
-  assumes step: "\<And>n. i \<le> n \<Longrightarrow> n < j \<Longrightarrow> P (Suc n) \<Longrightarrow> P n"
+    and base: "P j"
+    and step: "\<And>n. i \<le> n \<Longrightarrow> n < j \<Longrightarrow> P (Suc n) \<Longrightarrow> P n"
   shows "P i"
   using less step
 proof (induct "j - i" arbitrary: i)
@@ -1873,7 +1826,7 @@
   from \<open>Suc d = j - n\<close> have "d + 1 = j - n" by simp
   then have "d + 1 - 1 = j - n - 1" by simp
   then have "d = j - n - 1" by simp
-  then have "d = j - (n + 1)" 
+  then have "d = j - (n + 1)"
     by (simp add: diff_diff_eq)
   then have "d = j - Suc n"
     by simp
@@ -1892,11 +1845,11 @@
   with order_refl \<open>n < j\<close> show "P n"
     by (rule Suc.prems)
 qed
-    
+
 lemma strict_inc_induct [consumes 1, case_names base step]:
   assumes less: "i < j"
-  assumes base: "\<And>i. j = Suc i \<Longrightarrow> P i"
-  assumes step: "\<And>i. i < j \<Longrightarrow> P (Suc i) \<Longrightarrow> P i"
+    and base: "\<And>i. j = Suc i \<Longrightarrow> P i"
+    and step: "\<And>i. i < j \<Longrightarrow> P (Suc i) \<Longrightarrow> P i"
   shows "P i"
 using less proof (induct "j - i - 1" arbitrary: i)
   case (0 i)
@@ -1922,10 +1875,10 @@
   with \<open>i < j\<close> show "P i" by (rule step)
 qed
 
-lemma zero_induct_lemma: "P k ==> (!!n. P (Suc n) ==> P n) ==> P (k - i)"
+lemma zero_induct_lemma: "P k \<Longrightarrow> (\<And>n. P (Suc n) \<Longrightarrow> P n) \<Longrightarrow> P (k - i)"
   using inc_induct[of "k - i" k P, simplified] by blast
 
-lemma zero_induct: "P k ==> (!!n. P (Suc n) ==> P n) ==> P 0"
+lemma zero_induct: "P k \<Longrightarrow> (\<And>n. P (Suc n) \<Longrightarrow> P n) \<Longrightarrow> P 0"
   using inc_induct[of 0 k P] by blast
 
 text \<open>Further induction rule similar to @{thm inc_induct}\<close>
@@ -1933,13 +1886,15 @@
 lemma dec_induct [consumes 1, case_names base step]:
   "i \<le> j \<Longrightarrow> P i \<Longrightarrow> (\<And>n. i \<le> n \<Longrightarrow> n < j \<Longrightarrow> P n \<Longrightarrow> P (Suc n)) \<Longrightarrow> P j"
 proof (induct j arbitrary: i)
-  case 0 then show ?case by simp
+  case 0
+  then show ?case by simp
 next
   case (Suc j)
-  from Suc.prems have "i \<le> j \<or> i = Suc j"
-    by (simp add: le_Suc_eq)
-  then show ?case proof
-    assume "i \<le> j"
+  from Suc.prems consider "i \<le> j" | "i = Suc j"
+    by (auto simp add: le_Suc_eq)
+  then show ?case
+  proof cases
+    case 1
     moreover have "j < Suc j" by simp
     moreover have "P j" using \<open>i \<le> j\<close> \<open>P i\<close>
     proof (rule Suc.hyps)
@@ -1954,70 +1909,62 @@
     ultimately show "P (Suc j)"
       by (rule Suc.prems)
   next
-    assume "i = Suc j"
+    case 2
     with \<open>P i\<close> show "P (Suc j)" by simp
   qed
 qed
 
 
-subsection \<open> Monotonicity of funpow \<close>
+subsection \<open>Monotonicity of \<open>funpow\<close>\<close>
 
 lemma funpow_increasing:
-  fixes f :: "'a \<Rightarrow> ('a::{lattice, order_top})"
+  fixes f :: "'a \<Rightarrow> 'a::{lattice,order_top}"
   shows "m \<le> n \<Longrightarrow> mono f \<Longrightarrow> (f ^^ n) \<top> \<le> (f ^^ m) \<top>"
   by (induct rule: inc_induct)
      (auto simp del: funpow.simps(2) simp add: funpow_Suc_right
            intro: order_trans[OF _ funpow_mono])
 
 lemma funpow_decreasing:
-  fixes f :: "'a \<Rightarrow> ('a::{lattice, order_bot})"
+  fixes f :: "'a \<Rightarrow> 'a::{lattice,order_bot}"
   shows "m \<le> n \<Longrightarrow> mono f \<Longrightarrow> (f ^^ m) \<bottom> \<le> (f ^^ n) \<bottom>"
   by (induct rule: dec_induct)
      (auto simp del: funpow.simps(2) simp add: funpow_Suc_right
            intro: order_trans[OF _ funpow_mono])
 
 lemma mono_funpow:
-  fixes Q :: "'a::{lattice, order_bot} \<Rightarrow> 'a"
+  fixes Q :: "'a::{lattice,order_bot} \<Rightarrow> 'a"
   shows "mono Q \<Longrightarrow> mono (\<lambda>i. (Q ^^ i) \<bottom>)"
   by (auto intro!: funpow_decreasing simp: mono_def)
 
 lemma antimono_funpow:
-  fixes Q :: "'a::{lattice, order_top} \<Rightarrow> 'a"
+  fixes Q :: "'a::{lattice,order_top} \<Rightarrow> 'a"
   shows "mono Q \<Longrightarrow> antimono (\<lambda>i. (Q ^^ i) \<top>)"
   by (auto intro!: funpow_increasing simp: antimono_def)
 
+
 subsection \<open>The divides relation on @{typ nat}\<close>
 
-lemma dvd_1_left [iff]:
-  "Suc 0 dvd k"
+lemma dvd_1_left [iff]: "Suc 0 dvd k"
   by (simp add: dvd_def)
 
-lemma dvd_1_iff_1 [simp]:
-  "m dvd Suc 0 \<longleftrightarrow> m = Suc 0"
+lemma dvd_1_iff_1 [simp]: "m dvd Suc 0 \<longleftrightarrow> m = Suc 0"
   by (simp add: dvd_def)
 
-lemma nat_dvd_1_iff_1 [simp]:
-  "m dvd (1::nat) \<longleftrightarrow> m = 1"
+lemma nat_dvd_1_iff_1 [simp]: "m dvd 1 \<longleftrightarrow> m = 1" for m :: nat
   by (simp add: dvd_def)
 
-lemma dvd_antisym:
-  "m dvd n \<Longrightarrow> n dvd m \<Longrightarrow> m = (n::nat)"
-  unfolding dvd_def
-  by (force dest: mult_eq_self_implies_10 simp add: mult.assoc)
-
-lemma dvd_diff_nat [simp]:
-  "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m - n :: nat)"
-  unfolding dvd_def
-  by (blast intro: right_diff_distrib' [symmetric])
-
-lemma dvd_diffD:
-  "k dvd m - n \<Longrightarrow> k dvd n \<Longrightarrow> n \<le> m \<Longrightarrow> k dvd (m::nat)"
+lemma dvd_antisym: "m dvd n \<Longrightarrow> n dvd m \<Longrightarrow> m = n" for m n :: nat
+  unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult.assoc)
+
+lemma dvd_diff_nat [simp]: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m - n)" for k m n :: nat
+  unfolding dvd_def by (blast intro: right_diff_distrib' [symmetric])
+
+lemma dvd_diffD: "k dvd m - n \<Longrightarrow> k dvd n \<Longrightarrow> n \<le> m \<Longrightarrow> k dvd m" for k m n :: nat
   apply (erule linorder_not_less [THEN iffD2, THEN add_diff_inverse, THEN subst])
   apply (blast intro: dvd_add)
   done
 
-lemma dvd_diffD1:
-  "k dvd m - n \<Longrightarrow> k dvd m \<Longrightarrow> n \<le> m \<Longrightarrow> k dvd (n::nat)"
+lemma dvd_diffD1: "k dvd m - n \<Longrightarrow> k dvd m \<Longrightarrow> n \<le> m \<Longrightarrow> k dvd n" for k m n :: nat
   by (drule_tac m = m in dvd_diff_nat) auto
 
 lemma dvd_mult_cancel:
@@ -2030,25 +1977,20 @@
   with \<open>0 < k\<close> have "n = m * q" by (auto simp add: mult_left_cancel)
   then show ?thesis ..
 qed
-  
-lemma dvd_mult_cancel1:
-  "0 < m \<Longrightarrow> m * n dvd m \<longleftrightarrow> n = (1::nat)"
+
+lemma dvd_mult_cancel1: "0 < m \<Longrightarrow> m * n dvd m \<longleftrightarrow> n = 1" for m n :: nat
   apply auto
-   apply (subgoal_tac "m*n dvd m*1")
+   apply (subgoal_tac "m * n dvd m * 1")
    apply (drule dvd_mult_cancel, auto)
   done
 
-lemma dvd_mult_cancel2:
-  "0 < m \<Longrightarrow> n * m dvd m \<longleftrightarrow> n = (1::nat)"
+lemma dvd_mult_cancel2: "0 < m \<Longrightarrow> n * m dvd m \<longleftrightarrow> n = 1" for m n :: nat
   using dvd_mult_cancel1 [of m n] by (simp add: ac_simps)
 
-lemma dvd_imp_le:
-  "k dvd n \<Longrightarrow> 0 < n \<Longrightarrow> k \<le> (n::nat)"
+lemma dvd_imp_le: "k dvd n \<Longrightarrow> 0 < n \<Longrightarrow> k \<le> n" for k n :: nat
   by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc)
 
-lemma nat_dvd_not_less:
-  fixes m n :: nat
-  shows "0 < m \<Longrightarrow> m < n \<Longrightarrow> \<not> n dvd m"
+lemma nat_dvd_not_less: "0 < m \<Longrightarrow> m < n \<Longrightarrow> \<not> n dvd m" for m n :: nat
   by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc)
 
 lemma less_eq_dvd_minus:
@@ -2061,9 +2003,7 @@
   then show ?thesis by (simp add: add.commute [of m])
 qed
 
-lemma dvd_minus_self:
-  fixes m n :: nat
-  shows "m dvd n - m \<longleftrightarrow> n < m \<or> m dvd n"
+lemma dvd_minus_self: "m dvd n - m \<longleftrightarrow> n < m \<or> m dvd n" for m n :: nat
   by (cases "n < m") (auto elim!: dvdE simp add: not_less le_imp_diff_is_add dest: less_imp_le)
 
 lemma dvd_minus_add:
@@ -2082,61 +2022,55 @@
 
 subsection \<open>Aliasses\<close>
 
-lemma nat_mult_1: "(1::nat) * n = n"
+lemma nat_mult_1: "1 * n = n" for n :: nat
   by (fact mult_1_left)
 
-lemma nat_mult_1_right: "n * (1::nat) = n"
+lemma nat_mult_1_right: "n * 1 = n" for n :: nat
   by (fact mult_1_right)
 
-lemma nat_add_left_cancel:
-  fixes k m n :: nat
-  shows "k + m = k + n \<longleftrightarrow> m = n"
+lemma nat_add_left_cancel: "k + m = k + n \<longleftrightarrow> m = n" for k m n :: nat
   by (fact add_left_cancel)
 
-lemma nat_add_right_cancel:
-  fixes k m n :: nat
-  shows "m + k = n + k \<longleftrightarrow> m = n"
+lemma nat_add_right_cancel: "m + k = n + k \<longleftrightarrow> m = n" for k m n :: nat
   by (fact add_right_cancel)
 
-lemma diff_mult_distrib:
-  "((m::nat) - n) * k = (m * k) - (n * k)"
+lemma diff_mult_distrib: "(m - n) * k = (m * k) - (n * k)" for k m n :: nat
   by (fact left_diff_distrib')
 
-lemma diff_mult_distrib2:
-  "k * ((m::nat) - n) = (k * m) - (k * n)"
+lemma diff_mult_distrib2: "k * (m - n) = (k * m) - (k * n)" for k m n :: nat
   by (fact right_diff_distrib')
 
-lemma le_add_diff: "k \<le> (n::nat) ==> m \<le> n + m - k"
-  by (fact le_add_diff) \<comment> \<open>FIXME delete\<close>
-
-lemma le_diff_conv2: "k \<le> j ==> (i \<le> j-k) = (i+k \<le> (j::nat))"
-  by (fact le_diff_conv2) \<comment> \<open>FIXME delete\<close>
-
-lemma diff_self_eq_0 [simp]: "(m::nat) - m = 0"
+lemma le_add_diff: "k \<le> n \<Longrightarrow> m \<le> n + m - k" for k m n :: nat
+  by (fact le_add_diff)  (* FIXME delete *)
+
+lemma le_diff_conv2: "k \<le> j \<Longrightarrow> (i \<le> j - k) = (i + k \<le> j)" for i j k :: nat
+  by (fact le_diff_conv2) (* FIXME delete *)
+
+lemma diff_self_eq_0 [simp]: "m - m = 0" for m :: nat
   by (fact diff_cancel)
 
-lemma diff_diff_left [simp]: "(i::nat) - j - k = i - (j + k)"
+lemma diff_diff_left [simp]: "i - j - k = i - (j + k)" for i j k :: nat
   by (fact diff_diff_add)
 
-lemma diff_commute: "(i::nat) - j - k = i - k - j"
+lemma diff_commute: "i - j - k = i - k - j" for i j k :: nat
   by (fact diff_right_commute)
 
-lemma diff_add_inverse: "(n + m) - n = (m::nat)"
+lemma diff_add_inverse: "(n + m) - n = m" for m n :: nat
   by (fact add_diff_cancel_left')
 
-lemma diff_add_inverse2: "(m + n) - n = (m::nat)"
+lemma diff_add_inverse2: "(m + n) - n = m" for m n :: nat
   by (fact add_diff_cancel_right')
 
-lemma diff_cancel: "(k + m) - (k + n) = m - (n::nat)"
+lemma diff_cancel: "(k + m) - (k + n) = m - n" for k m n :: nat
   by (fact add_diff_cancel_left)
 
-lemma diff_cancel2: "(m + k) - (n + k) = m - (n::nat)"
+lemma diff_cancel2: "(m + k) - (n + k) = m - n" for k m n :: nat
   by (fact add_diff_cancel_right)
 
-lemma diff_add_0: "n - (n + m) = (0::nat)"
+lemma diff_add_0: "n - (n + m) = 0" for m n :: nat
   by (fact diff_add_zero)
 
-lemma add_mult_distrib2: "k * (m + n) = (k * m) + ((k * n)::nat)"
+lemma add_mult_distrib2: "k * (m + n) = (k * m) + (k * n)" for k m n :: nat
   by (fact distrib_left)
 
 lemmas nat_distrib =
@@ -2151,8 +2085,7 @@
 instantiation nat :: size
 begin
 
-definition size_nat where
-  [simp, code]: "size (n::nat) = n"
+definition size_nat where [simp, code]: "size (n::nat) = n"
 
 instance ..
 
--- a/src/HOL/Wellfounded.thy	Mon May 23 15:33:24 2016 +0100
+++ b/src/HOL/Wellfounded.thy	Mon May 23 15:46:30 2016 +0100
@@ -14,34 +14,33 @@
 
 subsection \<open>Basic Definitions\<close>
 
-definition wf :: "('a * 'a) set => bool" where
-  "wf r \<longleftrightarrow> (!P. (!x. (!y. (y,x):r --> P(y)) --> P(x)) --> (!x. P(x)))"
+definition wf :: "('a \<times> 'a) set \<Rightarrow> bool"
+  where "wf r \<longleftrightarrow> (\<forall>P. (\<forall>x. (\<forall>y. (y, x) \<in> r \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> (\<forall>x. P x))"
 
-definition wfP :: "('a => 'a => bool) => bool" where
-  "wfP r \<longleftrightarrow> wf {(x, y). r x y}"
+definition wfP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
+  where "wfP r \<longleftrightarrow> wf {(x, y). r x y}"
 
 lemma wfP_wf_eq [pred_set_conv]: "wfP (\<lambda>x y. (x, y) \<in> r) = wf r"
   by (simp add: wfP_def)
 
-lemma wfUNIVI: 
-   "(!!P x. (ALL x. (ALL y. (y,x) : r --> P(y)) --> P(x)) ==> P(x)) ==> wf(r)"
+lemma wfUNIVI: "(\<And>P x. (\<forall>x. (\<forall>y. (y, x) \<in> r \<longrightarrow> P y) \<longrightarrow> P x) \<Longrightarrow> P x) \<Longrightarrow> wf r"
   unfolding wf_def by blast
 
 lemmas wfPUNIVI = wfUNIVI [to_pred]
 
-text\<open>Restriction to domain @{term A} and range @{term B}.  If @{term r} is
-    well-founded over their intersection, then @{term "wf r"}\<close>
-lemma wfI: 
- "[| r \<subseteq> A \<times> B; 
-     !!x P. [|\<forall>x. (\<forall>y. (y,x) : r --> P y) --> P x;  x : A; x : B |] ==> P x |]
-  ==>  wf r"
-  unfolding wf_def by blast
+text \<open>Restriction to domain \<open>A\<close> and range \<open>B\<close>.
+  If \<open>r\<close> is well-founded over their intersection, then \<open>wf r\<close>.\<close>
+lemma wfI:
+  assumes "r \<subseteq> A \<times> B"
+    and "\<And>x P. \<lbrakk>\<forall>x. (\<forall>y. (y, x) \<in> r \<longrightarrow> P y) \<longrightarrow> P x;  x \<in> A; x \<in> B\<rbrakk> \<Longrightarrow> P x"
+  shows "wf r"
+  using assms unfolding wf_def by blast
 
-lemma wf_induct: 
-    "[| wf(r);           
-        !!x.[| ALL y. (y,x): r --> P(y) |] ==> P(x)  
-     |]  ==>  P(a)"
-  unfolding wf_def by blast
+lemma wf_induct:
+  assumes "wf r"
+    and "\<And>x. \<forall>y. (y, x) \<in> r \<longrightarrow> P y \<Longrightarrow> P x"
+  shows "P a"
+  using assms unfolding wf_def by blast
 
 lemmas wfP_induct = wf_induct [to_pred]
 
@@ -49,7 +48,7 @@
 
 lemmas wfP_induct_rule = wf_induct_rule [to_pred, induct set: wfP]
 
-lemma wf_not_sym: "wf r ==> (a, x) : r ==> (x, a) ~: r"
+lemma wf_not_sym: "wf r \<Longrightarrow> (a, x) \<in> r \<Longrightarrow> (x, a) \<notin> r"
   by (induct a arbitrary: x set: wf) blast
 
 lemma wf_asym:
@@ -57,22 +56,25 @@
   obtains "(x, a) \<notin> r"
   by (drule wf_not_sym[OF assms])
 
-lemma wf_not_refl [simp]: "wf r ==> (a, a) ~: r"
+lemma wf_not_refl [simp]: "wf r \<Longrightarrow> (a, a) \<notin> r"
   by (blast elim: wf_asym)
 
 lemma wf_irrefl: assumes "wf r" obtains "(a, a) \<notin> r"
-by (drule wf_not_refl[OF assms])
+  by (drule wf_not_refl[OF assms])
 
 lemma wf_wellorderI:
   assumes wf: "wf {(x::'a::ord, y). x < y}"
   assumes lin: "OFCLASS('a::ord, linorder_class)"
   shows "OFCLASS('a::ord, wellorder_class)"
-using lin by (rule wellorder_class.intro)
-  (rule class.wellorder_axioms.intro, rule wf_induct_rule [OF wf], simp)
+  using lin
+  apply (rule wellorder_class.intro)
+  apply (rule class.wellorder_axioms.intro)
+  apply (rule wf_induct_rule [OF wf])
+  apply simp
+  done
 
-lemma (in wellorder) wf:
-  "wf {(x, y). x < y}"
-unfolding wf_def by (blast intro: less_induct)
+lemma (in wellorder) wf: "wf {(x, y). x < y}"
+  unfolding wf_def by (blast intro: less_induct)
 
 
 subsection \<open>Basic Results\<close>
@@ -84,14 +86,13 @@
   assumes a: "A \<subseteq> R `` A"
   shows "A = {}"
 proof -
-  { fix x
-    from wf have "x \<notin> A"
-    proof induct
-      fix x assume "\<And>y. (y, x) \<in> R \<Longrightarrow> y \<notin> A"
-      then have "x \<notin> R `` A" by blast
-      with a show "x \<notin> A" by blast
-    qed
-  } thus ?thesis by auto
+  from wf have "x \<notin> A" for x
+  proof induct
+    fix x assume "\<And>y. (y, x) \<in> R \<Longrightarrow> y \<notin> A"
+    then have "x \<notin> R `` A" by blast
+    with a show "x \<notin> A" by blast
+  qed
+  then show ?thesis by auto
 qed
 
 lemma wfI_pf:
@@ -105,7 +106,8 @@
   with a show "P x" by blast
 qed
 
-text\<open>Minimal-element characterization of well-foundedness\<close>
+
+subsubsection \<open>Minimal-element characterization of well-foundedness\<close>
 
 lemma wfE_min:
   assumes wf: "wf R" and Q: "x \<in> Q"
@@ -120,14 +122,14 @@
   assumes a: "\<And>x Q. x \<in> Q \<Longrightarrow> \<exists>z\<in>Q. \<forall>y. (y, z) \<in> R \<longrightarrow> y \<notin> Q"
   shows "wf R"
 proof (rule wfI_pf)
-  fix A assume b: "A \<subseteq> R `` A"
-  { fix x assume "x \<in> A"
-    from a[OF this] b have "False" by blast
-  }
-  thus "A = {}" by blast
+  fix A
+  assume b: "A \<subseteq> R `` A"
+  have False if "x \<in> A" for x
+    using a[OF that] b by blast
+  then show "A = {}" by blast
 qed
 
-lemma wf_eq_minimal: "wf r = (\<forall>Q x. x\<in>Q --> (\<exists>z\<in>Q. \<forall>y. (y,z)\<in>r --> y\<notin>Q))"
+lemma wf_eq_minimal: "wf r \<longleftrightarrow> (\<forall>Q x. x \<in> Q \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. (y, z) \<in> r \<longrightarrow> y \<notin> Q))"
 apply auto
 apply (erule wfE_min, assumption, blast)
 apply (rule wfI_min, auto)
@@ -135,51 +137,52 @@
 
 lemmas wfP_eq_minimal = wf_eq_minimal [to_pred]
 
-text\<open>Well-foundedness of transitive closure\<close>
+
+subsubsection \<open>Well-foundedness of transitive closure\<close>
 
 lemma wf_trancl:
   assumes "wf r"
-  shows "wf (r^+)"
+  shows "wf (r\<^sup>+)"
 proof -
-  {
-    fix P and x
-    assume induct_step: "!!x. (!!y. (y, x) : r^+ ==> P y) ==> P x"
-    have "P x"
-    proof (rule induct_step)
-      fix y assume "(y, x) : r^+"
-      with \<open>wf r\<close> show "P y"
-      proof (induct x arbitrary: y)
-        case (less x)
-        note hyp = \<open>\<And>x' y'. (x', x) : r ==> (y', x') : r^+ ==> P y'\<close>
-        from \<open>(y, x) : r^+\<close> show "P y"
-        proof cases
-          case base
-          show "P y"
-          proof (rule induct_step)
-            fix y' assume "(y', y) : r^+"
-            with \<open>(y, x) : r\<close> 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])
+  have "P x" if induct_step: "\<And>x. (\<And>y. (y, x) \<in> r\<^sup>+ \<Longrightarrow> P y) \<Longrightarrow> P x" for P x
+  proof (rule induct_step)
+    show "P y" if "(y, x) \<in> r\<^sup>+" for y
+      using \<open>wf r\<close> and that
+    proof (induct x arbitrary: y)
+      case (less x)
+      note hyp = \<open>\<And>x' y'. (x', x) \<in> r \<Longrightarrow> (y', x') \<in> r\<^sup>+ \<Longrightarrow> P y'\<close>
+      from \<open>(y, x) \<in> r\<^sup>+\<close> show "P y"
+      proof cases
+        case base
+        show "P y"
+        proof (rule induct_step)
+          fix y'
+          assume "(y', y) \<in> r\<^sup>+"
+          with \<open>(y, x) \<in> r\<close> show "P y'"
+            by (rule hyp [of y y'])
         qed
+      next
+        case step
+        then obtain x' where "(x', x) \<in> r" and "(y, x') \<in> r\<^sup>+"
+          by simp
+        then show "P y" by (rule hyp [of x' y])
       qed
     qed
-  } then show ?thesis unfolding wf_def by blast
+  qed
+  then show ?thesis unfolding wf_def by blast
 qed
 
 lemmas wfP_trancl = wf_trancl [to_pred]
 
-lemma wf_converse_trancl: "wf (r^-1) ==> wf ((r^+)^-1)"
+lemma wf_converse_trancl: "wf (r\<inverse>) \<Longrightarrow> wf ((r\<^sup>+)\<inverse>)"
   apply (subst trancl_converse [symmetric])
   apply (erule wf_trancl)
   done
 
 text \<open>Well-foundedness of subsets\<close>
 
-lemma wf_subset: "[| wf(r);  p<=r |] ==> wf(p)"
-  apply (simp (no_asm_use) add: wf_eq_minimal)
+lemma wf_subset: "wf r \<Longrightarrow> p \<subseteq> r \<Longrightarrow> wf p"
+  apply (simp add: wf_eq_minimal)
   apply fast
   done
 
@@ -197,15 +200,15 @@
   then show ?thesis by (simp add: bot_fun_def)
 qed
 
-lemma wf_Int1: "wf r ==> wf (r Int r')"
+lemma wf_Int1: "wf r \<Longrightarrow> wf (r Int r')"
   apply (erule wf_subset)
   apply (rule Int_lower1)
   done
 
-lemma wf_Int2: "wf r ==> wf (r' Int r)"
+lemma wf_Int2: "wf r \<Longrightarrow> wf (r' Int r)"
   apply (erule wf_subset)
   apply (rule Int_lower2)
-  done  
+  done
 
 text \<open>Exponentiation\<close>
 
@@ -221,33 +224,34 @@
 
 text \<open>Well-foundedness of insert\<close>
 
-lemma wf_insert [iff]: "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)"
+lemma wf_insert [iff]: "wf (insert (y, x) r) \<longleftrightarrow> wf r \<and> (x, y) \<notin> r\<^sup>*"
 apply (rule iffI)
  apply (blast elim: wf_trancl [THEN wf_irrefl]
-              intro: rtrancl_into_trancl1 wf_subset 
+              intro: rtrancl_into_trancl1 wf_subset
                      rtrancl_mono [THEN [2] rev_subsetD])
 apply (simp add: wf_eq_minimal, safe)
-apply (rule allE, assumption, erule impE, blast) 
+apply (rule allE, assumption, erule impE, blast)
 apply (erule bexE)
 apply (rename_tac "a", case_tac "a = x")
  prefer 2
-apply blast 
-apply (case_tac "y:Q")
+apply blast
+apply (case_tac "y \<in> Q")
  prefer 2 apply blast
-apply (rule_tac x = "{z. z:Q & (z,y) : r^*}" in allE)
+apply (rule_tac x = "{z. z \<in> Q \<and> (z,y) \<in> r\<^sup>*}" in allE)
  apply assumption
-apply (erule_tac V = "ALL Q. (EX x. x : Q) --> P Q" for P in thin_rl) 
-  \<comment>\<open>essential for speed\<close>
-txt\<open>Blast with new substOccur fails\<close>
+apply (erule_tac V = "\<forall>Q. (\<exists>x. x \<in> Q) \<longrightarrow> P Q" for P in thin_rl)
+  (*essential for speed*)
+(*blast with new substOccur fails*)
 apply (fast intro: converse_rtrancl_into_rtrancl)
 done
 
-text\<open>Well-foundedness of image\<close>
+
+subsubsection \<open>Well-foundedness of image\<close>
 
-lemma wf_map_prod_image: "[| wf r; inj f |] ==> wf (map_prod f f ` r)"
+lemma wf_map_prod_image: "wf r \<Longrightarrow> inj f \<Longrightarrow> wf (map_prod f f ` r)"
 apply (simp only: wf_eq_minimal, clarify)
-apply (case_tac "EX p. f p : Q")
-apply (erule_tac x = "{p. f p : Q}" in allE)
+apply (case_tac "\<exists>p. f p \<in> Q")
+apply (erule_tac x = "{p. f p \<in> Q}" in allE)
 apply (fast dest: inj_onD, blast)
 done
 
@@ -259,25 +263,23 @@
   assumes "R O S \<subseteq> R"
   shows "wf (R \<union> S)"
 proof (rule wfI_min)
-  fix x :: 'a and Q 
+  fix x :: 'a and Q
   let ?Q' = "{x \<in> Q. \<forall>y. (y, x) \<in> R \<longrightarrow> y \<notin> Q}"
   assume "x \<in> Q"
   obtain a where "a \<in> ?Q'"
     by (rule wfE_min [OF \<open>wf R\<close> \<open>x \<in> Q\<close>]) blast
-  with \<open>wf S\<close>
-  obtain z where "z \<in> ?Q'" and zmin: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> ?Q'" by (erule wfE_min)
-  { 
+  with \<open>wf S\<close> obtain z where "z \<in> ?Q'" and zmin: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> ?Q'"
+    by (erule wfE_min)
+  {
     fix y assume "(y, z) \<in> S"
     then have "y \<notin> ?Q'" by (rule zmin)
-
     have "y \<notin> Q"
-    proof 
+    proof
       assume "y \<in> Q"
-      with \<open>y \<notin> ?Q'\<close> 
-      obtain w where "(w, y) \<in> R" and "w \<in> Q" by auto
+      with \<open>y \<notin> ?Q'\<close> obtain w where "(w, y) \<in> R" and "w \<in> Q" by auto
       from \<open>(w, y) \<in> R\<close> \<open>(y, z) \<in> S\<close> have "(w, z) \<in> R O S" by (rule relcompI)
       with \<open>R O S \<subseteq> R\<close> have "(w, z) \<in> R" ..
-      with \<open>z \<in> ?Q'\<close> have "w \<notin> Q" by blast 
+      with \<open>z \<in> ?Q'\<close> have "w \<notin> Q" by blast
       with \<open>w \<in> Q\<close> show False by contradiction
     qed
   }
@@ -287,18 +289,21 @@
 
 text \<open>Well-foundedness of indexed union with disjoint domains and ranges\<close>
 
-lemma wf_UN: "[| ALL i:I. wf(r i);  
-         ALL i:I. ALL j:I. r i ~= r j --> Domain(r i) Int Range(r j) = {}  
-      |] ==> wf(UN i:I. r i)"
-apply (simp only: wf_eq_minimal, clarify)
-apply (rename_tac A a, case_tac "EX i:I. EX a:A. EX b:A. (b,a) : r i")
- prefer 2
- apply force 
-apply clarify
-apply (drule bspec, assumption)  
-apply (erule_tac x="{a. a:A & (EX b:A. (b,a) : r i) }" in allE)
-apply (blast elim!: allE)  
-done
+lemma wf_UN:
+  assumes "\<forall>i\<in>I. wf (r i)"
+    and "\<forall>i\<in>I. \<forall>j\<in>I. r i \<noteq> r j \<longrightarrow> Domain (r i) \<inter> Range (r j) = {}"
+  shows "wf (\<Union>i\<in>I. r i)"
+  using assms
+  apply (simp only: wf_eq_minimal)
+  apply clarify
+  apply (rename_tac A a, case_tac "\<exists>i\<in>I. \<exists>a\<in>A. \<exists>b\<in>A. (b, a) \<in> r i")
+   prefer 2
+   apply force
+  apply clarify
+  apply (drule bspec, assumption)
+  apply (erule_tac x="{a. a \<in> A \<and> (\<exists>b\<in>A. (b, a) \<in> r i) }" in allE)
+  apply (blast elim!: allE)
+  done
 
 lemma wfP_SUP:
   "\<forall>i. wfP (r i) \<Longrightarrow> \<forall>i j. r i \<noteq> r j \<longrightarrow> inf (DomainP (r i)) (RangeP (r j)) = bot \<Longrightarrow> wfP (SUPREMUM UNIV r)"
@@ -306,72 +311,66 @@
   apply simp_all
   done
 
-lemma wf_Union: 
- "[| ALL r:R. wf r;  
-     ALL r:R. ALL s:R. r ~= s --> Domain r Int Range s = {}  
-  |] ==> wf (\<Union> R)"
-  using wf_UN[of R "\<lambda>i. i"] by simp
+lemma wf_Union:
+  assumes "\<forall>r\<in>R. wf r"
+    and "\<forall>r\<in>R. \<forall>s\<in>R. r \<noteq> s \<longrightarrow> Domain r \<inter> Range s = {}"
+  shows "wf (\<Union>R)"
+  using assms wf_UN[of R "\<lambda>i. i"] by simp
 
-(*Intuition: we find an (R u S)-min element of a nonempty subset A
-             by case distinction.
-  1. There is a step a -R-> b with a,b : A.
-     Pick an R-min element z of the (nonempty) set {a:A | EX b:A. a -R-> b}.
-     By definition, there is z':A s.t. z -R-> z'. Because z is R-min in the
-     subset, z' must be R-min in A. Because z' has an R-predecessor, it cannot
-     have an S-successor and is thus S-min in A as well.
-  2. There is no such step.
-     Pick an S-min element of A. In this case it must be an R-min
-     element of A as well.
-*)
-lemma wf_Un:
-     "[| wf r; wf s; Domain r Int Range s = {} |] ==> wf(r Un s)"
-  using wf_union_compatible[of s r] 
+text \<open>
+  Intuition: We find an \<open>R \<union> S\<close>-min element of a nonempty subset \<open>A\<close> by case distinction.
+  \<^enum> There is a step \<open>a \<midarrow>R\<rightarrow> b\<close> with \<open>a, b \<in> A\<close>.
+    Pick an \<open>R\<close>-min element \<open>z\<close> of the (nonempty) set \<open>{a\<in>A | \<exists>b\<in>A. a \<midarrow>R\<rightarrow> b}\<close>.
+    By definition, there is \<open>z' \<in> A\<close> s.t. \<open>z \<midarrow>R\<rightarrow> z'\<close>. Because \<open>z\<close> is \<open>R\<close>-min in the
+    subset, \<open>z'\<close> must be \<open>R\<close>-min in \<open>A\<close>. Because \<open>z'\<close> has an \<open>R\<close>-predecessor, it cannot
+    have an \<open>S\<close>-successor and is thus \<open>S\<close>-min in \<open>A\<close> as well.
+  \<^enum> There is no such step.
+    Pick an \<open>S\<close>-min element of \<open>A\<close>. In this case it must be an \<open>R\<close>-min
+    element of \<open>A\<close> as well.
+\<close>
+lemma wf_Un: "wf r \<Longrightarrow> wf s \<Longrightarrow> Domain r \<inter> Range s = {} \<Longrightarrow> wf (r \<union> s)"
+  using wf_union_compatible[of s r]
   by (auto simp: Un_ac)
 
-lemma wf_union_merge: 
-  "wf (R \<union> S) = wf (R O R \<union> S O R \<union> S)" (is "wf ?A = wf ?B")
+lemma wf_union_merge: "wf (R \<union> S) = wf (R O R \<union> S O R \<union> S)"
+  (is "wf ?A = wf ?B")
 proof
   assume "wf ?A"
-  with wf_trancl have wfT: "wf (?A^+)" .
-  moreover have "?B \<subseteq> ?A^+"
+  with wf_trancl have wfT: "wf (?A\<^sup>+)" .
+  moreover have "?B \<subseteq> ?A\<^sup>+"
     by (subst trancl_unfold, subst trancl_unfold) blast
   ultimately show "wf ?B" by (rule wf_subset)
 next
   assume "wf ?B"
-
   show "wf ?A"
   proof (rule wfI_min)
-    fix Q :: "'a set" and x 
+    fix Q :: "'a set" and x
     assume "x \<in> Q"
-
-    with \<open>wf ?B\<close>
-    obtain z where "z \<in> Q" and "\<And>y. (y, z) \<in> ?B \<Longrightarrow> y \<notin> Q" 
+    with \<open>wf ?B\<close> obtain z where "z \<in> Q" and "\<And>y. (y, z) \<in> ?B \<Longrightarrow> y \<notin> Q"
       by (erule wfE_min)
-    then have A1: "\<And>y. (y, z) \<in> R O R \<Longrightarrow> y \<notin> Q"
-      and A2: "\<And>y. (y, z) \<in> S O R \<Longrightarrow> y \<notin> Q"
-      and A3: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> Q"
+    then have 1: "\<And>y. (y, z) \<in> R O R \<Longrightarrow> y \<notin> Q"
+      and 2: "\<And>y. (y, z) \<in> S O R \<Longrightarrow> y \<notin> Q"
+      and 3: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> Q"
       by auto
-    
     show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> ?A \<longrightarrow> y \<notin> Q"
     proof (cases "\<forall>y. (y, z) \<in> R \<longrightarrow> y \<notin> Q")
       case True
-      with \<open>z \<in> Q\<close> A3 show ?thesis by blast
+      with \<open>z \<in> Q\<close> 3 show ?thesis by blast
     next
-      case False 
+      case False
       then obtain z' where "z'\<in>Q" "(z', z) \<in> R" by blast
-
       have "\<forall>y. (y, z') \<in> ?A \<longrightarrow> y \<notin> Q"
       proof (intro allI impI)
         fix y assume "(y, z') \<in> ?A"
         then show "y \<notin> Q"
         proof
-          assume "(y, z') \<in> R" 
+          assume "(y, z') \<in> R"
           then have "(y, z) \<in> R O R" using \<open>(z', z) \<in> R\<close> ..
-          with A1 show "y \<notin> Q" .
+          with 1 show "y \<notin> Q" .
         next
-          assume "(y, z') \<in> S" 
+          assume "(y, z') \<in> S"
           then have "(y, z) \<in> S O R" using  \<open>(z', z) \<in> R\<close> ..
-          with A2 show "y \<notin> Q" .
+          with 2 show "y \<notin> Q" .
         qed
       qed
       with \<open>z' \<in> Q\<close> show ?thesis ..
@@ -389,40 +388,55 @@
 
 lemma qc_wf_relto_iff:
   assumes "R O S \<subseteq> (R \<union> S)\<^sup>* O R" \<comment> \<open>R quasi-commutes over S\<close>
-  shows "wf (S\<^sup>* O R O S\<^sup>*) \<longleftrightarrow> wf R" (is "wf ?S \<longleftrightarrow> _")
+  shows "wf (S\<^sup>* O R O S\<^sup>*) \<longleftrightarrow> wf R"
+  (is "wf ?S \<longleftrightarrow> _")
 proof
-  assume "wf ?S"
-  moreover have "R \<subseteq> ?S" by auto
-  ultimately show "wf R" using wf_subset by auto
+  show "wf R" if "wf ?S"
+  proof -
+    have "R \<subseteq> ?S" by auto
+    with that show "wf R" using wf_subset by auto
+  qed
 next
-  assume "wf R"
-  show "wf ?S"
+  show "wf ?S" if "wf R"
   proof (rule wfI_pf)
-    fix A assume A: "A \<subseteq> ?S `` A"
+    fix A
+    assume A: "A \<subseteq> ?S `` A"
     let ?X = "(R \<union> S)\<^sup>* `` A"
     have *: "R O (R \<union> S)\<^sup>* \<subseteq> (R \<union> S)\<^sup>* O R"
-      proof -
-        { fix x y z assume "(y, z) \<in> (R \<union> S)\<^sup>*" and "(x, y) \<in> R"
-          then have "(x, z) \<in> (R \<union> S)\<^sup>* O R"
-          proof (induct y z)
-            case rtrancl_refl then show ?case by auto
-          next
-            case (rtrancl_into_rtrancl a b c)
-            then have "(x, c) \<in> ((R \<union> S)\<^sup>* O (R \<union> S)\<^sup>*) O R" using assms by blast
-            then show ?case by simp
-          qed }
-        then show ?thesis by auto
+    proof -
+      have "(x, z) \<in> (R \<union> S)\<^sup>* O R" if "(y, z) \<in> (R \<union> S)\<^sup>*" and "(x, y) \<in> R" for x y z
+        using that
+      proof (induct y z)
+        case rtrancl_refl
+        then show ?case by auto
+      next
+        case (rtrancl_into_rtrancl a b c)
+        then have "(x, c) \<in> ((R \<union> S)\<^sup>* O (R \<union> S)\<^sup>*) O R"
+          using assms by blast
+        then show ?case by simp
       qed
-    then have "R O S\<^sup>* \<subseteq> (R \<union> S)\<^sup>* O R" using rtrancl_Un_subset by blast
-    then have "?S \<subseteq> (R \<union> S)\<^sup>* O (R \<union> S)\<^sup>* O R" by (simp add: relcomp_mono rtrancl_mono)
-    also have "\<dots> = (R \<union> S)\<^sup>* O R" by (simp add: O_assoc[symmetric])
-    finally have "?S O (R \<union> S)\<^sup>* \<subseteq> (R \<union> S)\<^sup>* O R O (R \<union> S)\<^sup>*" by (simp add: O_assoc[symmetric] relcomp_mono)
-    also have "\<dots> \<subseteq> (R \<union> S)\<^sup>* O (R \<union> S)\<^sup>* O R" using * by (simp add: relcomp_mono)
-    finally have "?S O (R \<union> S)\<^sup>* \<subseteq> (R \<union> S)\<^sup>* O R" by (simp add: O_assoc[symmetric])
-    then have "(?S O (R \<union> S)\<^sup>*) `` A \<subseteq> ((R \<union> S)\<^sup>* O R) `` A" by (simp add: Image_mono)
-    moreover have "?X \<subseteq> (?S O (R \<union> S)\<^sup>*) `` A" using A by (auto simp: relcomp_Image)
-    ultimately have "?X \<subseteq> R `` ?X" by (auto simp: relcomp_Image)
-    then have "?X = {}" using \<open>wf R\<close> by (simp add: wfE_pf)
+      then show ?thesis by auto
+    qed
+    then have "R O S\<^sup>* \<subseteq> (R \<union> S)\<^sup>* O R"
+      using rtrancl_Un_subset by blast
+    then have "?S \<subseteq> (R \<union> S)\<^sup>* O (R \<union> S)\<^sup>* O R"
+      by (simp add: relcomp_mono rtrancl_mono)
+    also have "\<dots> = (R \<union> S)\<^sup>* O R"
+      by (simp add: O_assoc[symmetric])
+    finally have "?S O (R \<union> S)\<^sup>* \<subseteq> (R \<union> S)\<^sup>* O R O (R \<union> S)\<^sup>*"
+      by (simp add: O_assoc[symmetric] relcomp_mono)
+    also have "\<dots> \<subseteq> (R \<union> S)\<^sup>* O (R \<union> S)\<^sup>* O R"
+      using * by (simp add: relcomp_mono)
+    finally have "?S O (R \<union> S)\<^sup>* \<subseteq> (R \<union> S)\<^sup>* O R"
+      by (simp add: O_assoc[symmetric])
+    then have "(?S O (R \<union> S)\<^sup>*) `` A \<subseteq> ((R \<union> S)\<^sup>* O R) `` A"
+      by (simp add: Image_mono)
+    moreover have "?X \<subseteq> (?S O (R \<union> S)\<^sup>*) `` A"
+      using A by (auto simp: relcomp_Image)
+    ultimately have "?X \<subseteq> R `` ?X"
+      by (auto simp: relcomp_Image)
+    then have "?X = {}"
+      using \<open>wf R\<close> by (simp add: wfE_pf)
     moreover have "A \<subseteq> ?X" by auto
     ultimately show "A = {}" by simp
   qed
@@ -443,22 +457,23 @@
 
 subsection \<open>Acyclic relations\<close>
 
-lemma wf_acyclic: "wf r ==> acyclic r"
+lemma wf_acyclic: "wf r \<Longrightarrow> acyclic r"
 apply (simp add: acyclic_def)
 apply (blast elim: wf_trancl [THEN wf_irrefl])
 done
 
 lemmas wfP_acyclicP = wf_acyclic [to_pred]
 
-text\<open>Wellfoundedness of finite acyclic relations\<close>
+
+subsubsection \<open>Wellfoundedness of finite acyclic relations\<close>
 
-lemma finite_acyclic_wf [rule_format]: "finite r ==> acyclic r --> wf r"
+lemma finite_acyclic_wf [rule_format]: "finite r \<Longrightarrow> acyclic r \<longrightarrow> wf r"
 apply (erule finite_induct, blast)
-apply (simp (no_asm_simp) only: split_tupled_all)
+apply (simp only: split_tupled_all)
 apply simp
 done
 
-lemma finite_acyclic_wf_converse: "[|finite r; acyclic r|] ==> wf (r^-1)"
+lemma finite_acyclic_wf_converse: "finite r \<Longrightarrow> acyclic r \<Longrightarrow> wf (r\<inverse>)"
 apply (erule finite_converse [THEN iffD2, THEN finite_acyclic_wf])
 apply (erule acyclic_converse [THEN iffD2])
 done
@@ -477,44 +492,39 @@
   with \<open>finite r\<close> show ?thesis by (rule finite_acyclic_wf_converse)
 qed
 
-lemma wf_iff_acyclic_if_finite: "finite r ==> wf r = acyclic r"
+lemma wf_iff_acyclic_if_finite: "finite r \<Longrightarrow> wf r = acyclic r"
 by (blast intro: finite_acyclic_wf wf_acyclic)
 
 
 subsection \<open>@{typ nat} is well-founded\<close>
 
-lemma less_nat_rel: "op < = (\<lambda>m n. n = Suc m)^++"
+lemma less_nat_rel: "op < = (\<lambda>m n. n = Suc m)\<^sup>+\<^sup>+"
 proof (rule ext, rule ext, rule iffI)
   fix n m :: nat
-  assume "m < n"
-  then show "(\<lambda>m n. n = Suc m)^++ m n"
+  show "(\<lambda>m n. n = Suc m)\<^sup>+\<^sup>+ m n" if "m < n"
+    using that
   proof (induct n)
-    case 0 then show ?case by auto
+    case 0
+    then show ?case by auto
   next
-    case (Suc n) then show ?case
+    case (Suc n)
+    then show ?case
       by (auto simp add: less_Suc_eq_le le_less intro: tranclp.trancl_into_trancl)
   qed
-next
-  fix n m :: nat
-  assume "(\<lambda>m n. n = Suc m)^++ m n"
-  then show "m < n"
-    by (induct n)
-      (simp_all add: less_Suc_eq_le reflexive le_less)
+  show "m < n" if "(\<lambda>m n. n = Suc m)\<^sup>+\<^sup>+ m n"
+    using that by (induct n) (simp_all add: less_Suc_eq_le reflexive le_less)
 qed
 
-definition
-  pred_nat :: "(nat * nat) set" where
-  "pred_nat = {(m, n). n = Suc m}"
+definition pred_nat :: "(nat \<times> nat) set"
+  where "pred_nat = {(m, n). n = Suc m}"
 
-definition
-  less_than :: "(nat * nat) set" where
-  "less_than = pred_nat^+"
+definition less_than :: "(nat \<times> nat) set"
+  where "less_than = pred_nat\<^sup>+"
 
-lemma less_eq: "(m, n) \<in> pred_nat^+ \<longleftrightarrow> m < n"
+lemma less_eq: "(m, n) \<in> pred_nat\<^sup>+ \<longleftrightarrow> m < n"
   unfolding less_nat_rel pred_nat_def trancl_def by simp
 
-lemma pred_nat_trancl_eq_le:
-  "(m, n) \<in> pred_nat^* \<longleftrightarrow> m \<le> n"
+lemma pred_nat_trancl_eq_le: "(m, n) \<in> pred_nat\<^sup>* \<longleftrightarrow> m \<le> n"
   unfolding less_eq rtrancl_eq_or_trancl by auto
 
 lemma wf_pred_nat: "wf pred_nat"
@@ -528,7 +538,7 @@
 lemma trans_less_than [iff]: "trans less_than"
   by (simp add: less_than_def)
 
-lemma less_than_iff [iff]: "((x,y): less_than) = (x<y)"
+lemma less_than_iff [iff]: "((x,y) \<in> less_than) = (x<y)"
   by (simp add: less_than_def less_eq)
 
 lemma wf_less: "wf {(x, y::nat). x < y}"
@@ -538,28 +548,22 @@
 subsection \<open>Accessible Part\<close>
 
 text \<open>
- Inductive definition of the accessible part @{term "acc r"} of a
- relation; see also @{cite "paulin-tlca"}.
+  Inductive definition of the accessible part \<open>acc r\<close> of a
+  relation; see also @{cite "paulin-tlca"}.
 \<close>
 
-inductive_set
-  acc :: "('a * 'a) set => 'a set"
-  for r :: "('a * 'a) set"
-  where
-    accI: "(!!y. (y, x) : r ==> y : acc r) ==> x : acc r"
+inductive_set acc :: "('a \<times> 'a) set \<Rightarrow> 'a set" for r :: "('a \<times> 'a) set"
+  where accI: "(\<And>y. (y, x) \<in> r \<Longrightarrow> y \<in> acc r) \<Longrightarrow> x \<in> acc r"
 
-abbreviation
-  termip :: "('a => 'a => bool) => 'a => bool" where
-  "termip r \<equiv> accp (r\<inverse>\<inverse>)"
+abbreviation termip :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
+  where "termip r \<equiv> accp (r\<inverse>\<inverse>)"
 
-abbreviation
-  termi :: "('a * 'a) set => 'a set" where
-  "termi r \<equiv> acc (r\<inverse>)"
+abbreviation termi :: "('a \<times> 'a) set \<Rightarrow> 'a set"
+  where "termi r \<equiv> acc (r\<inverse>)"
 
 lemmas accpI = accp.accI
 
-lemma accp_eq_acc [code]:
-  "accp r = (\<lambda>x. x \<in> Wellfounded.acc {(x, y). r x y})"
+lemma accp_eq_acc [code]: "accp r = (\<lambda>x. x \<in> Wellfounded.acc {(x, y). r x y})"
   by (simp add: acc_def)
 
 
@@ -567,7 +571,7 @@
 
 theorem accp_induct:
   assumes major: "accp r a"
-  assumes hyp: "!!x. accp r x ==> \<forall>y. r y x --> P y ==> P x"
+  assumes hyp: "\<And>x. accp r x \<Longrightarrow> \<forall>y. r y x \<longrightarrow> P y \<Longrightarrow> P x"
   shows "P a"
   apply (rule major [THEN accp.induct])
   apply (rule hyp)
@@ -578,7 +582,7 @@
 
 lemmas accp_induct_rule = accp_induct [rule_format, induct set: accp]
 
-theorem accp_downward: "accp r b ==> r a b ==> accp r a"
+theorem accp_downward: "accp r b \<Longrightarrow> r a b \<Longrightarrow> accp r a"
   apply (erule accp.cases)
   apply fast
   done
@@ -588,13 +592,11 @@
   obtains z where "R z x" and "\<not> accp R z"
 proof -
   assume a: "\<And>z. \<lbrakk>R z x; \<not> accp R z\<rbrakk> \<Longrightarrow> thesis"
-
   show thesis
   proof (cases "\<forall>z. R z x \<longrightarrow> accp R z")
     case True
-    hence "\<And>z. R z x \<Longrightarrow> accp R z" by auto
-    hence "accp R x"
-      by (rule accp.accI)
+    then have "\<And>z. R z x \<Longrightarrow> accp R z" by auto
+    then have "accp R x" by (rule accp.accI)
     with na show thesis ..
   next
     case False then obtain z where "R z x" and "\<not> accp R z"
@@ -603,24 +605,24 @@
   qed
 qed
 
-lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a ==> accp r a --> accp r b"
+lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a \<Longrightarrow> accp r a \<longrightarrow> accp r b"
   apply (erule rtranclp_induct)
    apply blast
   apply (blast dest: accp_downward)
   done
 
-theorem accp_downwards: "accp r a ==> r\<^sup>*\<^sup>* b a ==> accp r b"
+theorem accp_downwards: "accp r a \<Longrightarrow> r\<^sup>*\<^sup>* b a \<Longrightarrow> accp r b"
   apply (blast dest: accp_downwards_aux)
   done
 
-theorem accp_wfPI: "\<forall>x. accp r x ==> wfP r"
+theorem accp_wfPI: "\<forall>x. accp r x \<Longrightarrow> wfP r"
   apply (rule wfPUNIVI)
   apply (rule_tac P=P in accp_induct)
    apply blast
   apply blast
   done
 
-theorem accp_wfPD: "wfP r ==> accp r x"
+theorem accp_wfPD: "wfP r \<Longrightarrow> accp r x"
   apply (erule wfP_induct_rule)
   apply (rule accp.accI)
   apply blast
@@ -673,25 +675,15 @@
 text \<open>Set versions of the above theorems\<close>
 
 lemmas acc_induct = accp_induct [to_set]
-
 lemmas acc_induct_rule = acc_induct [rule_format, induct set: acc]
-
 lemmas acc_downward = accp_downward [to_set]
-
 lemmas not_acc_down = not_accp_down [to_set]
-
 lemmas acc_downwards_aux = accp_downwards_aux [to_set]
-
 lemmas acc_downwards = accp_downwards [to_set]
-
 lemmas acc_wfI = accp_wfPI [to_set]
-
 lemmas acc_wfD = accp_wfPD [to_set]
-
 lemmas wf_acc_iff = wfP_accp_iff [to_set]
-
 lemmas acc_subset = accp_subset [to_set]
-
 lemmas acc_subset_induct = accp_subset_induct [to_set]
 
 
@@ -699,10 +691,10 @@
 
 text \<open>Inverse Image\<close>
 
-lemma wf_inv_image [simp,intro!]: "wf(r) ==> wf(inv_image r (f::'a=>'b))"
-apply (simp (no_asm_use) add: inv_image_def wf_eq_minimal)
+lemma wf_inv_image [simp,intro!]: "wf r \<Longrightarrow> wf (inv_image r f)" for f :: "'a \<Rightarrow> 'b"
+apply (simp add: inv_image_def wf_eq_minimal)
 apply clarify
-apply (subgoal_tac "EX (w::'b) . w : {w. EX (x::'a) . x: Q & (f x = w) }")
+apply (subgoal_tac "\<exists>w::'b. w \<in> {w. \<exists>x::'a. x \<in> Q \<and> f x = w}")
 prefer 2 apply (blast del: allE)
 apply (erule allE)
 apply (erule (1) notE impE)
@@ -711,10 +703,10 @@
 
 text \<open>Measure functions into @{typ nat}\<close>
 
-definition measure :: "('a => nat) => ('a * 'a)set"
-where "measure = inv_image less_than"
+definition measure :: "('a \<Rightarrow> nat) \<Rightarrow> ('a \<times> 'a) set"
+  where "measure = inv_image less_than"
 
-lemma in_measure[simp, code_unfold]: "((x,y) : measure f) = (f x < f y)"
+lemma in_measure[simp, code_unfold]: "(x, y) \<in> measure f \<longleftrightarrow> f x < f y"
   by (simp add:measure_def)
 
 lemma wf_measure [iff]: "wf (measure f)"
@@ -722,8 +714,8 @@
 apply (rule wf_less_than [THEN wf_inv_image])
 done
 
-lemma wf_if_measure: fixes f :: "'a \<Rightarrow> nat"
-shows "(!!x. P x \<Longrightarrow> f(g x) < f x) \<Longrightarrow> wf {(y,x). P x \<and> y = g x}"
+lemma wf_if_measure: "(\<And>x. P x \<Longrightarrow> f(g x) < f x) \<Longrightarrow> wf {(y,x). P x \<and> y = g x}"
+  for f :: "'a \<Rightarrow> nat"
 apply(insert wf_measure[of f])
 apply(simp only: measure_def inv_image_def less_than_def less_eq)
 apply(erule wf_subset)
@@ -731,69 +723,66 @@
 done
 
 
-text\<open>Lexicographic combinations\<close>
+subsubsection \<open>Lexicographic combinations\<close>
 
-definition lex_prod :: "('a \<times>'a) set \<Rightarrow> ('b \<times> 'b) set \<Rightarrow> (('a \<times> 'b) \<times> ('a \<times> 'b)) set" (infixr "<*lex*>" 80) where
-  "ra <*lex*> rb = {((a, b), (a', b')). (a, a') \<in> ra \<or> a = a' \<and> (b, b') \<in> rb}"
+definition lex_prod :: "('a \<times>'a) set \<Rightarrow> ('b \<times> 'b) set \<Rightarrow> (('a \<times> 'b) \<times> ('a \<times> 'b)) set"
+    (infixr "<*lex*>" 80)
+  where "ra <*lex*> rb = {((a, b), (a', b')). (a, a') \<in> ra \<or> a = a' \<and> (b, b') \<in> rb}"
 
-lemma wf_lex_prod [intro!]: "[| wf(ra); wf(rb) |] ==> wf(ra <*lex*> rb)"
-apply (unfold wf_def lex_prod_def) 
+lemma wf_lex_prod [intro!]: "wf ra \<Longrightarrow> wf rb \<Longrightarrow> wf (ra <*lex*> rb)"
+apply (unfold wf_def lex_prod_def)
 apply (rule allI, rule impI)
-apply (simp (no_asm_use) only: split_paired_All)
-apply (drule spec, erule mp) 
+apply (simp only: split_paired_All)
+apply (drule spec, erule mp)
 apply (rule allI, rule impI)
-apply (drule spec, erule mp, blast) 
+apply (drule spec, erule mp, blast)
 done
 
-lemma in_lex_prod[simp]: 
-  "(((a,b),(a',b')): r <*lex*> s) = ((a,a'): r \<or> (a = a' \<and> (b, b') : s))"
+lemma in_lex_prod[simp]: "((a, b), (a', b')) \<in> r <*lex*> s \<longleftrightarrow> (a, a') \<in> r \<or> a = a' \<and> (b, b') \<in> s"
   by (auto simp:lex_prod_def)
 
-text\<open>@{term "op <*lex*>"} preserves transitivity\<close>
+text \<open>\<open><*lex*>\<close> preserves transitivity\<close>
+lemma trans_lex_prod [intro!]: "trans R1 \<Longrightarrow> trans R2 \<Longrightarrow> trans (R1 <*lex*> R2)"
+  unfolding trans_def lex_prod_def by blast
 
-lemma trans_lex_prod [intro!]: 
-    "[| trans R1; trans R2 |] ==> trans (R1 <*lex*> R2)"
-by (unfold trans_def lex_prod_def, blast) 
 
 text \<open>lexicographic combinations with measure functions\<close>
 
-definition 
-  mlex_prod :: "('a \<Rightarrow> nat) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set" (infixr "<*mlex*>" 80)
-where
-  "f <*mlex*> R = inv_image (less_than <*lex*> R) (%x. (f x, x))"
+definition mlex_prod :: "('a \<Rightarrow> nat) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set" (infixr "<*mlex*>" 80)
+  where "f <*mlex*> R = inv_image (less_than <*lex*> R) (\<lambda>x. (f x, x))"
 
 lemma wf_mlex: "wf R \<Longrightarrow> wf (f <*mlex*> R)"
-unfolding mlex_prod_def
-by auto
+  unfolding mlex_prod_def
+  by auto
 
 lemma mlex_less: "f x < f y \<Longrightarrow> (x, y) \<in> f <*mlex*> R"
-unfolding mlex_prod_def by simp
+  unfolding mlex_prod_def by simp
 
 lemma mlex_leq: "f x \<le> f y \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> (x, y) \<in> f <*mlex*> R"
-unfolding mlex_prod_def by auto
+  unfolding mlex_prod_def by auto
 
 text \<open>proper subset relation on finite sets\<close>
 
-definition finite_psubset  :: "('a set * 'a set) set"
-where "finite_psubset = {(A,B). A < B & finite B}"
+definition finite_psubset :: "('a set \<times> 'a set) set"
+  where "finite_psubset = {(A,B). A < B \<and> finite B}"
 
-lemma wf_finite_psubset[simp]: "wf(finite_psubset)"
-apply (unfold finite_psubset_def)
-apply (rule wf_measure [THEN wf_subset])
-apply (simp add: measure_def inv_image_def less_than_def less_eq)
-apply (fast elim!: psubset_card_mono)
-done
+lemma wf_finite_psubset[simp]: "wf finite_psubset"
+  apply (unfold finite_psubset_def)
+  apply (rule wf_measure [THEN wf_subset])
+  apply (simp add: measure_def inv_image_def less_than_def less_eq)
+  apply (fast elim!: psubset_card_mono)
+  done
 
 lemma trans_finite_psubset: "trans finite_psubset"
-by (simp add: finite_psubset_def less_le trans_def, blast)
+  by (auto simp add: finite_psubset_def less_le trans_def)
 
-lemma in_finite_psubset[simp]: "(A, B) \<in> finite_psubset = (A < B & finite B)"
-unfolding finite_psubset_def by auto
+lemma in_finite_psubset[simp]: "(A, B) \<in> finite_psubset \<longleftrightarrow> A < B \<and> finite B"
+  unfolding finite_psubset_def by auto
 
 text \<open>max- and min-extension of order to finite sets\<close>
 
-inductive_set max_ext :: "('a \<times> 'a) set \<Rightarrow> ('a set \<times> 'a set) set" 
-for R :: "('a \<times> 'a) set"
+inductive_set max_ext :: "('a \<times> 'a) set \<Rightarrow> ('a set \<times> 'a set) set"
+  for R :: "('a \<times> 'a) set"
 where
   max_extI[intro]: "finite X \<Longrightarrow> finite Y \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> \<exists>y\<in>Y. (x, y) \<in> R) \<Longrightarrow> (X, Y) \<in> max_ext R"
 
@@ -801,10 +790,11 @@
   assumes wf: "wf r"
   shows "wf (max_ext r)"
 proof (rule acc_wfI, intro allI)
-  fix M show "M \<in> acc (max_ext r)" (is "_ \<in> ?W")
+  fix M
+  show "M \<in> acc (max_ext r)" (is "_ \<in> ?W")
   proof cases
     assume "finite M"
-    thus ?thesis
+    then show ?thesis
     proof (induct M)
       show "{} \<in> ?W"
         by (rule accI) (auto elim: max_ext.cases)
@@ -815,42 +805,37 @@
         fix M a
         assume "M \<in> ?W"  and  [intro]: "finite M"
         assume hyp: "\<And>b M. (b, a) \<in> r \<Longrightarrow> M \<in> ?W \<Longrightarrow> finite M \<Longrightarrow> insert b M \<in> ?W"
-        {
-          fix N M :: "'a set"
-          assume "finite N" "finite M"
-          then
-          have "\<lbrakk>M \<in> ?W ; (\<And>y. y \<in> N \<Longrightarrow> (y, a) \<in> r)\<rbrakk> \<Longrightarrow>  N \<union> M \<in> ?W"
-            by (induct N arbitrary: M) (auto simp: hyp)
-        }
-        note add_less = this
-        
+        have add_less: "\<lbrakk>M \<in> ?W ; (\<And>y. y \<in> N \<Longrightarrow> (y, a) \<in> r)\<rbrakk> \<Longrightarrow> N \<union> M \<in> ?W"
+          if "finite N" "finite M" for N M :: "'a set"
+          using that by (induct N arbitrary: M) (auto simp: hyp)
+
         show "insert a M \<in> ?W"
         proof (rule accI)
-          fix N assume Nless: "(N, insert a M) \<in> max_ext r"
-          hence asm1: "\<And>x. x \<in> N \<Longrightarrow> (x, a) \<in> r \<or> (\<exists>y \<in> M. (x, y) \<in> r)"
+          fix N
+          assume Nless: "(N, insert a M) \<in> max_ext r"
+          then have *: "\<And>x. x \<in> N \<Longrightarrow> (x, a) \<in> r \<or> (\<exists>y \<in> M. (x, y) \<in> r)"
             by (auto elim!: max_ext.cases)
 
-          let ?N1 = "{ n \<in> N. (n, a) \<in> r }"
-          let ?N2 = "{ n \<in> N. (n, a) \<notin> r }"
+          let ?N1 = "{n \<in> N. (n, a) \<in> r}"
+          let ?N2 = "{n \<in> N. (n, a) \<notin> r}"
           have N: "?N1 \<union> ?N2 = N" by (rule set_eqI) auto
           from Nless have "finite N" by (auto elim: max_ext.cases)
           then have finites: "finite ?N1" "finite ?N2" by auto
-          
+
           have "?N2 \<in> ?W"
           proof cases
             assume [simp]: "M = {}"
             have Mw: "{} \<in> ?W" by (rule accI) (auto elim: max_ext.cases)
 
-            from asm1 have "?N2 = {}" by auto
+            from * have "?N2 = {}" by auto
             with Mw show "?N2 \<in> ?W" by (simp only:)
           next
             assume "M \<noteq> {}"
-            from asm1 finites have N2: "(?N2, M) \<in> max_ext r" 
+            from * finites have N2: "(?N2, M) \<in> max_ext r"
               by (rule_tac max_extI[OF _ _ \<open>M \<noteq> {}\<close>]) auto
-
             with \<open>M \<in> ?W\<close> show "?N2 \<in> ?W" by (rule acc_downward)
           qed
-          with finites have "?N1 \<union> ?N2 \<in> ?W" 
+          with finites have "?N1 \<union> ?N2 \<in> ?W"
             by (rule add_less) simp
           then show "N \<in> ?W" by (simp only: N)
         qed
@@ -863,32 +848,30 @@
   qed
 qed
 
-lemma max_ext_additive: 
- "(A, B) \<in> max_ext R \<Longrightarrow> (C, D) \<in> max_ext R \<Longrightarrow>
-  (A \<union> C, B \<union> D) \<in> max_ext R"
-by (force elim!: max_ext.cases)
+lemma max_ext_additive:
+  "(A, B) \<in> max_ext R \<Longrightarrow> (C, D) \<in> max_ext R \<Longrightarrow>
+    (A \<union> C, B \<union> D) \<in> max_ext R"
+  by (force elim!: max_ext.cases)
 
 
-definition min_ext :: "('a \<times> 'a) set \<Rightarrow> ('a set \<times> 'a set) set"  where
-  "min_ext r = {(X, Y) | X Y. X \<noteq> {} \<and> (\<forall>y \<in> Y. (\<exists>x \<in> X. (x, y) \<in> r))}"
+definition min_ext :: "('a \<times> 'a) set \<Rightarrow> ('a set \<times> 'a set) set"
+  where "min_ext r = {(X, Y) | X Y. X \<noteq> {} \<and> (\<forall>y \<in> Y. (\<exists>x \<in> X. (x, y) \<in> r))}"
 
 lemma min_ext_wf:
   assumes "wf r"
   shows "wf (min_ext r)"
 proof (rule wfI_min)
-  fix Q :: "'a set set"
-  fix x
-  assume nonempty: "x \<in> Q"
-  show "\<exists>m \<in> Q. (\<forall> n. (n, m) \<in> min_ext r \<longrightarrow> n \<notin> Q)"
-  proof cases
-    assume "Q = {{}}" thus ?thesis by (simp add: min_ext_def)
+  show "\<exists>m \<in> Q. (\<forall> n. (n, m) \<in> min_ext r \<longrightarrow> n \<notin> Q)" if nonempty: "x \<in> Q"
+    for Q :: "'a set set" and x
+  proof (cases "Q = {{}}")
+    case True
+    then show ?thesis by (simp add: min_ext_def)
   next
-    assume "Q \<noteq> {{}}"
-    with nonempty
-    obtain e x where "x \<in> Q" "e \<in> x" by force
+    case False
+    with nonempty obtain e x where "x \<in> Q" "e \<in> x" by force
     then have eU: "e \<in> \<Union>Q" by auto
-    with \<open>wf r\<close> 
-    obtain z where z: "z \<in> \<Union>Q" "\<And>y. (y, z) \<in> r \<Longrightarrow> y \<notin> \<Union>Q" 
+    with \<open>wf r\<close>
+    obtain z where z: "z \<in> \<Union>Q" "\<And>y. (y, z) \<in> r \<Longrightarrow> y \<notin> \<Union>Q"
       by (erule wfE_min)
     from z obtain m where "m \<in> Q" "z \<in> m" by auto
     from \<open>m \<in> Q\<close>
@@ -898,36 +881,38 @@
       assume smaller: "(n, m) \<in> min_ext r"
       with \<open>z \<in> m\<close> obtain y where y: "y \<in> n" "(y, z) \<in> r" by (auto simp: min_ext_def)
       then show "n \<notin> Q" using z(2) by auto
-    qed      
+    qed
   qed
 qed
 
-text\<open>Bounded increase must terminate:\<close>
+
+subsubsection \<open>Bounded increase must terminate\<close>
 
 lemma wf_bounded_measure:
-fixes ub :: "'a \<Rightarrow> nat" and f :: "'a \<Rightarrow> nat"
-assumes "!!a b. (b,a) : r \<Longrightarrow> ub b \<le> ub a & ub a \<ge> f b & f b > f a"
-shows "wf r"
-apply(rule wf_subset[OF wf_measure[of "%a. ub a - f a"]])
-apply (auto dest: assms)
-done
+  fixes ub :: "'a \<Rightarrow> nat"
+    and f :: "'a \<Rightarrow> nat"
+  assumes "\<And>a b. (b, a) \<in> r \<Longrightarrow> ub b \<le> ub a \<and> ub a \<ge> f b \<and> f b > f a"
+  shows "wf r"
+  apply (rule wf_subset[OF wf_measure[of "\<lambda>a. ub a - f a"]])
+  apply (auto dest: assms)
+  done
 
 lemma wf_bounded_set:
-fixes ub :: "'a \<Rightarrow> 'b set" and f :: "'a \<Rightarrow> 'b set"
-assumes "!!a b. (b,a) : r \<Longrightarrow>
-  finite(ub a) & ub b \<subseteq> ub a & ub a \<supseteq> f b & f b \<supset> f a"
-shows "wf r"
-apply(rule wf_bounded_measure[of r "%a. card(ub a)" "%a. card(f a)"])
-apply(drule assms)
-apply (blast intro: card_mono finite_subset psubset_card_mono dest: psubset_eq[THEN iffD2])
-done
+  fixes ub :: "'a \<Rightarrow> 'b set"
+    and f :: "'a \<Rightarrow> 'b set"
+  assumes "\<And>a b. (b,a) \<in> r \<Longrightarrow> finite (ub a) \<and> ub b \<subseteq> ub a \<and> ub a \<supseteq> f b \<and> f b \<supset> f a"
+  shows "wf r"
+  apply(rule wf_bounded_measure[of r "\<lambda>a. card(ub a)" "\<lambda>a. card(f a)"])
+  apply(drule assms)
+  apply (blast intro: card_mono finite_subset psubset_card_mono dest: psubset_eq[THEN iffD2])
+  done
 
 lemma finite_subset_wf:
   assumes "finite A"
   shows   "wf {(X,Y). X \<subset> Y \<and> Y \<subseteq> A}"
 proof (intro finite_acyclic_wf)
   have "{(X,Y). X \<subset> Y \<and> Y \<subseteq> A} \<subseteq> Pow A \<times> Pow A" by blast
-  thus "finite {(X,Y). X \<subset> Y \<and> Y \<subseteq> A}" 
+  then show "finite {(X,Y). X \<subset> Y \<and> Y \<subseteq> A}"
     by (rule finite_subset) (auto simp: assms finite_cartesian_product)
 next
   have "{(X, Y). X \<subset> Y \<and> Y \<subseteq> A}\<^sup>+ = {(X, Y). X \<subset> Y \<and> Y \<subseteq> A}"