src/HOL/MacLaurin.thy
changeset 63569 7e0b0db5e9ac
parent 63365 5340fb6633d0
child 63570 1826a90b9cbc
--- a/src/HOL/MacLaurin.thy	Sat Jul 30 21:10:02 2016 +0200
+++ b/src/HOL/MacLaurin.thy	Sun Jul 31 17:25:38 2016 +0200
@@ -1,124 +1,105 @@
-(*  Author      : Jacques D. Fleuriot
-    Copyright   : 2001 University of Edinburgh
-    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
-    Conversion of Mac Laurin to Isar by Lukas Bulwahn and Bernhard Häupler, 2005
+(*  Title:      HOL/MacLaurin.thy
+    Author:     Jacques D. Fleuriot, 2001 University of Edinburgh
+    Author:     Lawrence C Paulson, 2004
+    Author:     Lukas Bulwahn and Bernhard Häupler, 2005
 *)
 
-section\<open>MacLaurin Series\<close>
+section \<open>MacLaurin Series\<close>
 
 theory MacLaurin
 imports Transcendental
 begin
 
-subsection\<open>Maclaurin's Theorem with Lagrange Form of Remainder\<close>
+subsection \<open>Maclaurin's Theorem with Lagrange Form of Remainder\<close>
 
-text\<open>This is a very long, messy proof even now that it's been broken down
-into lemmas.\<close>
+text \<open>This is a very long, messy proof even now that it's been broken down
+  into lemmas.\<close>
 
 lemma Maclaurin_lemma:
-    "0 < h ==>
-     \<exists>B::real. f h = (\<Sum>m<n. (j m / (fact m)) * (h^m)) +
-               (B * ((h^n) /(fact n)))"
-by (rule exI[where x = "(f h - (\<Sum>m<n. (j m / (fact m)) * h^m)) * (fact n) / (h^n)"]) simp
+  "0 < h \<Longrightarrow>
+    \<exists>B::real. f h = (\<Sum>m<n. (j m / (fact m)) * (h^m)) + (B * ((h^n) /(fact n)))"
+  by (rule exI[where x = "(f h - (\<Sum>m<n. (j m / (fact m)) * h^m)) * (fact n) / (h^n)"]) simp
 
-lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))"
-by arith
+lemma eq_diff_eq': "x = y - z \<longleftrightarrow> y = x + z"
+  for x y z :: real
+  by arith
 
-lemma fact_diff_Suc:
-  "n < Suc m \<Longrightarrow> fact (Suc m - n) = (Suc m - n) * fact (m - n)"
-  by (subst fact_reduce, auto)
+lemma fact_diff_Suc: "n < Suc m \<Longrightarrow> fact (Suc m - n) = (Suc m - n) * fact (m - n)"
+  by (subst fact_reduce) auto
 
 lemma Maclaurin_lemma2:
   fixes B
-  assumes DERIV : "\<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
-      and INIT : "n = Suc k"
+  assumes DERIV: "\<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
+    and INIT: "n = Suc k"
   defines "difg \<equiv>
-      (\<lambda>m t::real. diff m t -
-         ((\<Sum>p<n - m. diff (m + p) 0 / (fact p) * t ^ p) + B * (t ^ (n - m) / (fact (n - m)))))"
-        (is "difg \<equiv> (\<lambda>m t. diff m t - ?difg m t)")
-  shows "\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (difg m) t :> difg (Suc m) t"
+    (\<lambda>m t::real. diff m t -
+      ((\<Sum>p<n - m. diff (m + p) 0 / fact p * t ^ p) + B * (t ^ (n - m) / fact (n - m))))"
+    (is "difg \<equiv> (\<lambda>m t. diff m t - ?difg m t)")
+  shows "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
 proof (rule allI impI)+
-  fix m and t::real
-  assume INIT2: "m < n & 0 \<le> t & t \<le> h"
+  fix m t
+  assume INIT2: "m < n \<and> 0 \<le> t \<and> t \<le> h"
   have "DERIV (difg m) t :> diff (Suc m) t -
-    ((\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / (fact x)) +
-     real (n - m) * t ^ (n - Suc m) * B / (fact (n - m)))"
-    unfolding difg_def
-    by (auto intro!: derivative_eq_intros DERIV[rule_format, OF INIT2])
+    ((\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / fact x) +
+     real (n - m) * t ^ (n - Suc m) * B / fact (n - m))"
+    by (auto simp: difg_def intro!: derivative_eq_intros DERIV[rule_format, OF INIT2])
   moreover
   from INIT2 have intvl: "{..<n - m} = insert 0 (Suc ` {..<n - Suc m})" and "0 < n - m"
     unfolding atLeast0LessThan[symmetric] by auto
-  have "(\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / (fact x)) =
-      (\<Sum>x<n - Suc m. real (Suc x) * t ^ x * diff (Suc m + x) 0 / (fact (Suc x)))"
-    unfolding intvl
-    by (subst setsum.insert) (auto simp add: setsum.reindex)
+  have "(\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / fact x) =
+      (\<Sum>x<n - Suc m. real (Suc x) * t ^ x * diff (Suc m + x) 0 / fact (Suc x))"
+    unfolding intvl by (subst setsum.insert) (auto simp add: setsum.reindex)
   moreover
   have fact_neq_0: "\<And>x. (fact x) + real x * (fact x) \<noteq> 0"
-    by (metis add_pos_pos fact_gt_zero less_add_same_cancel1 less_add_same_cancel2 less_numeral_extra(3) mult_less_0_iff of_nat_less_0_iff)
-  have "\<And>x. (Suc x) * t ^ x * diff (Suc m + x) 0 / (fact (Suc x)) =
-            diff (Suc m + x) 0 * t^x / (fact x)"
+    by (metis add_pos_pos fact_gt_zero less_add_same_cancel1 less_add_same_cancel2
+        less_numeral_extra(3) mult_less_0_iff of_nat_less_0_iff)
+  have "\<And>x. (Suc x) * t ^ x * diff (Suc m + x) 0 / fact (Suc x) = diff (Suc m + x) 0 * t^x / fact x"
     by (rule nonzero_divide_eq_eq[THEN iffD2]) auto
   moreover
-  have "(n - m) * t ^ (n - Suc m) * B / (fact (n - m)) =
-        B * (t ^ (n - Suc m) / (fact (n - Suc m)))"
-    using \<open>0 < n - m\<close>
-    by (simp add: divide_simps fact_reduce)
+  have "(n - m) * t ^ (n - Suc m) * B / fact (n - m) = B * (t ^ (n - Suc m) / fact (n - Suc m))"
+    using \<open>0 < n - m\<close> by (simp add: divide_simps fact_reduce)
   ultimately show "DERIV (difg m) t :> difg (Suc m) t"
     unfolding difg_def  by (simp add: mult.commute)
 qed
 
 lemma Maclaurin:
   assumes h: "0 < h"
-  assumes n: "0 < n"
-  assumes diff_0: "diff 0 = f"
-  assumes diff_Suc:
-    "\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t"
+    and n: "0 < n"
+    and diff_0: "diff 0 = f"
+    and diff_Suc: "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
   shows
-    "\<exists>t::real. 0 < t & t < h &
-              f h =
-              setsum (%m. (diff m 0 / (fact m)) * h ^ m) {..<n} +
-              (diff n t / (fact n)) * h ^ n"
+    "\<exists>t::real. 0 < t \<and> t < h \<and>
+      f h = setsum (\<lambda>m. (diff m 0 / fact m) * h ^ m) {..<n} + (diff n t / fact n) * h ^ n"
 proof -
   from n obtain m where m: "n = Suc m"
     by (cases n) (simp add: n)
+  from m have "m < n" by simp
 
-  obtain B where f_h: "f h =
-        (\<Sum>m<n. diff m (0::real) / (fact m) * h ^ m) + B * (h ^ n / (fact n))"
+  obtain B where f_h: "f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + B * (h ^ n / fact n)"
     using Maclaurin_lemma [OF h] ..
 
   define g where [abs_def]: "g t =
-    f t - (setsum (\<lambda>m. (diff m 0 / (fact m)) * t^m) {..<n} + (B * (t^n / (fact n))))" for t
-
-  have g2: "g 0 = 0 & g h = 0"
-    by (simp add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 setsum.reindex)
+    f t - (setsum (\<lambda>m. (diff m 0 / fact m) * t^m) {..<n} + B * (t^n / fact n))" for t
+  have g2: "g 0 = 0" "g h = 0"
+    by (simp_all add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 setsum.reindex)
 
   define difg where [abs_def]: "difg m t =
-    diff m t - (setsum (%p. (diff (m + p) 0 / (fact p)) * (t ^ p)) {..<n-m}
-      + (B * ((t ^ (n - m)) / (fact (n - m)))))" for m t
-
+    diff m t - (setsum (\<lambda>p. (diff (m + p) 0 / fact p) * (t ^ p)) {..<n-m} +
+      B * ((t ^ (n - m)) / fact (n - m)))" for m t
   have difg_0: "difg 0 = g"
-    unfolding difg_def g_def by (simp add: diff_0)
-
-  have difg_Suc: "\<forall>(m::nat) t::real.
-        m < n \<and> (0::real) \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
+    by (simp add: difg_def g_def diff_0)
+  have difg_Suc: "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
     using diff_Suc m unfolding difg_def [abs_def] by (rule Maclaurin_lemma2)
-
   have difg_eq_0: "\<forall>m<n. difg m 0 = 0"
     by (auto simp: difg_def m Suc_diff_le lessThan_Suc_eq_insert_0 image_iff setsum.reindex)
-
-  have isCont_difg: "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> isCont (difg m) x"
+  have isCont_difg: "\<And>m x. m < n \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<le> h \<Longrightarrow> isCont (difg m) x"
     by (rule DERIV_isCont [OF difg_Suc [rule_format]]) simp
-
-  have differentiable_difg:
-    "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> difg m differentiable (at x)"
+  have differentiable_difg: "\<And>m x. m < n \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<le> h \<Longrightarrow> difg m differentiable (at x)"
     by (rule differentiableI [OF difg_Suc [rule_format]]) simp
-
-  have difg_Suc_eq_0: "\<And>m t. \<lbrakk>m < n; 0 \<le> t; t \<le> h; DERIV (difg m) t :> 0\<rbrakk>
-        \<Longrightarrow> difg (Suc m) t = 0"
+  have difg_Suc_eq_0:
+    "\<And>m t. m < n \<Longrightarrow> 0 \<le> t \<Longrightarrow> t \<le> h \<Longrightarrow> DERIV (difg m) t :> 0 \<Longrightarrow> difg (Suc m) t = 0"
     by (rule DERIV_unique [OF difg_Suc [rule_format]]) simp
 
-  have "m < n" using m by simp
-
   have "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m) t :> 0"
   using \<open>m < n\<close>
   proof (induct m)
@@ -126,7 +107,8 @@
     show ?case
     proof (rule Rolle)
       show "0 < h" by fact
-      show "difg 0 0 = difg 0 h" by (simp add: difg_0 g2)
+      show "difg 0 0 = difg 0 h"
+        by (simp add: difg_0 g2)
       show "\<forall>x. 0 \<le> x \<and> x \<le> h \<longrightarrow> isCont (difg (0::nat)) x"
         by (simp add: isCont_difg n)
       show "\<forall>x. 0 < x \<and> x < h \<longrightarrow> difg (0::nat) differentiable (at x)"
@@ -134,8 +116,10 @@
     qed
   next
     case (Suc m')
-    hence "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m') t :> 0" by simp
-    then obtain t where t: "0 < t" "t < h" "DERIV (difg m') t :> 0" by fast
+    then have "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m') t :> 0"
+      by simp
+    then obtain t where t: "0 < t" "t < h" "DERIV (difg m') t :> 0"
+      by fast
     have "\<exists>t'. 0 < t' \<and> t' < t \<and> DERIV (difg (Suc m')) t' :> 0"
     proof (rule Rolle)
       show "0 < t" by fact
@@ -146,448 +130,429 @@
       show "\<forall>x. 0 < x \<and> x < t \<longrightarrow> difg (Suc m') differentiable (at x)"
         using \<open>t < h\<close> \<open>Suc m' < n\<close> by (simp add: differentiable_difg)
     qed
-    thus ?case
-      using \<open>t < h\<close> by auto
+    with \<open>t < h\<close> show ?case
+      by auto
   qed
-  then obtain t where "0 < t" "t < h" "DERIV (difg m) t :> 0" by fast
-
-  hence "difg (Suc m) t = 0"
-    using \<open>m < n\<close> by (simp add: difg_Suc_eq_0)
-
+  then obtain t where "0 < t" "t < h" "DERIV (difg m) t :> 0"
+    by fast
+  with \<open>m < n\<close> have "difg (Suc m) t = 0"
+    by (simp add: difg_Suc_eq_0)
   show ?thesis
   proof (intro exI conjI)
     show "0 < t" by fact
     show "t < h" by fact
     show "f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / (fact n) * h ^ n"
-      using \<open>difg (Suc m) t = 0\<close>
-      by (simp add: m f_h difg_def)
+      using \<open>difg (Suc m) t = 0\<close> by (simp add: m f_h difg_def)
   qed
 qed
 
 lemma Maclaurin_objl:
-  "0 < h & n>0 & diff 0 = f &
-  (\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
-   --> (\<exists>t::real. 0 < t & t < h &
-            f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
-                  diff n t / (fact n) * h ^ n)"
-by (blast intro: Maclaurin)
-
+  "0 < h \<and> n > 0 \<and> diff 0 = f \<and>
+    (\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t) \<longrightarrow>
+    (\<exists>t. 0 < t \<and> t < h \<and> f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / fact n * h ^ n)"
+  for n :: nat and h :: real
+  by (blast intro: Maclaurin)
 
 lemma Maclaurin2:
-  assumes INIT1: "0 < h " and INIT2: "diff 0 = f"
-  and DERIV: "\<forall>m t::real.
-  m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t"
-  shows "\<exists>t. 0 < t \<and> t \<le> h \<and> f h =
-  (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
-  diff n t / (fact n) * h ^ n"
-proof (cases "n")
-  case 0 with INIT1 INIT2 show ?thesis by fastforce
+  fixes n :: nat
+    and h :: real
+  assumes INIT1: "0 < h"
+    and INIT2: "diff 0 = f"
+    and DERIV: "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
+  shows "\<exists>t. 0 < t \<and> t \<le> h \<and> f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / fact n * h ^ n"
+proof (cases n)
+  case 0
+  with INIT1 INIT2 show ?thesis by fastforce
 next
   case Suc
-  hence "n > 0" by simp
-  from INIT1 this INIT2 DERIV have "\<exists>t>0. t < h \<and>
-    f h =
-    (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / (fact n) * h ^ n"
+  then have "n > 0" by simp
+  from INIT1 this INIT2 DERIV
+  have "\<exists>t>0. t < h \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n"
     by (rule Maclaurin)
-  thus ?thesis by fastforce
+  then show ?thesis by fastforce
 qed
 
 lemma Maclaurin2_objl:
-     "0 < h & diff 0 = f &
-       (\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
-    --> (\<exists>t::real. 0 < t &
-              t \<le> h &
-              f h =
-              (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
-              diff n t / (fact n) * h ^ n)"
-by (blast intro: Maclaurin2)
+  "0 < h \<and> diff 0 = f \<and>
+    (\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t) \<longrightarrow>
+    (\<exists>t. 0 < t \<and> t \<le> h \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n)"
+  for n :: nat and h :: real
+  by (blast intro: Maclaurin2)
 
 lemma Maclaurin_minus:
-  fixes h::real
+  fixes n :: nat and h :: real
   assumes "h < 0" "0 < n" "diff 0 = f"
-  and DERIV: "\<forall>m t. m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t"
-  shows "\<exists>t. h < t & t < 0 &
-         f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
-         diff n t / (fact n) * h ^ n"
+    and DERIV: "\<forall>m t. m < n \<and> h \<le> t \<and> t \<le> 0 \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
+  shows "\<exists>t. h < t \<and> t < 0 \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n"
 proof -
-  txt "Transform \<open>ABL'\<close> into \<open>derivative_intros\<close> format."
+  txt \<open>Transform \<open>ABL'\<close> into \<open>derivative_intros\<close> format.\<close>
   note DERIV' = DERIV_chain'[OF _ DERIV[rule_format], THEN DERIV_cong]
-  from assms
-  have "\<exists>t>0. t < - h \<and>
-    f (- (- h)) =
-    (\<Sum>m<n.
-    (- 1) ^ m * diff m (- 0) / (fact m) * (- h) ^ m) +
+  let ?sum = "\<lambda>t.
+    (\<Sum>m<n. (- 1) ^ m * diff m (- 0) / (fact m) * (- h) ^ m) +
     (- 1) ^ n * diff n (- t) / (fact n) * (- h) ^ n"
+  from assms have "\<exists>t>0. t < - h \<and> f (- (- h)) = ?sum t"
     by (intro Maclaurin) (auto intro!: derivative_eq_intros DERIV')
-  then guess t ..
-  moreover
-  have "(- 1) ^ n * diff n (- t) * (- h) ^ n / (fact n) = diff n (- t) * h ^ n / (fact n)"
-    by (auto simp add: power_mult_distrib[symmetric])
+  then obtain t where "0 < t" "t < - h" "f (- (- h)) = ?sum t"
+    by blast
+  moreover have "(- 1) ^ n * diff n (- t) * (- h) ^ n / fact n = diff n (- t) * h ^ n / fact n"
+    by (auto simp: power_mult_distrib[symmetric])
   moreover
-  have "(\<Sum>m<n. (- 1) ^ m * diff m 0 * (- h) ^ m / (fact m)) = (\<Sum>m<n. diff m 0 * h ^ m / (fact m))"
+    have "(\<Sum>m<n. (- 1) ^ m * diff m 0 * (- h) ^ m / fact m) = (\<Sum>m<n. diff m 0 * h ^ m / fact m)"
     by (auto intro: setsum.cong simp add: power_mult_distrib[symmetric])
-  ultimately have " h < - t \<and>
-    - t < 0 \<and>
-    f h =
-    (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n (- t) / (fact n) * h ^ n"
+  ultimately have "h < - t \<and> - t < 0 \<and>
+    f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n (- t) / (fact n) * h ^ n"
     by auto
-  thus ?thesis ..
+  then show ?thesis ..
 qed
 
 lemma Maclaurin_minus_objl:
-  fixes h::real
+  fixes n :: nat and h :: real
   shows
-     "(h < 0 & n > 0 & diff 0 = f &
-       (\<forall>m t.
-          m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t))
-    --> (\<exists>t. h < t &
-              t < 0 &
-              f h =
-              (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) +
-              diff n t / (fact n) * h ^ n)"
-by (blast intro: Maclaurin_minus)
+    "h < 0 \<and> n > 0 \<and> diff 0 = f \<and>
+      (\<forall>m t. m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t) \<longrightarrow>
+    (\<exists>t. h < t \<and> t < 0 \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n)"
+  by (blast intro: Maclaurin_minus)
 
 
-subsection\<open>More Convenient "Bidirectional" Version.\<close>
+subsection \<open>More Convenient "Bidirectional" Version.\<close>
 
 (* not good for PVS sin_approx, cos_approx *)
 
 lemma Maclaurin_bi_le_lemma:
-  "n>0 \<Longrightarrow>
-   diff 0 0 =
-   (\<Sum>m<n. diff m 0 * 0 ^ m / (fact m)) + diff n 0 * 0 ^ n / (fact n :: real)"
-by (induct "n") auto
+  "n > 0 \<Longrightarrow>
+    diff 0 0 = (\<Sum>m<n. diff m 0 * 0 ^ m / (fact m)) + diff n 0 * 0 ^ n / (fact n :: real)"
+  by (induct n) auto
 
 lemma Maclaurin_bi_le:
-   assumes "diff 0 = f"
-   and DERIV : "\<forall>m t::real. m < n & \<bar>t\<bar> \<le> \<bar>x\<bar> --> DERIV (diff m) t :> diff (Suc m) t"
-   shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> &
-              f x =
-              (\<Sum>m<n. diff m 0 / (fact m) * x ^ m) +
-     diff n t / (fact n) * x ^ n" (is "\<exists>t. _ \<and> f x = ?f x t")
-proof cases
-  assume "n = 0" with \<open>diff 0 = f\<close> show ?thesis by force
+  fixes n :: nat and x :: real
+  assumes "diff 0 = f"
+    and DERIV : "\<forall>m t. m < n \<and> \<bar>t\<bar> \<le> \<bar>x\<bar> \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
+  shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = (\<Sum>m<n. diff m 0 / (fact m) * x ^ m) + diff n t / (fact n) * x ^ n"
+    (is "\<exists>t. _ \<and> f x = ?f x t")
+proof (cases "n = 0")
+  case True
+  with \<open>diff 0 = f\<close> show ?thesis by force
 next
-  assume "n \<noteq> 0"
+  case False
   show ?thesis
   proof (cases rule: linorder_cases)
-    assume "x = 0" with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV
-    have "\<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0" by (auto simp add: Maclaurin_bi_le_lemma)
-    thus ?thesis ..
+    assume "x = 0"
+    with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV have "\<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0"
+      by (auto simp add: Maclaurin_bi_le_lemma)
+    then show ?thesis ..
   next
     assume "x < 0"
-    with \<open>n \<noteq> 0\<close> DERIV
-    have "\<exists>t>x. t < 0 \<and> diff 0 x = ?f x t" by (intro Maclaurin_minus) auto
-    then guess t ..
-    with \<open>x < 0\<close> \<open>diff 0 = f\<close> have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
-    thus ?thesis ..
+    with \<open>n \<noteq> 0\<close> DERIV have "\<exists>t>x. t < 0 \<and> diff 0 x = ?f x t"
+      by (intro Maclaurin_minus) auto
+    then obtain t where "x < t" "t < 0"
+      "diff 0 x = (\<Sum>m<n. diff m 0 / fact m * x ^ m) + diff n t / fact n * x ^ n"
+      by blast
+    with \<open>x < 0\<close> \<open>diff 0 = f\<close> have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t"
+      by simp
+    then show ?thesis ..
   next
     assume "x > 0"
-    with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV
-    have "\<exists>t>0. t < x \<and> diff 0 x = ?f x t" by (intro Maclaurin) auto
-    then guess t ..
+    with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV have "\<exists>t>0. t < x \<and> diff 0 x = ?f x t"
+      by (intro Maclaurin) auto
+    then obtain t where "0 < t" "t < x"
+      "diff 0 x = (\<Sum>m<n. diff m 0 / fact m * x ^ m) + diff n t / fact n * x ^ n"
+      by blast
     with \<open>x > 0\<close> \<open>diff 0 = f\<close> have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
-    thus ?thesis ..
+    then show ?thesis ..
   qed
 qed
 
 lemma Maclaurin_all_lt:
-  fixes x::real
-  assumes INIT1: "diff 0 = f" and INIT2: "0 < n" and INIT3: "x \<noteq> 0"
-  and DERIV: "\<forall>m x. DERIV (diff m) x :> diff(Suc m) x"
-  shows "\<exists>t. 0 < \<bar>t\<bar> & \<bar>t\<bar> < \<bar>x\<bar> & f x =
-    (\<Sum>m<n. (diff m 0 / (fact m)) * x ^ m) +
-                (diff n t / (fact n)) * x ^ n" (is "\<exists>t. _ \<and> _ \<and> f x = ?f x t")
+  fixes x :: real
+  assumes INIT1: "diff 0 = f"
+    and INIT2: "0 < n"
+    and INIT3: "x \<noteq> 0"
+    and DERIV: "\<forall>m x. DERIV (diff m) x :> diff(Suc m) x"
+  shows "\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x =
+      (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) + (diff n t / fact n) * x ^ n"
+    (is "\<exists>t. _ \<and> _ \<and> f x = ?f x t")
 proof (cases rule: linorder_cases)
-  assume "x = 0" with INIT3 show "?thesis"..
+  assume "x = 0"
+  with INIT3 show ?thesis ..
 next
   assume "x < 0"
-  with assms have "\<exists>t>x. t < 0 \<and> f x = ?f x t" by (intro Maclaurin_minus) auto
-  then guess t ..
-  with \<open>x < 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
-  thus ?thesis ..
+  with assms have "\<exists>t>x. t < 0 \<and> f x = ?f x t"
+    by (intro Maclaurin_minus) auto
+  then obtain t where "t > x" "t < 0" "f x = ?f x t"
+    by blast
+  with \<open>x < 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
+    by simp
+  then show ?thesis ..
 next
   assume "x > 0"
-  with assms have "\<exists>t>0. t < x \<and> f x = ?f x t " by (intro Maclaurin) auto
-  then guess t ..
-  with \<open>x > 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
-  thus ?thesis ..
+  with assms have "\<exists>t>0. t < x \<and> f x = ?f x t"
+    by (intro Maclaurin) auto
+  then obtain t where "t > 0" "t < x" "f x = ?f x t"
+    by blast
+  with \<open>x > 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
+    by simp
+  then show ?thesis ..
 qed
 
 
 lemma Maclaurin_all_lt_objl:
-  fixes x::real
+  fixes x :: real
   shows
-     "diff 0 = f &
-      (\<forall>m x. DERIV (diff m) x :> diff(Suc m) x) &
-      x ~= 0 & n > 0
-      --> (\<exists>t. 0 < \<bar>t\<bar> & \<bar>t\<bar> < \<bar>x\<bar> &
-               f x = (\<Sum>m<n. (diff m 0 / (fact m)) * x ^ m) +
-                     (diff n t / (fact n)) * x ^ n)"
-by (blast intro: Maclaurin_all_lt)
+    "diff 0 = f \<and> (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x) \<and> x \<noteq> 0 \<and> n > 0 \<longrightarrow>
+    (\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and>
+      f x = (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) + (diff n t / fact n) * x ^ n)"
+  by (blast intro: Maclaurin_all_lt)
 
-lemma Maclaurin_zero [rule_format]:
-     "x = (0::real)
-      ==> n \<noteq> 0 -->
-          (\<Sum>m<n. (diff m (0::real) / (fact m)) * x ^ m) =
-          diff 0 0"
-by (induct n, auto)
+lemma Maclaurin_zero: "x = 0 \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) = diff 0 0"
+  for x :: real and n :: nat
+  by (induct n) auto
 
 
 lemma Maclaurin_all_le:
+  fixes x :: real and n :: nat
   assumes INIT: "diff 0 = f"
-  and DERIV: "\<forall>m x::real. DERIV (diff m) x :> diff (Suc m) x"
-  shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> & f x =
-    (\<Sum>m<n. (diff m 0 / (fact m)) * x ^ m) +
-    (diff n t / (fact n)) * x ^ n" (is "\<exists>t. _ \<and> f x = ?f x t")
-proof cases
-  assume "n = 0" with INIT show ?thesis by force
-  next
-  assume "n \<noteq> 0"
+    and DERIV: "\<forall>m x. DERIV (diff m) x :> diff (Suc m) x"
+  shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) + (diff n t / fact n) * x ^ n"
+    (is "\<exists>t. _ \<and> f x = ?f x t")
+proof (cases "n = 0")
+  case True
+  with INIT show ?thesis by force
+next
+  case False
   show ?thesis
-  proof cases
-    assume "x = 0"
+  proof (cases "x = 0")
+    case True
     with \<open>n \<noteq> 0\<close> have "(\<Sum>m<n. diff m 0 / (fact m) * x ^ m) = diff 0 0"
       by (intro Maclaurin_zero) auto
-    with INIT \<open>x = 0\<close> \<open>n \<noteq> 0\<close> have " \<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0" by force
-    thus ?thesis ..
+    with INIT \<open>x = 0\<close> \<open>n \<noteq> 0\<close> have " \<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0"
+      by force
+    then show ?thesis ..
   next
-    assume "x \<noteq> 0"
+    case False
     with INIT \<open>n \<noteq> 0\<close> DERIV have "\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
       by (intro Maclaurin_all_lt) auto
-    then guess t ..
-    hence "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
-    thus ?thesis ..
+    then obtain t where "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" ..
+    then have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t"
+      by simp
+    then show ?thesis ..
   qed
 qed
 
 lemma Maclaurin_all_le_objl:
-  "diff 0 = f &
-      (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x)
-      --> (\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> &
-              f x = (\<Sum>m<n. (diff m 0 / (fact m)) * x ^ m) +
-                    (diff n t / (fact n)) * x ^ n)"
-by (blast intro: Maclaurin_all_le)
+  "diff 0 = f \<and> (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x) \<longrightarrow>
+    (\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) + (diff n t / fact n) * x ^ n)"
+  for x :: real and n :: nat
+  by (blast intro: Maclaurin_all_le)
 
 
-subsection\<open>Version for Exponential Function\<close>
+subsection \<open>Version for Exponential Function\<close>
 
 lemma Maclaurin_exp_lt:
-  fixes x::real
+  fixes x :: real and n :: nat
   shows
-  "[| x ~= 0; n > 0 |]
-      ==> (\<exists>t. 0 < \<bar>t\<bar> &
-                \<bar>t\<bar> < \<bar>x\<bar> &
-                exp x = (\<Sum>m<n. (x ^ m) / (fact m)) +
-                        (exp t / (fact n)) * x ^ n)"
-by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_lt_objl, auto)
-
+    "x \<noteq> 0 \<Longrightarrow> n > 0 \<Longrightarrow>
+      (\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> exp x = (\<Sum>m<n. (x ^ m) / fact m) + (exp t / fact n) * x ^ n)"
+ using Maclaurin_all_lt_objl [where diff = "\<lambda>n. exp" and f = exp and x = x and n = n] by auto
 
 lemma Maclaurin_exp_le:
-     "\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> &
-            exp x = (\<Sum>m<n. (x ^ m) / (fact m)) +
-                       (exp t / (fact n)) * x ^ n"
-by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_le_objl, auto)
+  fixes x :: real and n :: nat
+  shows "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and> exp x = (\<Sum>m<n. (x ^ m) / fact m) + (exp t / fact n) * x ^ n"
+  using Maclaurin_all_le_objl [where diff = "\<lambda>n. exp" and f = exp and x = x and n = n] by auto
+
+lemma exp_lower_taylor_quadratic: "0 \<le> x \<Longrightarrow> 1 + x + x\<^sup>2 / 2 \<le> exp x"
+  for x :: real
+  using Maclaurin_exp_le [of x 3] by (auto simp: numeral_3_eq_3 power2_eq_square)
+
+
+subsection \<open>Version for Sine Function\<close>
 
-lemma exp_lower_taylor_quadratic:
-  fixes x::real
-  shows "0 \<le> x \<Longrightarrow> 1 + x + x\<^sup>2 / 2 \<le> exp x"
-  using Maclaurin_exp_le [of x 3]
-  by (auto simp: numeral_3_eq_3 power2_eq_square power_Suc)
+lemma mod_exhaust_less_4: "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = 3"
+  for m :: nat
+  by auto
+
+lemma Suc_Suc_mult_two_diff_two [simp]: "n \<noteq> 0 \<Longrightarrow> Suc (Suc (2 * n - 2)) = 2 * n"
+  by (induct n) auto
+
+lemma lemma_Suc_Suc_4n_diff_2 [simp]: "n \<noteq> 0 \<Longrightarrow> Suc (Suc (4 * n - 2)) = 4 * n"
+  by (induct n) auto
+
+lemma Suc_mult_two_diff_one [simp]: "n \<noteq> 0 \<Longrightarrow> Suc (2 * n - 1) = 2 * n"
+  by (induct n) auto
 
 
-subsection\<open>Version for Sine Function\<close>
-
-lemma mod_exhaust_less_4:
-  "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = (3::nat)"
-by auto
-
-lemma Suc_Suc_mult_two_diff_two [rule_format, simp]:
-  "n\<noteq>0 --> Suc (Suc (2 * n - 2)) = 2*n"
-by (induct "n", auto)
+text \<open>It is unclear why so many variant results are needed.\<close>
 
-lemma lemma_Suc_Suc_4n_diff_2 [rule_format, simp]:
-  "n\<noteq>0 --> Suc (Suc (4*n - 2)) = 4*n"
-by (induct "n", auto)
-
-lemma Suc_mult_two_diff_one [rule_format, simp]:
-  "n\<noteq>0 --> Suc (2 * n - 1) = 2*n"
-by (induct "n", auto)
-
-
-text\<open>It is unclear why so many variant results are needed.\<close>
-
-lemma sin_expansion_lemma:
-     "sin (x + real (Suc m) * pi / 2) =
-      cos (x + real (m) * pi / 2)"
-by (simp only: cos_add sin_add of_nat_Suc add_divide_distrib distrib_right, auto)
+lemma sin_expansion_lemma: "sin (x + real (Suc m) * pi / 2) = cos (x + real m * pi / 2)"
+  by (auto simp: cos_add sin_add add_divide_distrib distrib_right)
 
 lemma Maclaurin_sin_expansion2:
-     "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> &
-       sin x =
-       (\<Sum>m<n. sin_coeff m * x ^ m)
-      + ((sin(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
-apply (cut_tac f = sin and n = n and x = x
-        and diff = "%n x. sin (x + 1/2*real n * pi)" in Maclaurin_all_lt_objl)
-apply safe
-    apply (simp)
-   apply (simp add: sin_expansion_lemma del: of_nat_Suc)
-   apply (force intro!: derivative_eq_intros)
-  apply (subst (asm) setsum.neutral, auto)[1]
- apply (rule ccontr, simp)
- apply (drule_tac x = x in spec, simp)
-apply (erule ssubst)
-apply (rule_tac x = t in exI, simp)
-apply (rule setsum.cong[OF refl])
-apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
-done
+  "\<exists>t. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and>
+    sin x = (\<Sum>m<n. sin_coeff m * x ^ m) + (sin (t + 1/2 * real n * pi) / fact n) * x ^ n"
+  using Maclaurin_all_lt_objl
+    [where f = sin and n = n and x = x and diff = "\<lambda>n x. sin (x + 1/2 * real n * pi)"]
+  apply safe
+      apply simp
+     apply (simp add: sin_expansion_lemma del: of_nat_Suc)
+     apply (force intro!: derivative_eq_intros)
+    apply (subst (asm) setsum.neutral; auto)
+   apply (rule ccontr)
+   apply simp
+   apply (drule_tac x = x in spec)
+   apply simp
+  apply (erule ssubst)
+  apply (rule_tac x = t in exI)
+  apply simp
+  apply (rule setsum.cong[OF refl])
+  apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
+  done
 
 lemma Maclaurin_sin_expansion:
-     "\<exists>t. sin x =
-       (\<Sum>m<n. sin_coeff m * x ^ m)
-      + ((sin(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
-apply (insert Maclaurin_sin_expansion2 [of x n])
-apply (blast intro: elim:)
-done
+  "\<exists>t. sin x = (\<Sum>m<n. sin_coeff m * x ^ m) + (sin (t + 1/2 * real n * pi) / fact n) * x ^ n"
+  using Maclaurin_sin_expansion2 [of x n] by blast
 
 lemma Maclaurin_sin_expansion3:
-     "[| n > 0; 0 < x |] ==>
-       \<exists>t. 0 < t & t < x &
-       sin x =
-       (\<Sum>m<n. sin_coeff m * x ^ m)
-      + ((sin(t + 1/2 * real(n) *pi) / (fact n)) * x ^ n)"
-apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin_objl)
-apply safe
+  "n > 0 \<Longrightarrow> 0 < x \<Longrightarrow>
+    \<exists>t. 0 < t \<and> t < x \<and>
+       sin x = (\<Sum>m<n. sin_coeff m * x ^ m) + (sin (t + 1/2 * real n * pi) / fact n) * x ^ n"
+  using Maclaurin_objl
+    [where f = sin and n = n and h = x and diff = "\<lambda>n x. sin (x + 1/2 * real n * pi)"]
+  apply safe
+    apply simp
+   apply (simp (no_asm) add: sin_expansion_lemma del: of_nat_Suc)
+   apply (force intro!: derivative_eq_intros)
+  apply (erule ssubst)
+  apply (rule_tac x = t in exI)
+  apply simp
+  apply (rule setsum.cong[OF refl])
+  apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
+  done
+
+lemma Maclaurin_sin_expansion4:
+  "0 < x \<Longrightarrow>
+    \<exists>t. 0 < t \<and> t \<le> x \<and>
+      sin x = (\<Sum>m<n. sin_coeff m * x ^ m) + (sin (t + 1/2 * real n * pi) / fact n) * x ^ n"
+  using Maclaurin2_objl
+    [where f = sin and n = n and h = x and diff = "\<lambda>n x. sin (x + 1/2 * real n * pi)"]
+  apply safe
     apply simp
    apply (simp (no_asm) add: sin_expansion_lemma del: of_nat_Suc)
    apply (force intro!: derivative_eq_intros)
   apply (erule ssubst)
-  apply (rule_tac x = t in exI, simp)
- apply (rule setsum.cong[OF refl])
- apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
-done
-
-lemma Maclaurin_sin_expansion4:
-     "0 < x ==>
-       \<exists>t. 0 < t & t \<le> x &
-       sin x =
-       (\<Sum>m<n. sin_coeff m * x ^ m)
-      + ((sin(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
-apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin2_objl)
-apply safe
-    apply simp
-   apply (simp (no_asm) add: sin_expansion_lemma del: of_nat_Suc)
-   apply (force intro!: derivative_eq_intros)
-  apply (erule ssubst)
-  apply (rule_tac x = t in exI, simp)
- apply (rule setsum.cong[OF refl])
- apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
-done
+  apply (rule_tac x = t in exI)
+  apply simp
+  apply (rule setsum.cong[OF refl])
+  apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
+  done
 
 
-subsection\<open>Maclaurin Expansion for Cosine Function\<close>
+subsection \<open>Maclaurin Expansion for Cosine Function\<close>
 
-lemma sumr_cos_zero_one [simp]:
-  "(\<Sum>m<(Suc n). cos_coeff m * 0 ^ m) = 1"
-by (induct "n", auto)
+lemma sumr_cos_zero_one [simp]: "(\<Sum>m<Suc n. cos_coeff m * 0 ^ m) = 1"
+  by (induct n) auto
 
-lemma cos_expansion_lemma:
-  "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)"
-by (simp only: cos_add sin_add of_nat_Suc distrib_right add_divide_distrib, auto)
+lemma cos_expansion_lemma: "cos (x + real (Suc m) * pi / 2) = - sin (x + real m * pi / 2)"
+  by (auto simp: cos_add sin_add distrib_right add_divide_distrib)
 
 lemma Maclaurin_cos_expansion:
-     "\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> &
-       cos x =
-       (\<Sum>m<n. cos_coeff m * x ^ m)
-      + ((cos(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
-apply (cut_tac f = cos and n = n and x = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_all_lt_objl)
-apply safe
-    apply (simp (no_asm))
-   apply (simp (no_asm) add: cos_expansion_lemma del: of_nat_Suc)
-  apply (case_tac "n", simp)
-  apply (simp del: setsum_lessThan_Suc)
-apply (rule ccontr, simp)
-apply (drule_tac x = x in spec, simp)
-apply (erule ssubst)
-apply (rule_tac x = t in exI, simp)
-apply (rule setsum.cong[OF refl])
-apply (auto simp add: cos_coeff_def cos_zero_iff elim: evenE)
-done
+  "\<exists>t::real. \<bar>t\<bar> \<le> \<bar>x\<bar> \<and>
+    cos x = (\<Sum>m<n. cos_coeff m * x ^ m) + (cos(t + 1/2 * real n * pi) / fact n) * x ^ n"
+  using Maclaurin_all_lt_objl
+    [where f = cos and n = n and x = x and diff = "\<lambda>n x. cos (x + 1/2 * real n * pi)"]
+  apply safe
+      apply simp
+     apply (simp add: cos_expansion_lemma del: of_nat_Suc)
+    apply (cases n)
+     apply simp
+    apply (simp del: setsum_lessThan_Suc)
+   apply (rule ccontr)
+   apply simp
+   apply (drule_tac x = x in spec)
+   apply simp
+  apply (erule ssubst)
+  apply (rule_tac x = t in exI)
+  apply simp
+  apply (rule setsum.cong[OF refl])
+  apply (auto simp add: cos_coeff_def cos_zero_iff elim: evenE)
+  done
 
 lemma Maclaurin_cos_expansion2:
-     "[| 0 < x; n > 0 |] ==>
-       \<exists>t. 0 < t & t < x &
-       cos x =
-       (\<Sum>m<n. cos_coeff m * x ^ m)
-      + ((cos(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
-apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_objl)
-apply safe
+  "0 < x \<Longrightarrow> n > 0 \<Longrightarrow>
+    \<exists>t. 0 < t \<and> t < x \<and>
+      cos x = (\<Sum>m<n. cos_coeff m * x ^ m) + (cos (t + 1/2 * real n * pi) / fact n) * x ^ n"
+  using Maclaurin_objl
+    [where f = cos and n = n and h = x and diff = "\<lambda>n x. cos (x + 1/2 * real n * pi)"]
+  apply safe
+    apply simp
+   apply (simp add: cos_expansion_lemma del: of_nat_Suc)
+  apply (erule ssubst)
+  apply (rule_tac x = t in exI)
   apply simp
-  apply (simp (no_asm) add: cos_expansion_lemma del: of_nat_Suc)
- apply (erule ssubst)
- apply (rule_tac x = t in exI, simp)
-apply (rule setsum.cong[OF refl])
-apply (auto simp add: cos_coeff_def cos_zero_iff elim: evenE)
-done
+  apply (rule setsum.cong[OF refl])
+  apply (auto simp: cos_coeff_def cos_zero_iff elim: evenE)
+  done
 
 lemma Maclaurin_minus_cos_expansion:
-     "[| x < 0; n > 0 |] ==>
-       \<exists>t. x < t & t < 0 &
-       cos x =
-       (\<Sum>m<n. cos_coeff m * x ^ m)
-      + ((cos(t + 1/2 * real (n) *pi) / (fact n)) * x ^ n)"
-apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_minus_objl)
-apply safe
+  "x < 0 \<Longrightarrow> n > 0 \<Longrightarrow>
+    \<exists>t. x < t \<and> t < 0 \<and>
+      cos x = (\<Sum>m<n. cos_coeff m * x ^ m) + ((cos (t + 1/2 * real n * pi) / fact n) * x ^ n)"
+  using Maclaurin_minus_objl
+    [where f = cos and n = n and h = x and diff = "\<lambda>n x. cos (x + 1/2 * real n *pi)"]
+  apply safe
+    apply simp
+   apply (simp add: cos_expansion_lemma del: of_nat_Suc)
+  apply (erule ssubst)
+  apply (rule_tac x = t in exI)
   apply simp
- apply (simp (no_asm) add: cos_expansion_lemma del: of_nat_Suc)
-apply (erule ssubst)
-apply (rule_tac x = t in exI, simp)
-apply (rule setsum.cong[OF refl])
-apply (auto simp add: cos_coeff_def cos_zero_iff elim: evenE)
-done
+  apply (rule setsum.cong[OF refl])
+  apply (auto simp: cos_coeff_def cos_zero_iff elim: evenE)
+  done
+
 
-(* ------------------------------------------------------------------------- *)
-(* Version for ln(1 +/- x). Where is it??                                    *)
-(* ------------------------------------------------------------------------- *)
+(* Version for ln(1 +/- x). Where is it?? *)
+
 
-lemma sin_bound_lemma:
-    "[|x = y; \<bar>u\<bar> \<le> (v::real) |] ==> \<bar>(x + u) - y\<bar> \<le> v"
-by auto
+lemma sin_bound_lemma: "x = y \<Longrightarrow> \<bar>u\<bar> \<le> v \<Longrightarrow> \<bar>(x + u) - y\<bar> \<le> v"
+  for x y u v :: real
+  by auto
 
-lemma Maclaurin_sin_bound:
-  "\<bar>sin x - (\<Sum>m<n. sin_coeff m * x ^ m)\<bar> \<le> inverse((fact n)) * \<bar>x\<bar> ^ n"
+lemma Maclaurin_sin_bound: "\<bar>sin x - (\<Sum>m<n. sin_coeff m * x ^ m)\<bar> \<le> inverse (fact n) * \<bar>x\<bar> ^ n"
 proof -
-  have "!! x (y::real). x \<le> 1 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x * y \<le> 1 * y"
-    by (rule_tac mult_right_mono,simp_all)
-  note est = this[simplified]
-  let ?diff = "\<lambda>(n::nat) x. if n mod 4 = 0 then sin(x) else if n mod 4 = 1 then cos(x) else if n mod 4 = 2 then -sin(x) else -cos(x)"
+  have est: "x \<le> 1 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x * y \<le> 1 * y" for x y :: real
+    by (rule mult_right_mono) simp_all
+  let ?diff = "\<lambda>(n::nat) x.
+    if n mod 4 = 0 then sin x
+    else if n mod 4 = 1 then cos x
+    else if n mod 4 = 2 then - sin x
+    else - cos x"
   have diff_0: "?diff 0 = sin" by simp
   have DERIV_diff: "\<forall>m x. DERIV (?diff m) x :> ?diff (Suc m) x"
-    apply (clarify)
+    apply clarify
     apply (subst (1 2 3) mod_Suc_eq_Suc_mod)
     apply (cut_tac m=m in mod_exhaust_less_4)
-    apply (safe, auto intro!: derivative_eq_intros)
+    apply safe
+       apply (auto intro!: derivative_eq_intros)
     done
   from Maclaurin_all_le [OF diff_0 DERIV_diff]
-  obtain t where t1: "\<bar>t\<bar> \<le> \<bar>x\<bar>" and
-    t2: "sin x = (\<Sum>m<n. ?diff m 0 / (fact m) * x ^ m) +
-      ?diff n t / (fact n) * x ^ n" by fast
-  have diff_m_0:
-    "\<And>m. ?diff m 0 = (if even m then 0
-         else (- 1) ^ ((m - Suc 0) div 2))"
+  obtain t where t1: "\<bar>t\<bar> \<le> \<bar>x\<bar>"
+    and t2: "sin x = (\<Sum>m<n. ?diff m 0 / (fact m) * x ^ m) + ?diff n t / (fact n) * x ^ n"
+    by fast
+  have diff_m_0: "\<And>m. ?diff m 0 = (if even m then 0 else (- 1) ^ ((m - Suc 0) div 2))"
     apply (subst even_even_mod_4_iff)
     apply (cut_tac m=m in mod_exhaust_less_4)
-    apply (elim disjE, simp_all)
-    apply (safe dest!: mod_eqD, simp_all)
+    apply (elim disjE)
+       apply simp_all
+     apply (safe dest!: mod_eqD)
+     apply simp_all
     done
   show ?thesis
     unfolding sin_coeff_def
     apply (subst t2)
     apply (rule sin_bound_lemma)
-    apply (rule setsum.cong[OF refl])
-    apply (subst diff_m_0, simp)
+     apply (rule setsum.cong[OF refl])
+     apply (subst diff_m_0, simp)
+    using est
     apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
-                simp add: est ac_simps divide_inverse power_abs [symmetric] abs_mult)
+        simp: ac_simps divide_inverse power_abs [symmetric] abs_mult)
     done
 qed