src/HOL/Transcendental.thy
changeset 29740 c56a5571f60a
parent 29695 171146a93106
child 30019 43c5b7bfc791
     1.1 --- a/src/HOL/Transcendental.thy	Wed Feb 04 18:10:07 2009 +0100
     1.2 +++ b/src/HOL/Transcendental.thy	Thu Feb 05 11:34:42 2009 +0100
     1.3 @@ -113,6 +113,208 @@
     1.4        ==> summable (%n. f(n) * (z ^ n))"
     1.5  by (rule powser_insidea [THEN summable_norm_cancel])
     1.6  
     1.7 +lemma sum_split_even_odd: fixes f :: "nat \<Rightarrow> real" shows
     1.8 +  "(\<Sum> i = 0 ..< 2 * n. if even i then f i else g i) = 
     1.9 +   (\<Sum> i = 0 ..< n. f (2 * i)) + (\<Sum> i = 0 ..< n. g (2 * i + 1))"
    1.10 +proof (induct n)
    1.11 +  case (Suc n)
    1.12 +  have "(\<Sum> i = 0 ..< 2 * Suc n. if even i then f i else g i) = 
    1.13 +        (\<Sum> i = 0 ..< n. f (2 * i)) + (\<Sum> i = 0 ..< n. g (2 * i + 1)) + (f (2 * n) + g (2 * n + 1))"
    1.14 +    using Suc.hyps by auto
    1.15 +  also have "\<dots> = (\<Sum> i = 0 ..< Suc n. f (2 * i)) + (\<Sum> i = 0 ..< Suc n. g (2 * i + 1))" by auto
    1.16 +  finally show ?case .
    1.17 +qed auto
    1.18 +
    1.19 +lemma sums_if': fixes g :: "nat \<Rightarrow> real" assumes "g sums x"
    1.20 +  shows "(\<lambda> n. if even n then 0 else g ((n - 1) div 2)) sums x"
    1.21 +  unfolding sums_def
    1.22 +proof (rule LIMSEQ_I)
    1.23 +  fix r :: real assume "0 < r"
    1.24 +  from `g sums x`[unfolded sums_def, THEN LIMSEQ_D, OF this]
    1.25 +  obtain no where no_eq: "\<And> n. n \<ge> no \<Longrightarrow> (norm (setsum g { 0..<n } - x) < r)" by blast
    1.26 +
    1.27 +  let ?SUM = "\<lambda> m. \<Sum> i = 0 ..< m. if even i then 0 else g ((i - 1) div 2)"
    1.28 +  { fix m assume "m \<ge> 2 * no" hence "m div 2 \<ge> no" by auto
    1.29 +    have sum_eq: "?SUM (2 * (m div 2)) = setsum g { 0 ..< m div 2 }" 
    1.30 +      using sum_split_even_odd by auto
    1.31 +    hence "(norm (?SUM (2 * (m div 2)) - x) < r)" using no_eq unfolding sum_eq using `m div 2 \<ge> no` by auto
    1.32 +    moreover
    1.33 +    have "?SUM (2 * (m div 2)) = ?SUM m"
    1.34 +    proof (cases "even m")
    1.35 +      case True show ?thesis unfolding even_nat_div_two_times_two[OF True, unfolded numeral_2_eq_2[symmetric]] ..
    1.36 +    next
    1.37 +      case False hence "even (Suc m)" by auto
    1.38 +      from even_nat_div_two_times_two[OF this, unfolded numeral_2_eq_2[symmetric]] odd_nat_plus_one_div_two[OF False, unfolded numeral_2_eq_2[symmetric]]
    1.39 +      have eq: "Suc (2 * (m div 2)) = m" by auto
    1.40 +      hence "even (2 * (m div 2))" using `odd m` by auto
    1.41 +      have "?SUM m = ?SUM (Suc (2 * (m div 2)))" unfolding eq ..
    1.42 +      also have "\<dots> = ?SUM (2 * (m div 2))" using `even (2 * (m div 2))` by auto
    1.43 +      finally show ?thesis by auto
    1.44 +    qed
    1.45 +    ultimately have "(norm (?SUM m - x) < r)" by auto
    1.46 +  }
    1.47 +  thus "\<exists> no. \<forall> m \<ge> no. norm (?SUM m - x) < r" by blast
    1.48 +qed
    1.49 +
    1.50 +lemma sums_if: fixes g :: "nat \<Rightarrow> real" assumes "g sums x" and "f sums y"
    1.51 +  shows "(\<lambda> n. if even n then f (n div 2) else g ((n - 1) div 2)) sums (x + y)"
    1.52 +proof -
    1.53 +  let ?s = "\<lambda> n. if even n then 0 else f ((n - 1) div 2)"
    1.54 +  { fix B T E have "(if B then (0 :: real) else E) + (if B then T else 0) = (if B then T else E)"
    1.55 +      by (cases B) auto } note if_sum = this
    1.56 +  have g_sums: "(\<lambda> n. if even n then 0 else g ((n - 1) div 2)) sums x" using sums_if'[OF `g sums x`] .
    1.57 +  { 
    1.58 +    have "?s 0 = 0" by auto
    1.59 +    have Suc_m1: "\<And> n. Suc n - 1 = n" by auto
    1.60 +    { fix B T E have "(if \<not> B then T else E) = (if B then E else T)" by auto } note if_eq = this
    1.61 +
    1.62 +    have "?s sums y" using sums_if'[OF `f sums y`] .
    1.63 +    from this[unfolded sums_def, THEN LIMSEQ_Suc] 
    1.64 +    have "(\<lambda> n. if even n then f (n div 2) else 0) sums y"
    1.65 +      unfolding sums_def setsum_shift_lb_Suc0_0_upt[where f="?s", OF `?s 0 = 0`, symmetric]
    1.66 +                image_Suc_atLeastLessThan[symmetric] setsum_reindex[OF inj_Suc, unfolded comp_def]
    1.67 +                even_nat_Suc Suc_m1 if_eq .
    1.68 +  } from sums_add[OF g_sums this]
    1.69 +  show ?thesis unfolding if_sum .
    1.70 +qed
    1.71 +
    1.72 +subsection {* Alternating series test / Leibniz formula *}
    1.73 +
    1.74 +lemma sums_alternating_upper_lower:
    1.75 +  fixes a :: "nat \<Rightarrow> real"
    1.76 +  assumes mono: "\<And>n. a (Suc n) \<le> a n" and a_pos: "\<And>n. 0 \<le> a n" and "a ----> 0"
    1.77 +  shows "\<exists>l. ((\<forall>n. (\<Sum>i=0..<2*n. -1^i*a i) \<le> l) \<and> (\<lambda> n. \<Sum>i=0..<2*n. -1^i*a i) ----> l) \<and> 
    1.78 +             ((\<forall>n. l \<le> (\<Sum>i=0..<2*n + 1. -1^i*a i)) \<and> (\<lambda> n. \<Sum>i=0..<2*n + 1. -1^i*a i) ----> l)"
    1.79 +  (is "\<exists>l. ((\<forall>n. ?f n \<le> l) \<and> _) \<and> ((\<forall>n. l \<le> ?g n) \<and> _)")
    1.80 +proof -
    1.81 +  have fg_diff: "\<And>n. ?f n - ?g n = - a (2 * n)" by auto
    1.82 +
    1.83 +  have "\<forall> n. ?f n \<le> ?f (Suc n)"
    1.84 +  proof fix n show "?f n \<le> ?f (Suc n)" using mono[of "2*n"] by auto qed
    1.85 +  moreover
    1.86 +  have "\<forall> n. ?g (Suc n) \<le> ?g n"
    1.87 +  proof fix n show "?g (Suc n) \<le> ?g n" using mono[of "Suc (2*n)"] by auto qed
    1.88 +  moreover
    1.89 +  have "\<forall> n. ?f n \<le> ?g n" 
    1.90 +  proof fix n show "?f n \<le> ?g n" using fg_diff a_pos by auto qed
    1.91 +  moreover
    1.92 +  have "(\<lambda> n. ?f n - ?g n) ----> 0" unfolding fg_diff
    1.93 +  proof (rule LIMSEQ_I)
    1.94 +    fix r :: real assume "0 < r"
    1.95 +    with `a ----> 0`[THEN LIMSEQ_D] 
    1.96 +    obtain N where "\<And> n. n \<ge> N \<Longrightarrow> norm (a n - 0) < r" by auto
    1.97 +    hence "\<forall> n \<ge> N. norm (- a (2 * n) - 0) < r" by auto
    1.98 +    thus "\<exists> N. \<forall> n \<ge> N. norm (- a (2 * n) - 0) < r" by auto
    1.99 +  qed
   1.100 +  ultimately
   1.101 +  show ?thesis by (rule lemma_nest_unique)
   1.102 +qed 
   1.103 +
   1.104 +lemma summable_Leibniz': fixes a :: "nat \<Rightarrow> real"
   1.105 +  assumes a_zero: "a ----> 0" and a_pos: "\<And> n. 0 \<le> a n"
   1.106 +  and a_monotone: "\<And> n. a (Suc n) \<le> a n"
   1.107 +  shows summable: "summable (\<lambda> n. (-1)^n * a n)"
   1.108 +  and "\<And>n. (\<Sum>i=0..<2*n. (-1)^i*a i) \<le> (\<Sum>i. (-1)^i*a i)"
   1.109 +  and "(\<lambda>n. \<Sum>i=0..<2*n. (-1)^i*a i) ----> (\<Sum>i. (-1)^i*a i)"
   1.110 +  and "\<And>n. (\<Sum>i. (-1)^i*a i) \<le> (\<Sum>i=0..<2*n+1. (-1)^i*a i)"
   1.111 +  and "(\<lambda>n. \<Sum>i=0..<2*n+1. (-1)^i*a i) ----> (\<Sum>i. (-1)^i*a i)"
   1.112 +proof -
   1.113 +  let "?S n" = "(-1)^n * a n"
   1.114 +  let "?P n" = "\<Sum>i=0..<n. ?S i"
   1.115 +  let "?f n" = "?P (2 * n)"
   1.116 +  let "?g n" = "?P (2 * n + 1)"
   1.117 +  obtain l :: real where below_l: "\<forall> n. ?f n \<le> l" and "?f ----> l" and above_l: "\<forall> n. l \<le> ?g n" and "?g ----> l"
   1.118 +    using sums_alternating_upper_lower[OF a_monotone a_pos a_zero] by blast
   1.119 +  
   1.120 +  let ?Sa = "\<lambda> m. \<Sum> n = 0..<m. ?S n"
   1.121 +  have "?Sa ----> l"
   1.122 +  proof (rule LIMSEQ_I)
   1.123 +    fix r :: real assume "0 < r"
   1.124 +
   1.125 +    with `?f ----> l`[THEN LIMSEQ_D] 
   1.126 +    obtain f_no where f: "\<And> n. n \<ge> f_no \<Longrightarrow> norm (?f n - l) < r" by auto
   1.127 +
   1.128 +    from `0 < r` `?g ----> l`[THEN LIMSEQ_D] 
   1.129 +    obtain g_no where g: "\<And> n. n \<ge> g_no \<Longrightarrow> norm (?g n - l) < r" by auto
   1.130 +
   1.131 +    { fix n :: nat
   1.132 +      assume "n \<ge> (max (2 * f_no) (2 * g_no))" hence "n \<ge> 2 * f_no" and "n \<ge> 2 * g_no" by auto
   1.133 +      have "norm (?Sa n - l) < r"
   1.134 +      proof (cases "even n")
   1.135 +	case True from even_nat_div_two_times_two[OF this]
   1.136 +	have n_eq: "2 * (n div 2) = n" unfolding numeral_2_eq_2[symmetric] by auto
   1.137 +	with `n \<ge> 2 * f_no` have "n div 2 \<ge> f_no" by auto
   1.138 +	from f[OF this]
   1.139 +	show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost .
   1.140 +      next
   1.141 +	case False hence "even (n - 1)" using even_num_iff odd_pos by auto 
   1.142 +	from even_nat_div_two_times_two[OF this]
   1.143 +	have n_eq: "2 * ((n - 1) div 2) = n - 1" unfolding numeral_2_eq_2[symmetric] by auto
   1.144 +	hence range_eq: "n - 1 + 1 = n" using odd_pos[OF False] by auto
   1.145 +
   1.146 +	from n_eq `n \<ge> 2 * g_no` have "(n - 1) div 2 \<ge> g_no" by auto
   1.147 +	from g[OF this]
   1.148 +	show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost range_eq .
   1.149 +      qed
   1.150 +    }
   1.151 +    thus "\<exists> no. \<forall> n \<ge> no. norm (?Sa n - l) < r" by blast
   1.152 +  qed
   1.153 +  hence sums_l: "(\<lambda>i. (-1)^i * a i) sums l" unfolding sums_def atLeastLessThanSuc_atLeastAtMost[symmetric] .
   1.154 +  thus "summable ?S" using summable_def by auto
   1.155 +
   1.156 +  have "l = suminf ?S" using sums_unique[OF sums_l] .
   1.157 +
   1.158 +  { fix n show "suminf ?S \<le> ?g n" unfolding sums_unique[OF sums_l, symmetric] using above_l by auto }
   1.159 +  { fix n show "?f n \<le> suminf ?S" unfolding sums_unique[OF sums_l, symmetric] using below_l by auto }
   1.160 +  show "?g ----> suminf ?S" using `?g ----> l` `l = suminf ?S` by auto
   1.161 +  show "?f ----> suminf ?S" using `?f ----> l` `l = suminf ?S` by auto
   1.162 +qed
   1.163 +
   1.164 +theorem summable_Leibniz: fixes a :: "nat \<Rightarrow> real"
   1.165 +  assumes a_zero: "a ----> 0" and "monoseq a"
   1.166 +  shows "summable (\<lambda> n. (-1)^n * a n)" (is "?summable")
   1.167 +  and "0 < a 0 \<longrightarrow> (\<forall>n. (\<Sum>i. -1^i*a i) \<in> { \<Sum>i=0..<2*n. -1^i * a i .. \<Sum>i=0..<2*n+1. -1^i * a i})" (is "?pos")
   1.168 +  and "a 0 < 0 \<longrightarrow> (\<forall>n. (\<Sum>i. -1^i*a i) \<in> { \<Sum>i=0..<2*n+1. -1^i * a i .. \<Sum>i=0..<2*n. -1^i * a i})" (is "?neg")
   1.169 +  and "(\<lambda>n. \<Sum>i=0..<2*n. -1^i*a i) ----> (\<Sum>i. -1^i*a i)" (is "?f")
   1.170 +  and "(\<lambda>n. \<Sum>i=0..<2*n+1. -1^i*a i) ----> (\<Sum>i. -1^i*a i)" (is "?g")
   1.171 +proof -
   1.172 +  have "?summable \<and> ?pos \<and> ?neg \<and> ?f \<and> ?g"
   1.173 +  proof (cases "(\<forall> n. 0 \<le> a n) \<and> (\<forall>m. \<forall>n\<ge>m. a n \<le> a m)")
   1.174 +    case True
   1.175 +    hence ord: "\<And>n m. m \<le> n \<Longrightarrow> a n \<le> a m" and ge0: "\<And> n. 0 \<le> a n" by auto
   1.176 +    { fix n have "a (Suc n) \<le> a n" using ord[where n="Suc n" and m=n] by auto }
   1.177 +    note leibniz = summable_Leibniz'[OF `a ----> 0` ge0] and mono = this
   1.178 +    from leibniz[OF mono]
   1.179 +    show ?thesis using `0 \<le> a 0` by auto
   1.180 +  next
   1.181 +    let ?a = "\<lambda> n. - a n"
   1.182 +    case False
   1.183 +    with monoseq_le[OF `monoseq a` `a ----> 0`]
   1.184 +    have "(\<forall> n. a n \<le> 0) \<and> (\<forall>m. \<forall>n\<ge>m. a m \<le> a n)" by auto
   1.185 +    hence ord: "\<And>n m. m \<le> n \<Longrightarrow> ?a n \<le> ?a m" and ge0: "\<And> n. 0 \<le> ?a n" by auto
   1.186 +    { fix n have "?a (Suc n) \<le> ?a n" using ord[where n="Suc n" and m=n] by auto }
   1.187 +    note monotone = this
   1.188 +    note leibniz = summable_Leibniz'[OF _ ge0, of "\<lambda>x. x", OF LIMSEQ_minus[OF `a ----> 0`, unfolded minus_zero] monotone]
   1.189 +    have "summable (\<lambda> n. (-1)^n * ?a n)" using leibniz(1) by auto
   1.190 +    then obtain l where "(\<lambda> n. (-1)^n * ?a n) sums l" unfolding summable_def by auto
   1.191 +    from this[THEN sums_minus]
   1.192 +    have "(\<lambda> n. (-1)^n * a n) sums -l" by auto
   1.193 +    hence ?summable unfolding summable_def by auto
   1.194 +    moreover
   1.195 +    have "\<And> a b :: real. \<bar> - a - - b \<bar> = \<bar>a - b\<bar>" unfolding minus_diff_minus by auto
   1.196 +    
   1.197 +    from suminf_minus[OF leibniz(1), unfolded mult_minus_right minus_minus]
   1.198 +    have move_minus: "(\<Sum>n. - (-1 ^ n * a n)) = - (\<Sum>n. -1 ^ n * a n)" by auto
   1.199 +
   1.200 +    have ?pos using `0 \<le> ?a 0` by auto
   1.201 +    moreover have ?neg using leibniz(2,4) unfolding mult_minus_right setsum_negf move_minus neg_le_iff_le by auto
   1.202 +    moreover have ?f and ?g using leibniz(3,5)[unfolded mult_minus_right setsum_negf move_minus, THEN LIMSEQ_minus_cancel] by auto
   1.203 +    ultimately show ?thesis by auto
   1.204 +  qed
   1.205 +  from this[THEN conjunct1] this[THEN conjunct2, THEN conjunct1] this[THEN conjunct2, THEN conjunct2, THEN conjunct1] this[THEN conjunct2, THEN conjunct2, THEN conjunct2, THEN conjunct1]
   1.206 +       this[THEN conjunct2, THEN conjunct2, THEN conjunct2, THEN conjunct2]
   1.207 +  show ?summable and ?pos and ?neg and ?f and ?g .
   1.208 +qed
   1.209  
   1.210  subsection {* Term-by-Term Differentiability of Power Series *}
   1.211  
   1.212 @@ -432,6 +634,188 @@
   1.213  lemma inv_real_of_nat_fact_ge_zero [simp]: "0 \<le> inverse (real (fact n))"
   1.214  by (auto intro: order_less_imp_le)
   1.215  
   1.216 +subsection {* Derivability of power series *}
   1.217 +
   1.218 +lemma DERIV_series': fixes f :: "real \<Rightarrow> nat \<Rightarrow> real"
   1.219 +  assumes DERIV_f: "\<And> n. DERIV (\<lambda> x. f x n) x0 :> (f' x0 n)"
   1.220 +  and allf_summable: "\<And> x. x \<in> {a <..< b} \<Longrightarrow> summable (f x)" and x0_in_I: "x0 \<in> {a <..< b}"
   1.221 +  and "summable (f' x0)"
   1.222 +  and "summable L" and L_def: "\<And> n x y. \<lbrakk> x \<in> { a <..< b} ; y \<in> { a <..< b} \<rbrakk> \<Longrightarrow> \<bar> f x n - f y n \<bar> \<le> L n * \<bar> x - y \<bar>"
   1.223 +  shows "DERIV (\<lambda> x. suminf (f x)) x0 :> (suminf (f' x0))"
   1.224 +  unfolding deriv_def
   1.225 +proof (rule LIM_I)
   1.226 +  fix r :: real assume "0 < r" hence "0 < r/3" by auto
   1.227 +
   1.228 +  obtain N_L where N_L: "\<And> n. N_L \<le> n \<Longrightarrow> \<bar> \<Sum> i. L (i + n) \<bar> < r/3" 
   1.229 +    using suminf_exist_split[OF `0 < r/3` `summable L`] by auto
   1.230 +
   1.231 +  obtain N_f' where N_f': "\<And> n. N_f' \<le> n \<Longrightarrow> \<bar> \<Sum> i. f' x0 (i + n) \<bar> < r/3" 
   1.232 +    using suminf_exist_split[OF `0 < r/3` `summable (f' x0)`] by auto
   1.233 +
   1.234 +  let ?N = "Suc (max N_L N_f')"
   1.235 +  have "\<bar> \<Sum> i. f' x0 (i + ?N) \<bar> < r/3" (is "?f'_part < r/3") and
   1.236 +    L_estimate: "\<bar> \<Sum> i. L (i + ?N) \<bar> < r/3" using N_L[of "?N"] and N_f' [of "?N"] by auto
   1.237 +
   1.238 +  let "?diff i x" = "(f (x0 + x) i - f x0 i) / x"
   1.239 +
   1.240 +  let ?r = "r / (3 * real ?N)"
   1.241 +  have "0 < 3 * real ?N" by auto
   1.242 +  from divide_pos_pos[OF `0 < r` this]
   1.243 +  have "0 < ?r" .
   1.244 +
   1.245 +  let "?s n" = "SOME s. 0 < s \<and> (\<forall> x. x \<noteq> 0 \<and> \<bar> x \<bar> < s \<longrightarrow> \<bar> ?diff n x - f' x0 n \<bar> < ?r)"
   1.246 +  def S' \<equiv> "Min (?s ` { 0 ..< ?N })"
   1.247 +
   1.248 +  have "0 < S'" unfolding S'_def
   1.249 +  proof (rule iffD2[OF Min_gr_iff])
   1.250 +    show "\<forall> x \<in> (?s ` { 0 ..< ?N }). 0 < x"
   1.251 +    proof (rule ballI)
   1.252 +      fix x assume "x \<in> ?s ` {0..<?N}"
   1.253 +      then obtain n where "x = ?s n" and "n \<in> {0..<?N}" using image_iff[THEN iffD1] by blast
   1.254 +      from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def] 
   1.255 +      obtain s where s_bound: "0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> \<bar>x\<bar> < s \<longrightarrow> \<bar>?diff n x - f' x0 n\<bar> < ?r)" by auto
   1.256 +      have "0 < ?s n" by (rule someI2[where a=s], auto simp add: s_bound)
   1.257 +      thus "0 < x" unfolding `x = ?s n` .
   1.258 +    qed
   1.259 +  qed auto
   1.260 +
   1.261 +  def S \<equiv> "min (min (x0 - a) (b - x0)) S'"
   1.262 +  hence "0 < S" and S_a: "S \<le> x0 - a" and S_b: "S \<le> b - x0" and "S \<le> S'" using x0_in_I and `0 < S'`
   1.263 +    by auto
   1.264 +
   1.265 +  { fix x assume "x \<noteq> 0" and "\<bar> x \<bar> < S"
   1.266 +    hence x_in_I: "x0 + x \<in> { a <..< b }" using S_a S_b by auto
   1.267 +    
   1.268 +    note diff_smbl = summable_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
   1.269 +    note div_smbl = summable_divide[OF diff_smbl]
   1.270 +    note all_smbl = summable_diff[OF div_smbl `summable (f' x0)`]
   1.271 +    note ign = summable_ignore_initial_segment[where k="?N"]
   1.272 +    note diff_shft_smbl = summable_diff[OF ign[OF allf_summable[OF x_in_I]] ign[OF allf_summable[OF x0_in_I]]]
   1.273 +    note div_shft_smbl = summable_divide[OF diff_shft_smbl]
   1.274 +    note all_shft_smbl = summable_diff[OF div_smbl ign[OF `summable (f' x0)`]]
   1.275 +
   1.276 +    { fix n
   1.277 +      have "\<bar> ?diff (n + ?N) x \<bar> \<le> L (n + ?N) * \<bar> (x0 + x) - x0 \<bar> / \<bar> x \<bar>" 
   1.278 +	using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] unfolding abs_divide .
   1.279 +      hence "\<bar> ( \<bar> ?diff (n + ?N) x \<bar>) \<bar> \<le> L (n + ?N)" using `x \<noteq> 0` by auto
   1.280 +    } note L_ge = summable_le2[OF allI[OF this] ign[OF `summable L`]]
   1.281 +    from order_trans[OF summable_rabs[OF conjunct1[OF L_ge]] L_ge[THEN conjunct2]]
   1.282 +    have "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> (\<Sum> i. L (i + ?N))" .
   1.283 +    hence "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> r / 3" (is "?L_part \<le> r/3") using L_estimate by auto
   1.284 +
   1.285 +    have "\<bar>\<Sum>n \<in> { 0 ..< ?N}. ?diff n x - f' x0 n \<bar> \<le> (\<Sum>n \<in> { 0 ..< ?N}. \<bar>?diff n x - f' x0 n \<bar>)" ..
   1.286 +    also have "\<dots> < (\<Sum>n \<in> { 0 ..< ?N}. ?r)"
   1.287 +    proof (rule setsum_strict_mono)
   1.288 +      fix n assume "n \<in> { 0 ..< ?N}"
   1.289 +      have "\<bar> x \<bar> < S" using `\<bar> x \<bar> < S` .
   1.290 +      also have "S \<le> S'" using `S \<le> S'` .
   1.291 +      also have "S' \<le> ?s n" unfolding S'_def 
   1.292 +      proof (rule Min_le_iff[THEN iffD2])
   1.293 +	have "?s n \<in> (?s ` {0..<?N}) \<and> ?s n \<le> ?s n" using `n \<in> { 0 ..< ?N}` by auto
   1.294 +	thus "\<exists> a \<in> (?s ` {0..<?N}). a \<le> ?s n" by blast
   1.295 +      qed auto
   1.296 +      finally have "\<bar> x \<bar> < ?s n" .
   1.297 +
   1.298 +      from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def diff_0_right, unfolded some_eq_ex[symmetric], THEN conjunct2]
   1.299 +      have "\<forall>x. x \<noteq> 0 \<and> \<bar>x\<bar> < ?s n \<longrightarrow> \<bar>?diff n x - f' x0 n\<bar> < ?r" .
   1.300 +      with `x \<noteq> 0` and `\<bar>x\<bar> < ?s n`
   1.301 +      show "\<bar>?diff n x - f' x0 n\<bar> < ?r" by blast
   1.302 +    qed auto
   1.303 +    also have "\<dots> = of_nat (card {0 ..< ?N}) * ?r" by (rule setsum_constant)
   1.304 +    also have "\<dots> = real ?N * ?r" unfolding real_eq_of_nat by auto
   1.305 +    also have "\<dots> = r/3" by auto
   1.306 +    finally have "\<bar>\<Sum>n \<in> { 0 ..< ?N}. ?diff n x - f' x0 n \<bar> < r / 3" (is "?diff_part < r / 3") .
   1.307 +
   1.308 +    from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
   1.309 +    have "\<bar> (suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0) \<bar> = 
   1.310 +                    \<bar> \<Sum>n. ?diff n x - f' x0 n \<bar>" unfolding suminf_diff[OF div_smbl `summable (f' x0)`, symmetric] using suminf_divide[OF diff_smbl, symmetric] by auto
   1.311 +    also have "\<dots> \<le> ?diff_part + \<bar> (\<Sum>n. ?diff (n + ?N) x) - (\<Sum> n. f' x0 (n + ?N)) \<bar>" unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"] unfolding suminf_diff[OF div_shft_smbl ign[OF `summable (f' x0)`]] by (rule abs_triangle_ineq)
   1.312 +    also have "\<dots> \<le> ?diff_part + ?L_part + ?f'_part" using abs_triangle_ineq4 by auto
   1.313 +    also have "\<dots> < r /3 + r/3 + r/3" 
   1.314 +      using `?diff_part < r/3` `?L_part \<le> r/3` and `?f'_part < r/3` by auto
   1.315 +    finally have "\<bar> (suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0) \<bar> < r"
   1.316 +      by auto
   1.317 +  } thus "\<exists> s > 0. \<forall> x. x \<noteq> 0 \<and> norm (x - 0) < s \<longrightarrow> 
   1.318 +      norm (((\<Sum>n. f (x0 + x) n) - (\<Sum>n. f x0 n)) / x - (\<Sum>n. f' x0 n)) < r" using `0 < S`
   1.319 +    unfolding real_norm_def diff_0_right by blast
   1.320 +qed
   1.321 +
   1.322 +lemma DERIV_power_series': fixes f :: "nat \<Rightarrow> real"
   1.323 +  assumes converges: "\<And> x. x \<in> {-R <..< R} \<Longrightarrow> summable (\<lambda> n. f n * real (Suc n) * x^n)"
   1.324 +  and x0_in_I: "x0 \<in> {-R <..< R}" and "0 < R"
   1.325 +  shows "DERIV (\<lambda> x. (\<Sum> n. f n * x^(Suc n))) x0 :> (\<Sum> n. f n * real (Suc n) * x0^n)"
   1.326 +  (is "DERIV (\<lambda> x. (suminf (?f x))) x0 :> (suminf (?f' x0))")
   1.327 +proof -
   1.328 +  { fix R' assume "0 < R'" and "R' < R" and "-R' < x0" and "x0 < R'"
   1.329 +    hence "x0 \<in> {-R' <..< R'}" and "R' \<in> {-R <..< R}" and "x0 \<in> {-R <..< R}" by auto
   1.330 +    have "DERIV (\<lambda> x. (suminf (?f x))) x0 :> (suminf (?f' x0))"
   1.331 +    proof (rule DERIV_series')
   1.332 +      show "summable (\<lambda> n. \<bar>f n * real (Suc n) * R'^n\<bar>)"
   1.333 +      proof -
   1.334 +	have "(R' + R) / 2 < R" and "0 < (R' + R) / 2" using `0 < R'` `0 < R` `R' < R` by auto
   1.335 +	hence in_Rball: "(R' + R) / 2 \<in> {-R <..< R}" using `R' < R` by auto
   1.336 +	have "norm R' < norm ((R' + R) / 2)" using `0 < R'` `0 < R` `R' < R` by auto
   1.337 +	from powser_insidea[OF converges[OF in_Rball] this] show ?thesis by auto
   1.338 +      qed
   1.339 +      { fix n x y assume "x \<in> {-R' <..< R'}" and "y \<in> {-R' <..< R'}"
   1.340 +	show "\<bar>?f x n - ?f y n\<bar> \<le> \<bar>f n * real (Suc n) * R'^n\<bar> * \<bar>x-y\<bar>"
   1.341 +	proof -
   1.342 +	  have "\<bar>f n * x ^ (Suc n) - f n * y ^ (Suc n)\<bar> = (\<bar>f n\<bar> * \<bar>x-y\<bar>) * \<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar>" 
   1.343 +	    unfolding right_diff_distrib[symmetric] lemma_realpow_diff_sumr2 abs_mult by auto
   1.344 +	  also have "\<dots> \<le> (\<bar>f n\<bar> * \<bar>x-y\<bar>) * (\<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>)" 
   1.345 +	  proof (rule mult_left_mono)
   1.346 +	    have "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> (\<Sum>p = 0..<Suc n. \<bar>x ^ p * y ^ (n - p)\<bar>)" by (rule setsum_abs)
   1.347 +	    also have "\<dots> \<le> (\<Sum>p = 0..<Suc n. R' ^ n)"
   1.348 +	    proof (rule setsum_mono)
   1.349 +	      fix p assume "p \<in> {0..<Suc n}" hence "p \<le> n" by auto
   1.350 +	      { fix n fix x :: real assume "x \<in> {-R'<..<R'}"
   1.351 +		hence "\<bar>x\<bar> \<le> R'"  by auto
   1.352 +		hence "\<bar>x^n\<bar> \<le> R'^n" unfolding power_abs by (rule power_mono, auto)
   1.353 +	      } from mult_mono[OF this[OF `x \<in> {-R'<..<R'}`, of p] this[OF `y \<in> {-R'<..<R'}`, of "n-p"]] `0 < R'`
   1.354 +	      have "\<bar>x^p * y^(n-p)\<bar> \<le> R'^p * R'^(n-p)" unfolding abs_mult by auto
   1.355 +	      thus "\<bar>x^p * y^(n-p)\<bar> \<le> R'^n" unfolding power_add[symmetric] using `p \<le> n` by auto
   1.356 +	    qed
   1.357 +	    also have "\<dots> = real (Suc n) * R' ^ n" unfolding setsum_constant card_atLeastLessThan real_of_nat_def by auto
   1.358 +	    finally show "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> \<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>" unfolding abs_real_of_nat_cancel abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF `0 < R'`]]] .
   1.359 +	    show "0 \<le> \<bar>f n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult[symmetric] by auto
   1.360 +	  qed
   1.361 +	  also have "\<dots> = \<bar>f n * real (Suc n) * R' ^ n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult real_mult_assoc[symmetric] by algebra
   1.362 +	  finally show ?thesis .
   1.363 +	qed }
   1.364 +      { fix n
   1.365 +	from DERIV_pow[of "Suc n" x0, THEN DERIV_cmult[where c="f n"]]
   1.366 +	show "DERIV (\<lambda> x. ?f x n) x0 :> (?f' x0 n)" unfolding real_mult_assoc by auto }
   1.367 +      { fix x assume "x \<in> {-R' <..< R'}" hence "R' \<in> {-R <..< R}" and "norm x < norm R'" using assms `R' < R` by auto
   1.368 +	have "summable (\<lambda> n. f n * x^n)"
   1.369 +	proof (rule summable_le2[THEN conjunct1, OF _ powser_insidea[OF converges[OF `R' \<in> {-R <..< R}`] `norm x < norm R'`]], rule allI)
   1.370 +	  fix n
   1.371 +	  have le: "\<bar>f n\<bar> * 1 \<le> \<bar>f n\<bar> * real (Suc n)" by (rule mult_left_mono, auto)
   1.372 +	  show "\<bar>f n * x ^ n\<bar> \<le> norm (f n * real (Suc n) * x ^ n)" unfolding real_norm_def abs_mult
   1.373 +	    by (rule mult_right_mono, auto simp add: le[unfolded mult_1_right])
   1.374 +	qed
   1.375 +	from this[THEN summable_mult2[where c=x], unfolded real_mult_assoc, unfolded real_mult_commute]
   1.376 +	show "summable (?f x)" by auto }
   1.377 +      show "summable (?f' x0)" using converges[OF `x0 \<in> {-R <..< R}`] .
   1.378 +      show "x0 \<in> {-R' <..< R'}" using `x0 \<in> {-R' <..< R'}` .
   1.379 +    qed
   1.380 +  } note for_subinterval = this
   1.381 +  let ?R = "(R + \<bar>x0\<bar>) / 2"
   1.382 +  have "\<bar>x0\<bar> < ?R" using assms by auto
   1.383 +  hence "- ?R < x0"
   1.384 +  proof (cases "x0 < 0")
   1.385 +    case True
   1.386 +    hence "- x0 < ?R" using `\<bar>x0\<bar> < ?R` by auto
   1.387 +    thus ?thesis unfolding neg_less_iff_less[symmetric, of "- x0"] by auto
   1.388 +  next
   1.389 +    case False
   1.390 +    have "- ?R < 0" using assms by auto
   1.391 +    also have "\<dots> \<le> x0" using False by auto 
   1.392 +    finally show ?thesis .
   1.393 +  qed
   1.394 +  hence "0 < ?R" "?R < R" "- ?R < x0" and "x0 < ?R" using assms by auto
   1.395 +  from for_subinterval[OF this]
   1.396 +  show ?thesis .
   1.397 +qed
   1.398  
   1.399  subsection {* Exponential Function *}
   1.400  
   1.401 @@ -830,6 +1214,37 @@
   1.402  apply (simp_all add: abs_if isCont_ln)
   1.403  done
   1.404  
   1.405 +lemma ln_series: assumes "0 < x" and "x < 2"
   1.406 +  shows "ln x = (\<Sum> n. (-1)^n * (1 / real (n + 1)) * (x - 1)^(Suc n))" (is "ln x = suminf (?f (x - 1))")
   1.407 +proof -
   1.408 +  let "?f' x n" = "(-1)^n * (x - 1)^n"
   1.409 +
   1.410 +  have "ln x - suminf (?f (x - 1)) = ln 1 - suminf (?f (1 - 1))"
   1.411 +  proof (rule DERIV_isconst3[where x=x])
   1.412 +    fix x :: real assume "x \<in> {0 <..< 2}" hence "0 < x" and "x < 2" by auto
   1.413 +    have "norm (1 - x) < 1" using `0 < x` and `x < 2` by auto
   1.414 +    have "1 / x = 1 / (1 - (1 - x))" by auto
   1.415 +    also have "\<dots> = (\<Sum> n. (1 - x)^n)" using geometric_sums[OF `norm (1 - x) < 1`] by (rule sums_unique)
   1.416 +    also have "\<dots> = suminf (?f' x)" unfolding power_mult_distrib[symmetric] by (rule arg_cong[where f=suminf], rule arg_cong[where f="op ^"], auto)
   1.417 +    finally have "DERIV ln x :> suminf (?f' x)" using DERIV_ln[OF `0 < x`] unfolding real_divide_def by auto
   1.418 +    moreover
   1.419 +    have repos: "\<And> h x :: real. h - 1 + x = h + x - 1" by auto
   1.420 +    have "DERIV (\<lambda>x. suminf (?f x)) (x - 1) :> (\<Sum>n. (-1)^n * (1 / real (n + 1)) * real (Suc n) * (x - 1) ^ n)"
   1.421 +    proof (rule DERIV_power_series')
   1.422 +      show "x - 1 \<in> {- 1<..<1}" and "(0 :: real) < 1" using `0 < x` `x < 2` by auto
   1.423 +      { fix x :: real assume "x \<in> {- 1<..<1}" hence "norm (-x) < 1" by auto
   1.424 +	show "summable (\<lambda>n. -1 ^ n * (1 / real (n + 1)) * real (Suc n) * x ^ n)"
   1.425 +	  by (auto simp del: power_mult_distrib simp add: power_mult_distrib[symmetric] summable_geometric[OF `norm (-x) < 1`])
   1.426 +      }
   1.427 +    qed
   1.428 +    hence "DERIV (\<lambda>x. suminf (?f x)) (x - 1) :> suminf (?f' x)" by auto
   1.429 +    hence "DERIV (\<lambda>x. suminf (?f (x - 1))) x :> suminf (?f' x)" unfolding DERIV_iff repos .
   1.430 +    ultimately have "DERIV (\<lambda>x. ln x - suminf (?f (x - 1))) x :> (suminf (?f' x) - suminf (?f' x))"
   1.431 +      by (rule DERIV_diff)
   1.432 +    thus "DERIV (\<lambda>x. ln x - suminf (?f (x - 1))) x :> 0" by auto
   1.433 +  qed (auto simp add: assms)
   1.434 +  thus ?thesis by (auto simp add: suminf_zero)
   1.435 +qed
   1.436  
   1.437  subsection {* Sine and Cosine *}
   1.438  
   1.439 @@ -1378,6 +1793,12 @@
   1.440  lemma minus_pi_half_less_zero: "-(pi/2) < 0"
   1.441  by simp
   1.442  
   1.443 +lemma m2pi_less_pi: "- (2 * pi) < pi"
   1.444 +proof -
   1.445 +  have "- (2 * pi) < 0" and "0 < pi" by auto
   1.446 +  from order_less_trans[OF this] show ?thesis .
   1.447 +qed
   1.448 +
   1.449  lemma sin_pi_half [simp]: "sin(pi/2) = 1"
   1.450  apply (cut_tac x = "pi/2" in sin_cos_squared_add2)
   1.451  apply (cut_tac sin_gt_zero [OF pi_half_gt_zero pi_half_less_two])
   1.452 @@ -1487,6 +1908,24 @@
   1.453  apply (auto intro!: cos_gt_zero_pi simp del: sin_cos_eq [symmetric])
   1.454  done
   1.455  
   1.456 +
   1.457 +lemma pi_ge_two: "2 \<le> pi"
   1.458 +proof (rule ccontr)
   1.459 +  assume "\<not> 2 \<le> pi" hence "pi < 2" by auto
   1.460 +  have "\<exists>y > pi. y < 2 \<and> y < 2 * pi"
   1.461 +  proof (cases "2 < 2 * pi")
   1.462 +    case True with dense[OF `pi < 2`] show ?thesis by auto
   1.463 +  next
   1.464 +    case False have "pi < 2 * pi" by auto
   1.465 +    from dense[OF this] and False show ?thesis by auto
   1.466 +  qed
   1.467 +  then obtain y where "pi < y" and "y < 2" and "y < 2 * pi" by blast
   1.468 +  hence "0 < sin y" using sin_gt_zero by auto
   1.469 +  moreover 
   1.470 +  have "sin y < 0" using sin_gt_zero_pi[of "y - pi"] `pi < y` and `y < 2 * pi` sin_periodic_pi[of "y - pi"] by auto
   1.471 +  ultimately show False by auto
   1.472 +qed
   1.473 +
   1.474  lemma sin_ge_zero: "[| 0 \<le> x; x \<le> pi |] ==> 0 \<le> sin x"
   1.475  by (auto simp add: order_le_less sin_gt_zero_pi)
   1.476  
   1.477 @@ -1586,6 +2025,48 @@
   1.478  apply (auto simp add: even_mult_two_ex)
   1.479  done
   1.480  
   1.481 +lemma cos_monotone_0_pi: assumes "0 \<le> y" and "y < x" and "x \<le> pi"
   1.482 +  shows "cos x < cos y"
   1.483 +proof -
   1.484 +  have "- (x - y) < 0" by (auto!)
   1.485 +
   1.486 +  from MVT2[OF `y < x` DERIV_cos[THEN impI, THEN allI]]
   1.487 +  obtain z where "y < z" and "z < x" and cos_diff: "cos x - cos y = (x - y) * - sin z" by auto
   1.488 +  hence "0 < z" and "z < pi" by (auto!)
   1.489 +  hence "0 < sin z" using sin_gt_zero_pi by auto
   1.490 +  hence "cos x - cos y < 0" unfolding cos_diff minus_mult_commute[symmetric] using `- (x - y) < 0` by (rule mult_pos_neg2)
   1.491 +  thus ?thesis by auto
   1.492 +qed
   1.493 +
   1.494 +lemma cos_monotone_0_pi': assumes "0 \<le> y" and "y \<le> x" and "x \<le> pi" shows "cos x \<le> cos y"
   1.495 +proof (cases "y < x")
   1.496 +  case True show ?thesis using cos_monotone_0_pi[OF `0 \<le> y` True `x \<le> pi`] by auto
   1.497 +next
   1.498 +  case False hence "y = x" using `y \<le> x` by auto
   1.499 +  thus ?thesis by auto
   1.500 +qed
   1.501 +
   1.502 +lemma cos_monotone_minus_pi_0: assumes "-pi \<le> y" and "y < x" and "x \<le> 0"
   1.503 +  shows "cos y < cos x"
   1.504 +proof -
   1.505 +  have "0 \<le> -x" and "-x < -y" and "-y \<le> pi" by (auto!)
   1.506 +  from cos_monotone_0_pi[OF this]
   1.507 +  show ?thesis unfolding cos_minus .
   1.508 +qed
   1.509 +
   1.510 +lemma cos_monotone_minus_pi_0': assumes "-pi \<le> y" and "y \<le> x" and "x \<le> 0" shows "cos y \<le> cos x"
   1.511 +proof (cases "y < x")
   1.512 +  case True show ?thesis using cos_monotone_minus_pi_0[OF `-pi \<le> y` True `x \<le> 0`] by auto
   1.513 +next
   1.514 +  case False hence "y = x" using `y \<le> x` by auto
   1.515 +  thus ?thesis by auto
   1.516 +qed
   1.517 +
   1.518 +lemma sin_monotone_2pi': assumes "- (pi / 2) \<le> y" and "y \<le> x" and "x \<le> pi / 2" shows "sin y \<le> sin x"
   1.519 +proof -
   1.520 +  have "0 \<le> y + pi / 2" and "y + pi / 2 \<le> x + pi / 2" and "x + pi /2 \<le> pi" using pi_ge_two by (auto!)
   1.521 +  from cos_monotone_0_pi'[OF this] show ?thesis unfolding minus_sin_cos_eq[symmetric] by auto
   1.522 +qed
   1.523  
   1.524  subsection {* Tangent *}
   1.525  
   1.526 @@ -1653,6 +2134,22 @@
   1.527    thus ?thesis by simp
   1.528  qed
   1.529  
   1.530 +lemma tan_half: fixes x :: real assumes "- (pi / 2) < x" and "x < pi / 2"
   1.531 +  shows "tan x = sin (2 * x) / (cos (2 * x) + 1)"
   1.532 +proof -
   1.533 +  from cos_gt_zero_pi[OF `- (pi / 2) < x` `x < pi / 2`]
   1.534 +  have "cos x \<noteq> 0" by auto
   1.535 +
   1.536 +  have minus_cos_2x: "\<And>X. X - cos (2*x) = X - (cos x) ^ 2 + (sin x) ^ 2" unfolding cos_double by algebra
   1.537 +
   1.538 +  have "tan x = (tan x + tan x) / 2" by auto
   1.539 +  also have "\<dots> = sin (x + x) / (cos x * cos x) / 2" unfolding add_tan_eq[OF `cos x \<noteq> 0` `cos x \<noteq> 0`] ..
   1.540 +  also have "\<dots> = sin (2 * x) / ((cos x) ^ 2 + (cos x) ^ 2 + cos (2*x) - cos (2*x))" unfolding divide_divide_eq_left numeral_2_eq_2 by auto
   1.541 +  also have "\<dots> = sin (2 * x) / ((cos x) ^ 2 + cos (2*x) + (sin x)^2)" unfolding minus_cos_2x by auto
   1.542 +  also have "\<dots> = sin (2 * x) / (cos (2*x) + 1)" by auto
   1.543 +  finally show ?thesis .
   1.544 +qed
   1.545 +
   1.546  lemma lemma_DERIV_tan:
   1.547       "cos x \<noteq> 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse((cos x)\<twosuperior>)"
   1.548  apply (rule lemma_DERIV_subst)
   1.549 @@ -1726,6 +2223,73 @@
   1.550  	    simp add: cos_gt_zero_pi [THEN less_imp_neq, THEN not_sym]) 
   1.551  done
   1.552  
   1.553 +lemma tan_monotone: assumes "- (pi / 2) < y" and "y < x" and "x < pi / 2"
   1.554 +  shows "tan y < tan x"
   1.555 +proof -
   1.556 +  have "\<forall> x'. y \<le> x' \<and> x' \<le> x \<longrightarrow> DERIV tan x' :> inverse (cos x'^2)"
   1.557 +  proof (rule allI, rule impI)
   1.558 +    fix x' :: real assume "y \<le> x' \<and> x' \<le> x"
   1.559 +    hence "-(pi/2) < x'" and "x' < pi/2" by (auto!)
   1.560 +    from cos_gt_zero_pi[OF this]
   1.561 +    have "cos x' \<noteq> 0" by auto
   1.562 +    thus "DERIV tan x' :> inverse (cos x'^2)" by (rule DERIV_tan)
   1.563 +  qed
   1.564 +  from MVT2[OF `y < x` this] 
   1.565 +  obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<twosuperior>)" by auto
   1.566 +  hence "- (pi / 2) < z" and "z < pi / 2" by (auto!)
   1.567 +  hence "0 < cos z" using cos_gt_zero_pi by auto
   1.568 +  hence inv_pos: "0 < inverse ((cos z)\<twosuperior>)" by auto
   1.569 +  have "0 < x - y" using `y < x` by auto
   1.570 +  from real_mult_order[OF this inv_pos]
   1.571 +  have "0 < tan x - tan y" unfolding tan_diff by auto
   1.572 +  thus ?thesis by auto
   1.573 +qed
   1.574 +
   1.575 +lemma tan_monotone': assumes "- (pi / 2) < y" and "y < pi / 2" and "- (pi / 2) < x" and "x < pi / 2"
   1.576 +  shows "(y < x) = (tan y < tan x)"
   1.577 +proof
   1.578 +  assume "y < x" thus "tan y < tan x" using tan_monotone and `- (pi / 2) < y` and `x < pi / 2` by auto
   1.579 +next
   1.580 +  assume "tan y < tan x"
   1.581 +  show "y < x"
   1.582 +  proof (rule ccontr)
   1.583 +    assume "\<not> y < x" hence "x \<le> y" by auto
   1.584 +    hence "tan x \<le> tan y" 
   1.585 +    proof (cases "x = y")
   1.586 +      case True thus ?thesis by auto
   1.587 +    next
   1.588 +      case False hence "x < y" using `x \<le> y` by auto
   1.589 +      from tan_monotone[OF `- (pi/2) < x` this `y < pi / 2`] show ?thesis by auto
   1.590 +    qed
   1.591 +    thus False using `tan y < tan x` by auto
   1.592 +  qed
   1.593 +qed
   1.594 +
   1.595 +lemma tan_inverse: "1 / (tan y) = tan (pi / 2 - y)" unfolding tan_def sin_cos_eq[of y] cos_sin_eq[of y] by auto
   1.596 +
   1.597 +lemma tan_periodic_pi[simp]: "tan (x + pi) = tan x" 
   1.598 +  by (simp add: tan_def)
   1.599 +
   1.600 +lemma tan_periodic_nat[simp]: fixes n :: nat shows "tan (x + real n * pi) = tan x" 
   1.601 +proof (induct n arbitrary: x)
   1.602 +  case (Suc n)
   1.603 +  have split_pi_off: "x + real (Suc n) * pi = (x + real n * pi) + pi" unfolding Suc_plus1 real_of_nat_add real_of_one real_add_mult_distrib by auto
   1.604 +  show ?case unfolding split_pi_off using Suc by auto
   1.605 +qed auto
   1.606 +
   1.607 +lemma tan_periodic_int[simp]: fixes i :: int shows "tan (x + real i * pi) = tan x"
   1.608 +proof (cases "0 \<le> i")
   1.609 +  case True hence i_nat: "real i = real (nat i)" by auto
   1.610 +  show ?thesis unfolding i_nat by auto
   1.611 +next
   1.612 +  case False hence i_nat: "real i = - real (nat (-i))" by auto
   1.613 +  have "tan x = tan (x + real i * pi - real i * pi)" by auto
   1.614 +  also have "\<dots> = tan (x + real i * pi)" unfolding i_nat mult_minus_left diff_minus_eq_add by (rule tan_periodic_nat)
   1.615 +  finally show ?thesis by auto
   1.616 +qed
   1.617 +
   1.618 +lemma tan_periodic_n[simp]: "tan (x + number_of n * pi) = tan x"
   1.619 +  using tan_periodic_int[of _ "number_of n" ] unfolding real_number_of .
   1.620  
   1.621  subsection {* Inverse Trigonometric Functions *}
   1.622  
   1.623 @@ -1968,7 +2532,6 @@
   1.624  apply (simp, simp, simp, rule isCont_arctan)
   1.625  done
   1.626  
   1.627 -
   1.628  subsection {* More Theorems about Sin and Cos *}
   1.629  
   1.630  lemma cos_45: "cos (pi / 4) = sqrt 2 / 2"
   1.631 @@ -2115,6 +2678,464 @@
   1.632  lemma cos_one_sin_zero: "cos x = 1 ==> sin x = 0"
   1.633  by (cut_tac x = x in sin_cos_squared_add3, auto)
   1.634  
   1.635 +subsection {* Machins formula *}
   1.636 +
   1.637 +lemma tan_total_pi4: assumes "\<bar>x\<bar> < 1"
   1.638 +  shows "\<exists> z. - (pi / 4) < z \<and> z < pi / 4 \<and> tan z = x"
   1.639 +proof -
   1.640 +  obtain z where "- (pi / 2) < z" and "z < pi / 2" and "tan z = x" using tan_total by blast
   1.641 +  have "tan (pi / 4) = 1" and "tan (- (pi / 4)) = - 1" using tan_45 tan_minus by auto
   1.642 +  have "z \<noteq> pi / 4" 
   1.643 +  proof (rule ccontr)
   1.644 +    assume "\<not> (z \<noteq> pi / 4)" hence "z = pi / 4" by auto
   1.645 +    have "tan z = 1" unfolding `z = pi / 4` `tan (pi / 4) = 1` ..
   1.646 +    thus False unfolding `tan z = x` using `\<bar>x\<bar> < 1` by auto
   1.647 +  qed
   1.648 +  have "z \<noteq> - (pi / 4)"
   1.649 +  proof (rule ccontr)
   1.650 +    assume "\<not> (z \<noteq> - (pi / 4))" hence "z = - (pi / 4)" by auto
   1.651 +    have "tan z = - 1" unfolding `z = - (pi / 4)` `tan (- (pi / 4)) = - 1` ..
   1.652 +    thus False unfolding `tan z = x` using `\<bar>x\<bar> < 1` by auto
   1.653 +  qed
   1.654 +
   1.655 +  have "z < pi / 4"
   1.656 +  proof (rule ccontr)
   1.657 +    assume "\<not> (z < pi / 4)" hence "pi / 4 < z" using `z \<noteq> pi / 4` by auto
   1.658 +    have "- (pi / 2) < pi / 4" using m2pi_less_pi by auto
   1.659 +    from tan_monotone[OF this `pi / 4 < z` `z < pi / 2`] 
   1.660 +    have "1 < x" unfolding `tan z = x` `tan (pi / 4) = 1` .
   1.661 +    thus False using `\<bar>x\<bar> < 1` by auto
   1.662 +  qed
   1.663 +  moreover 
   1.664 +  have "-(pi / 4) < z"
   1.665 +  proof (rule ccontr)
   1.666 +    assume "\<not> (-(pi / 4) < z)" hence "z < - (pi / 4)" using `z \<noteq> - (pi / 4)` by auto
   1.667 +    have "-(pi / 4) < pi / 2" using m2pi_less_pi by auto
   1.668 +    from tan_monotone[OF `-(pi / 2) < z` `z < -(pi / 4)` this]
   1.669 +    have "x < - 1" unfolding `tan z = x` `tan (-(pi / 4)) = - 1` .
   1.670 +    thus False using `\<bar>x\<bar> < 1` by auto
   1.671 +  qed
   1.672 +  ultimately show ?thesis using `tan z = x` by auto
   1.673 +qed
   1.674 +
   1.675 +lemma arctan_add: assumes "\<bar>x\<bar> \<le> 1" and "\<bar>y\<bar> < 1"
   1.676 +  shows "arctan x + arctan y = arctan ((x + y) / (1 - x * y))"
   1.677 +proof -
   1.678 +  obtain y' where "-(pi/4) < y'" and "y' < pi/4" and "tan y' = y" using tan_total_pi4[OF `\<bar>y\<bar> < 1`] by blast
   1.679 +
   1.680 +  have "pi / 4 < pi / 2" by auto
   1.681 +
   1.682 +  have "\<exists> x'. -(pi/4) \<le> x' \<and> x' \<le> pi/4 \<and> tan x' = x"
   1.683 +  proof (cases "\<bar>x\<bar> < 1")
   1.684 +    case True from tan_total_pi4[OF this] obtain x' where "-(pi/4) < x'" and "x' < pi/4" and "tan x' = x" by blast
   1.685 +    hence "-(pi/4) \<le> x'" and "x' \<le> pi/4" and "tan x' = x" by auto
   1.686 +    thus ?thesis by auto
   1.687 +  next
   1.688 +    case False
   1.689 +    show ?thesis
   1.690 +    proof (cases "x = 1")
   1.691 +      case True hence "tan (pi/4) = x" using tan_45 by auto
   1.692 +      moreover 
   1.693 +      have "- pi \<le> pi" unfolding minus_le_self_iff by auto
   1.694 +      hence "-(pi/4) \<le> pi/4" and "pi/4 \<le> pi/4" by auto
   1.695 +      ultimately show ?thesis by blast
   1.696 +    next
   1.697 +      case False hence "x = -1" using `\<not> \<bar>x\<bar> < 1` and `\<bar>x\<bar> \<le> 1` by auto
   1.698 +      hence "tan (-(pi/4)) = x" using tan_45 tan_minus by auto
   1.699 +      moreover 
   1.700 +      have "- pi \<le> pi" unfolding minus_le_self_iff by auto
   1.701 +      hence "-(pi/4) \<le> pi/4" and "-(pi/4) \<le> -(pi/4)" by auto
   1.702 +      ultimately show ?thesis by blast
   1.703 +    qed
   1.704 +  qed
   1.705 +  then obtain x' where "-(pi/4) \<le> x'" and "x' \<le> pi/4" and "tan x' = x" by blast
   1.706 +  hence "-(pi/2) < x'" and "x' < pi/2" using order_le_less_trans[OF `x' \<le> pi/4` `pi / 4 < pi / 2`] by auto
   1.707 +
   1.708 +  have "cos x' \<noteq> 0" using cos_gt_zero_pi[THEN less_imp_neq] and `-(pi/2) < x'` and `x' < pi/2` by auto
   1.709 +  moreover have "cos y' \<noteq> 0" using cos_gt_zero_pi[THEN less_imp_neq] and `-(pi/4) < y'` and `y' < pi/4` by auto
   1.710 +  ultimately have "cos x' * cos y' \<noteq> 0" by auto
   1.711 +
   1.712 +  have divide_nonzero_divide: "\<And> A B C :: real. C \<noteq> 0 \<Longrightarrow> (A / C) / (B / C) = A / B" by auto
   1.713 +  have divide_mult_commute: "\<And> A B C D :: real. A * B / (C * D) = (A / C) * (B / D)" by auto
   1.714 +
   1.715 +  have "tan (x' + y') = sin (x' + y') / (cos x' * cos y' - sin x' * sin y')" unfolding tan_def cos_add ..
   1.716 +  also have "\<dots> = (tan x' + tan y') / ((cos x' * cos y' - sin x' * sin y') / (cos x' * cos y'))" unfolding add_tan_eq[OF `cos x' \<noteq> 0` `cos y' \<noteq> 0`] divide_nonzero_divide[OF `cos x' * cos y' \<noteq> 0`] ..
   1.717 +  also have "\<dots> = (tan x' + tan y') / (1 - tan x' * tan y')" unfolding tan_def diff_divide_distrib divide_self[OF `cos x' * cos y' \<noteq> 0`] unfolding divide_mult_commute ..
   1.718 +  finally have tan_eq: "tan (x' + y') = (x + y) / (1 - x * y)" unfolding `tan y' = y` `tan x' = x` .
   1.719 +
   1.720 +  have "arctan (tan (x' + y')) = x' + y'" using `-(pi/4) < y'` `-(pi/4) \<le> x'` `y' < pi/4` and `x' \<le> pi/4` by (auto intro!: arctan_tan)
   1.721 +  moreover have "arctan (tan (x')) = x'" using `-(pi/2) < x'` and `x' < pi/2` by (auto intro!: arctan_tan)
   1.722 +  moreover have "arctan (tan (y')) = y'" using `-(pi/4) < y'` and `y' < pi/4` by (auto intro!: arctan_tan)
   1.723 +  ultimately have "arctan x + arctan y = arctan (tan (x' + y'))" unfolding `tan y' = y` [symmetric] `tan x' = x`[symmetric] by auto
   1.724 +  thus "arctan x + arctan y = arctan ((x + y) / (1 - x * y))" unfolding tan_eq .
   1.725 +qed
   1.726 +
   1.727 +lemma arctan1_eq_pi4: "arctan 1 = pi / 4" unfolding tan_45[symmetric] by (rule arctan_tan, auto simp add: m2pi_less_pi)
   1.728 +
   1.729 +theorem machin: "pi / 4 = 4 * arctan (1/5) - arctan (1 / 239)"
   1.730 +proof -
   1.731 +  have "\<bar>1 / 5\<bar> < (1 :: real)" by auto
   1.732 +  from arctan_add[OF less_imp_le[OF this] this]
   1.733 +  have "2 * arctan (1 / 5) = arctan (5 / 12)" by auto
   1.734 +  moreover
   1.735 +  have "\<bar>5 / 12\<bar> < (1 :: real)" by auto
   1.736 +  from arctan_add[OF less_imp_le[OF this] this]
   1.737 +  have "2 * arctan (5 / 12) = arctan (120 / 119)" by auto
   1.738 +  moreover 
   1.739 +  have "\<bar>1\<bar> \<le> (1::real)" and "\<bar>1 / 239\<bar> < (1::real)" by auto
   1.740 +  from arctan_add[OF this]
   1.741 +  have "arctan 1 + arctan (1 / 239) = arctan (120 / 119)" by auto
   1.742 +  ultimately have "arctan 1 + arctan (1 / 239) = 4 * arctan (1 / 5)" by auto
   1.743 +  thus ?thesis unfolding arctan1_eq_pi4 by algebra
   1.744 +qed
   1.745 +subsection {* Introducing the arcus tangens power series *}
   1.746 +
   1.747 +lemma monoseq_arctan_series: fixes x :: real
   1.748 +  assumes "\<bar>x\<bar> \<le> 1" shows "monoseq (\<lambda> n. 1 / real (n*2+1) * x^(n*2+1))" (is "monoseq ?a")
   1.749 +proof (cases "x = 0") case True thus ?thesis unfolding monoseq_def by auto
   1.750 +next
   1.751 +  case False
   1.752 +  have "norm x \<le> 1" and "x \<le> 1" and "-1 \<le> x" using assms by auto
   1.753 +  show "monoseq ?a"
   1.754 +  proof -
   1.755 +    { fix n fix x :: real assume "0 \<le> x" and "x \<le> 1"
   1.756 +      have "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \<le> 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)"
   1.757 +      proof (rule mult_mono)
   1.758 +	show "1 / real (Suc (Suc n * 2)) \<le> 1 / real (Suc (n * 2))" by (rule frac_le) simp_all
   1.759 +	show "0 \<le> 1 / real (Suc (n * 2))" by auto
   1.760 +	show "x ^ Suc (Suc n * 2) \<le> x ^ Suc (n * 2)" by (rule power_decreasing) (simp_all add: `0 \<le> x` `x \<le> 1`)
   1.761 +	show "0 \<le> x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: `0 \<le> x`)
   1.762 +      qed
   1.763 +    } note mono = this
   1.764 +    
   1.765 +    show ?thesis
   1.766 +    proof (cases "0 \<le> x")
   1.767 +      case True from mono[OF this `x \<le> 1`, THEN allI]
   1.768 +      show ?thesis unfolding Suc_plus1[symmetric] by (rule mono_SucI2)
   1.769 +    next
   1.770 +      case False hence "0 \<le> -x" and "-x \<le> 1" using `-1 \<le> x` by auto
   1.771 +      from mono[OF this]
   1.772 +      have "\<And>n. 1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \<ge> 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" using `0 \<le> -x` by auto
   1.773 +      thus ?thesis unfolding Suc_plus1[symmetric] by (rule mono_SucI1[OF allI])
   1.774 +    qed
   1.775 +  qed
   1.776 +qed
   1.777 +
   1.778 +lemma zeroseq_arctan_series: fixes x :: real
   1.779 +  assumes "\<bar>x\<bar> \<le> 1" shows "(\<lambda> n. 1 / real (n*2+1) * x^(n*2+1)) ----> 0" (is "?a ----> 0")
   1.780 +proof (cases "x = 0") case True thus ?thesis by (auto simp add: LIMSEQ_const)
   1.781 +next
   1.782 +  case False
   1.783 +  have "norm x \<le> 1" and "x \<le> 1" and "-1 \<le> x" using assms by auto
   1.784 +  show "?a ----> 0"
   1.785 +  proof (cases "\<bar>x\<bar> < 1")
   1.786 +    case True hence "norm x < 1" by auto
   1.787 +    from LIMSEQ_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_power_zero[OF `norm x < 1`, THEN LIMSEQ_Suc]]
   1.788 +    show ?thesis unfolding inverse_eq_divide Suc_plus1 using LIMSEQ_linear[OF _ pos2] by auto
   1.789 +  next
   1.790 +    case False hence "x = -1 \<or> x = 1" using `\<bar>x\<bar> \<le> 1` by auto
   1.791 +    hence n_eq: "\<And> n. x ^ (n * 2 + 1) = x" by auto
   1.792 +    from LIMSEQ_mult[OF LIMSEQ_inverse_real_of_nat[THEN LIMSEQ_linear, OF pos2, unfolded inverse_eq_divide] LIMSEQ_const[of x]]
   1.793 +    show ?thesis unfolding n_eq by auto
   1.794 +  qed
   1.795 +qed
   1.796 +
   1.797 +lemma summable_arctan_series: fixes x :: real and n :: nat
   1.798 +  assumes "\<bar>x\<bar> \<le> 1" shows "summable (\<lambda> k. (-1)^k * (1 / real (k*2+1) * x ^ (k*2+1)))" (is "summable (?c x)")
   1.799 +  by (rule summable_Leibniz(1), rule zeroseq_arctan_series[OF assms], rule monoseq_arctan_series[OF assms])
   1.800 +
   1.801 +lemma less_one_imp_sqr_less_one: fixes x :: real assumes "\<bar>x\<bar> < 1" shows "x^2 < 1"
   1.802 +proof -
   1.803 +  from mult_mono1[OF less_imp_le[OF `\<bar>x\<bar> < 1`] abs_ge_zero[of x]]
   1.804 +  have "\<bar> x^2 \<bar> < 1" using `\<bar> x \<bar> < 1` unfolding numeral_2_eq_2 power_Suc2 by auto
   1.805 +  thus ?thesis using zero_le_power2 by auto
   1.806 +qed 
   1.807 +
   1.808 +lemma DERIV_arctan_series: assumes "\<bar> x \<bar> < 1"
   1.809 +  shows "DERIV (\<lambda> x'. \<Sum> k. (-1)^k * (1 / real (k*2+1) * x' ^ (k*2+1))) x :> (\<Sum> k. (-1)^k * x^(k*2))" (is "DERIV ?arctan _ :> ?Int")
   1.810 +proof -
   1.811 +  let "?f n" = "if even n then (-1)^(n div 2) * 1 / real (Suc n) else 0"
   1.812 +
   1.813 +  { fix n :: nat assume "even n" hence "2 * (n div 2) = n" by presburger } note n_even=this
   1.814 +  have if_eq: "\<And> n x'. ?f n * real (Suc n) * x'^n = (if even n then (-1)^(n div 2) * x'^(2 * (n div 2)) else 0)" using n_even by auto
   1.815 +
   1.816 +  { fix x :: real assume "\<bar>x\<bar> < 1" hence "x^2 < 1" by (rule less_one_imp_sqr_less_one)
   1.817 +    have "summable (\<lambda> n. -1 ^ n * (x^2) ^n)"
   1.818 +      by (rule summable_Leibniz(1), auto intro!: LIMSEQ_realpow_zero monoseq_realpow `x^2 < 1` order_less_imp_le[OF `x^2 < 1`])
   1.819 +    hence "summable (\<lambda> n. -1 ^ n * x^(2*n))" unfolding power_mult .
   1.820 +  } note summable_Integral = this
   1.821 +
   1.822 +  { fix f :: "nat \<Rightarrow> real"
   1.823 +    have "\<And> x. f sums x = (\<lambda> n. if even n then f (n div 2) else 0) sums x"
   1.824 +    proof
   1.825 +      fix x :: real assume "f sums x" 
   1.826 +      from sums_if[OF sums_zero this]
   1.827 +      show "(\<lambda> n. if even n then f (n div 2) else 0) sums x" by auto
   1.828 +    next
   1.829 +      fix x :: real assume "(\<lambda> n. if even n then f (n div 2) else 0) sums x"
   1.830 +      from LIMSEQ_linear[OF this[unfolded sums_def] pos2, unfolded sum_split_even_odd[unfolded mult_commute]]
   1.831 +      show "f sums x" unfolding sums_def by auto
   1.832 +    qed
   1.833 +    hence "op sums f = op sums (\<lambda> n. if even n then f (n div 2) else 0)" ..
   1.834 +  } note sums_even = this
   1.835 +
   1.836 +  have Int_eq: "(\<Sum> n. ?f n * real (Suc n) * x^n) = ?Int" unfolding if_eq mult_commute[of _ 2] suminf_def sums_even[of "\<lambda> n. -1 ^ n * x ^ (2 * n)", symmetric]
   1.837 +    by auto
   1.838 +
   1.839 +  { fix x :: real
   1.840 +    have if_eq': "\<And> n. (if even n then -1 ^ (n div 2) * 1 / real (Suc n) else 0) * x ^ Suc n = 
   1.841 +      (if even n then -1 ^ (n div 2) * (1 / real (Suc (2 * (n div 2))) * x ^ Suc (2 * (n div 2))) else 0)"
   1.842 +      using n_even by auto
   1.843 +    have idx_eq: "\<And> n. n * 2 + 1 = Suc (2 * n)" by auto 
   1.844 +    have "(\<Sum> n. ?f n * x^(Suc n)) = ?arctan x" unfolding if_eq' idx_eq suminf_def sums_even[of "\<lambda> n. -1 ^ n * (1 / real (Suc (2 * n)) * x ^ Suc (2 * n))", symmetric]
   1.845 +      by auto
   1.846 +  } note arctan_eq = this
   1.847 +
   1.848 +  have "DERIV (\<lambda> x. \<Sum> n. ?f n * x^(Suc n)) x :> (\<Sum> n. ?f n * real (Suc n) * x^n)"
   1.849 +  proof (rule DERIV_power_series')
   1.850 +    show "x \<in> {- 1 <..< 1}" using `\<bar> x \<bar> < 1` by auto
   1.851 +    { fix x' :: real assume x'_bounds: "x' \<in> {- 1 <..< 1}"
   1.852 +      hence "\<bar>x'\<bar> < 1" by auto
   1.853 +
   1.854 +      let ?S = "\<Sum> n. (-1)^n * x'^(2 * n)"
   1.855 +      show "summable (\<lambda> n. ?f n * real (Suc n) * x'^n)" unfolding if_eq
   1.856 +	by (rule sums_summable[where l="0 + ?S"], rule sums_if, rule sums_zero, rule summable_sums, rule summable_Integral[OF `\<bar>x'\<bar> < 1`])
   1.857 +    }
   1.858 +  qed auto
   1.859 +  thus ?thesis unfolding Int_eq arctan_eq .
   1.860 +qed
   1.861 +
   1.862 +lemma arctan_series: assumes "\<bar> x \<bar> \<le> 1"
   1.863 +  shows "arctan x = (\<Sum> k. (-1)^k * (1 / real (k*2+1) * x ^ (k*2+1)))" (is "_ = suminf (\<lambda> n. ?c x n)")
   1.864 +proof -
   1.865 +  let "?c' x n" = "(-1)^n * x^(n*2)"
   1.866 +
   1.867 +  { fix r x :: real assume "0 < r" and "r < 1" and "\<bar> x \<bar> < r"
   1.868 +    have "\<bar>x\<bar> < 1" using `r < 1` and `\<bar>x\<bar> < r` by auto
   1.869 +    from DERIV_arctan_series[OF this]
   1.870 +    have "DERIV (\<lambda> x. suminf (?c x)) x :> (suminf (?c' x))" .
   1.871 +  } note DERIV_arctan_suminf = this
   1.872 +
   1.873 +  { fix x :: real assume "\<bar>x\<bar> \<le> 1" note summable_Leibniz[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] }
   1.874 +  note arctan_series_borders = this
   1.875 +
   1.876 +  { fix x :: real assume "\<bar>x\<bar> < 1" have "arctan x = (\<Sum> k. ?c x k)"
   1.877 +  proof -
   1.878 +    obtain r where "\<bar>x\<bar> < r" and "r < 1" using dense[OF `\<bar>x\<bar> < 1`] by blast
   1.879 +    hence "0 < r" and "-r < x" and "x < r" by auto
   1.880 +
   1.881 +    have suminf_eq_arctan_bounded: "\<And> x a b. \<lbrakk> -r < a ; b < r ; a < b ; a \<le> x ; x \<le> b \<rbrakk> \<Longrightarrow> suminf (?c x) - arctan x = suminf (?c a) - arctan a"
   1.882 +    proof -
   1.883 +      fix x a b assume "-r < a" and "b < r" and "a < b" and "a \<le> x" and "x \<le> b"
   1.884 +      hence "\<bar>x\<bar> < r" by auto
   1.885 +      show "suminf (?c x) - arctan x = suminf (?c a) - arctan a"
   1.886 +      proof (rule DERIV_isconst2[of "a" "b"])
   1.887 +	show "a < b" and "a \<le> x" and "x \<le> b" using `a < b` `a \<le> x` `x \<le> b` by auto
   1.888 +	have "\<forall> x. -r < x \<and> x < r \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
   1.889 +	proof (rule allI, rule impI)
   1.890 +	  fix x assume "-r < x \<and> x < r" hence "\<bar>x\<bar> < r" by auto
   1.891 +	  hence "\<bar>x\<bar> < 1" using `r < 1` by auto
   1.892 +	  have "\<bar> - (x^2) \<bar> < 1" using less_one_imp_sqr_less_one[OF `\<bar>x\<bar> < 1`] by auto
   1.893 +	  hence "(\<lambda> n. (- (x^2)) ^ n) sums (1 / (1 - (- (x^2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums)
   1.894 +	  hence "(?c' x) sums (1 / (1 - (- (x^2))))" unfolding power_mult_distrib[symmetric] power_mult nat_mult_commute[of _ 2] by auto
   1.895 +	  hence suminf_c'_eq_geom: "inverse (1 + x^2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto
   1.896 +	  have "DERIV (\<lambda> x. suminf (?c x)) x :> (inverse (1 + x^2))" unfolding suminf_c'_eq_geom
   1.897 +	    by (rule DERIV_arctan_suminf[OF `0 < r` `r < 1` `\<bar>x\<bar> < r`])
   1.898 +	  from DERIV_add_minus[OF this DERIV_arctan]
   1.899 +	  show "DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0" unfolding diff_minus by auto
   1.900 +	qed
   1.901 +	hence DERIV_in_rball: "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0" using `-r < a` `b < r` by auto
   1.902 +	thus "\<forall> y. a < y \<and> y < b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0" using `\<bar>x\<bar> < r` by auto
   1.903 +	show "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> isCont (\<lambda> x. suminf (?c x) - arctan x) y" using DERIV_in_rball DERIV_isCont by auto
   1.904 +      qed
   1.905 +    qed
   1.906 +    
   1.907 +    have suminf_arctan_zero: "suminf (?c 0) - arctan 0 = 0"
   1.908 +      unfolding Suc_plus1[symmetric] power_Suc2 mult_zero_right arctan_zero_zero suminf_zero by auto
   1.909 +    
   1.910 +    have "suminf (?c x) - arctan x = 0"
   1.911 +    proof (cases "x = 0")
   1.912 +      case True thus ?thesis using suminf_arctan_zero by auto
   1.913 +    next
   1.914 +      case False hence "0 < \<bar>x\<bar>" and "- \<bar>x\<bar> < \<bar>x\<bar>" by auto
   1.915 +      have "suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>) = suminf (?c 0) - arctan 0"
   1.916 +	by (rule suminf_eq_arctan_bounded[where x="0" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>", symmetric], auto simp add: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>`)
   1.917 +      moreover
   1.918 +      have "suminf (?c x) - arctan x = suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>)"
   1.919 +	by (rule suminf_eq_arctan_bounded[where x="x" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>"], auto simp add: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>`)
   1.920 +      ultimately 
   1.921 +      show ?thesis using suminf_arctan_zero by auto
   1.922 +    qed
   1.923 +    thus ?thesis by auto
   1.924 +  qed } note when_less_one = this
   1.925 +
   1.926 +  show "arctan x = suminf (\<lambda> n. ?c x n)"
   1.927 +  proof (cases "\<bar>x\<bar> < 1")
   1.928 +    case True thus ?thesis by (rule when_less_one)
   1.929 +  next case False hence "\<bar>x\<bar> = 1" using `\<bar>x\<bar> \<le> 1` by auto
   1.930 +    let "?a x n" = "\<bar>1 / real (n*2+1) * x^(n*2+1)\<bar>"
   1.931 +    let "?diff x n" = "\<bar> arctan x - (\<Sum> i = 0..<n. ?c x i)\<bar>"
   1.932 +    { fix n :: nat
   1.933 +      have "0 < (1 :: real)" by auto
   1.934 +      moreover
   1.935 +      { fix x :: real assume "0 < x" and "x < 1" hence "\<bar>x\<bar> \<le> 1" and "\<bar>x\<bar> < 1" by auto
   1.936 +	from `0 < x` have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)" by auto
   1.937 +	note bounds = mp[OF arctan_series_borders(2)[OF `\<bar>x\<bar> \<le> 1`] this, unfolded when_less_one[OF `\<bar>x\<bar> < 1`, symmetric], THEN spec]
   1.938 +	have "0 < 1 / real (n*2+1) * x^(n*2+1)" by (rule mult_pos_pos, auto simp only: zero_less_power[OF `0 < x`], auto)
   1.939 +	hence a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)" by (rule abs_of_pos)
   1.940 +        have "?diff x n \<le> ?a x n"
   1.941 +	proof (cases "even n")
   1.942 +	  case True hence sgn_pos: "(-1)^n = (1::real)" by auto
   1.943 +	  from `even n` obtain m where "2 * m = n" unfolding even_mult_two_ex by auto
   1.944 +	  from bounds[of m, unfolded this atLeastAtMost_iff]
   1.945 +	  have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n + 1. (?c x i)) - (\<Sum>i = 0..<n. (?c x i))" by auto
   1.946 +	  also have "\<dots> = ?c x n" by auto
   1.947 +	  also have "\<dots> = ?a x n" unfolding sgn_pos a_pos by auto
   1.948 +	  finally show ?thesis .
   1.949 +	next
   1.950 +	  case False hence sgn_neg: "(-1)^n = (-1::real)" by auto
   1.951 +	  from `odd n` obtain m where m_def: "2 * m + 1 = n" unfolding odd_Suc_mult_two_ex by auto
   1.952 +	  hence m_plus: "2 * (m + 1) = n + 1" by auto
   1.953 +	  from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2]
   1.954 +	  have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n. (?c x i)) - (\<Sum>i = 0..<n+1. (?c x i))" by auto
   1.955 +	  also have "\<dots> = - ?c x n" by auto
   1.956 +	  also have "\<dots> = ?a x n" unfolding sgn_neg a_pos by auto
   1.957 +	  finally show ?thesis .
   1.958 +	qed
   1.959 +        hence "0 \<le> ?a x n - ?diff x n" by auto
   1.960 +      }
   1.961 +      hence "\<forall> x \<in> { 0 <..< 1 }. 0 \<le> ?a x n - ?diff x n" by auto
   1.962 +      moreover have "\<And>x. isCont (\<lambda> x. ?a x n - ?diff x n) x"
   1.963 +	unfolding real_diff_def divide_inverse
   1.964 +	by (auto intro!: isCont_add isCont_rabs isCont_ident isCont_minus isCont_arctan isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum)
   1.965 +      ultimately have "0 \<le> ?a 1 n - ?diff 1 n" by (rule LIM_less_bound)
   1.966 +      hence "?diff 1 n \<le> ?a 1 n" by auto
   1.967 +    }
   1.968 +    have "?a 1 ----> 0" unfolding LIMSEQ_rabs_zero power_one divide_inverse by (auto intro!: LIMSEQ_mult LIMSEQ_linear LIMSEQ_inverse_real_of_nat)
   1.969 +    have "?diff 1 ----> 0"
   1.970 +    proof (rule LIMSEQ_I)
   1.971 +      fix r :: real assume "0 < r"
   1.972 +      obtain N :: nat where N_I: "\<And> n. N \<le> n \<Longrightarrow> ?a 1 n < r" using LIMSEQ_D[OF `?a 1 ----> 0` `0 < r`] by auto
   1.973 +      { fix n assume "N \<le> n" from `?diff 1 n \<le> ?a 1 n` N_I[OF this]
   1.974 +	have "norm (?diff 1 n - 0) < r" by auto }
   1.975 +      thus "\<exists> N. \<forall> n \<ge> N. norm (?diff 1 n - 0) < r" by blast
   1.976 +    qed
   1.977 +    from this[unfolded LIMSEQ_rabs_zero real_diff_def add_commute[of "arctan 1"], THEN LIMSEQ_add_const, of "- arctan 1", THEN LIMSEQ_minus]
   1.978 +    have "(?c 1) sums (arctan 1)" unfolding sums_def by auto
   1.979 +    hence "arctan 1 = (\<Sum> i. ?c 1 i)" by (rule sums_unique)
   1.980 +    
   1.981 +    show ?thesis
   1.982 +    proof (cases "x = 1", simp add: `arctan 1 = (\<Sum> i. ?c 1 i)`)
   1.983 +      assume "x \<noteq> 1" hence "x = -1" using `\<bar>x\<bar> = 1` by auto
   1.984 +      
   1.985 +      have "- (pi / 2) < 0" using pi_gt_zero by auto
   1.986 +      have "- (2 * pi) < 0" using pi_gt_zero by auto
   1.987 +      
   1.988 +      have c_minus_minus: "\<And> i. ?c (- 1) i = - ?c 1 i" by auto
   1.989 +    
   1.990 +      have "arctan (- 1) = arctan (tan (-(pi / 4)))" unfolding tan_45 tan_minus ..
   1.991 +      also have "\<dots> = - (pi / 4)" by (rule arctan_tan, auto simp add: order_less_trans[OF `- (pi / 2) < 0` pi_gt_zero])
   1.992 +      also have "\<dots> = - (arctan (tan (pi / 4)))" unfolding neg_equal_iff_equal by (rule arctan_tan[symmetric], auto simp add: order_less_trans[OF `- (2 * pi) < 0` pi_gt_zero])
   1.993 +      also have "\<dots> = - (arctan 1)" unfolding tan_45 ..
   1.994 +      also have "\<dots> = - (\<Sum> i. ?c 1 i)" using `arctan 1 = (\<Sum> i. ?c 1 i)` by auto
   1.995 +      also have "\<dots> = (\<Sum> i. ?c (- 1) i)" using suminf_minus[OF sums_summable[OF `(?c 1) sums (arctan 1)`]] unfolding c_minus_minus by auto
   1.996 +      finally show ?thesis using `x = -1` by auto
   1.997 +    qed
   1.998 +  qed
   1.999 +qed
  1.1000 +
  1.1001 +lemma arctan_half: fixes x :: real
  1.1002 +  shows "arctan x = 2 * arctan (x / (1 + sqrt(1 + x^2)))"
  1.1003 +proof -
  1.1004 +  obtain y where low: "- (pi / 2) < y" and high: "y < pi / 2" and y_eq: "tan y = x" using tan_total by blast
  1.1005 +  hence low2: "- (pi / 2) < y / 2" and high2: "y / 2 < pi / 2" by auto
  1.1006 +
  1.1007 +  have divide_nonzero_divide: "\<And> A B C :: real. C \<noteq> 0 \<Longrightarrow> A / B = (A / C) / (B / C)" by auto
  1.1008 +  
  1.1009 +  have "0 < cos y" using cos_gt_zero_pi[OF low high] .
  1.1010 +  hence "cos y \<noteq> 0" and cos_sqrt: "sqrt ((cos y) ^ 2) = cos y" by auto
  1.1011 +
  1.1012 +  have "1 + (tan y)^2 = 1 + sin y^2 / cos y^2" unfolding tan_def power_divide ..
  1.1013 +  also have "\<dots> = cos y^2 / cos y^2 + sin y^2 / cos y^2" using `cos y \<noteq> 0` by auto
  1.1014 +  also have "\<dots> = 1 / cos y^2" unfolding add_divide_distrib[symmetric] sin_cos_squared_add2 ..
  1.1015 +  finally have "1 + (tan y)^2 = 1 / cos y^2" .
  1.1016 +
  1.1017 +  have "sin y / (cos y + 1) = tan y / ((cos y + 1) / cos y)" unfolding tan_def divide_nonzero_divide[OF `cos y \<noteq> 0`, symmetric] ..
  1.1018 +  also have "\<dots> = tan y / (1 + 1 / cos y)" using `cos y \<noteq> 0` unfolding add_divide_distrib by auto
  1.1019 +  also have "\<dots> = tan y / (1 + 1 / sqrt(cos y^2))" unfolding cos_sqrt ..
  1.1020 +  also have "\<dots> = tan y / (1 + sqrt(1 / cos y^2))" unfolding real_sqrt_divide by auto
  1.1021 +  finally have eq: "sin y / (cos y + 1) = tan y / (1 + sqrt(1 + (tan y)^2))" unfolding `1 + (tan y)^2 = 1 / cos y^2` .
  1.1022 +
  1.1023 +  have "arctan x = y" using arctan_tan low high y_eq by auto
  1.1024 +  also have "\<dots> = 2 * (arctan (tan (y/2)))" using arctan_tan[OF low2 high2] by auto
  1.1025 +  also have "\<dots> = 2 * (arctan (sin y / (cos y + 1)))" unfolding tan_half[OF low2 high2] by auto
  1.1026 +  finally show ?thesis unfolding eq `tan y = x` .
  1.1027 +qed
  1.1028 +
  1.1029 +lemma arctan_monotone: assumes "x < y"
  1.1030 +  shows "arctan x < arctan y"
  1.1031 +proof -
  1.1032 +  obtain z where "-(pi / 2) < z" and "z < pi / 2" and "tan z = x" using tan_total by blast
  1.1033 +  obtain w where "-(pi / 2) < w" and "w < pi / 2" and "tan w = y" using tan_total by blast
  1.1034 +  have "z < w" unfolding tan_monotone'[OF `-(pi / 2) < z` `z < pi / 2` `-(pi / 2) < w` `w < pi / 2`] `tan z = x` `tan w = y` using `x < y` .
  1.1035 +  thus ?thesis
  1.1036 +    unfolding `tan z = x`[symmetric] arctan_tan[OF `-(pi / 2) < z` `z < pi / 2`]
  1.1037 +    unfolding `tan w = y`[symmetric] arctan_tan[OF `-(pi / 2) < w` `w < pi / 2`] .
  1.1038 +qed
  1.1039 +
  1.1040 +lemma arctan_monotone': assumes "x \<le> y" shows "arctan x \<le> arctan y"
  1.1041 +proof (cases "x = y") 
  1.1042 +  case False hence "x < y" using `x \<le> y` by auto from arctan_monotone[OF this] show ?thesis by auto
  1.1043 +qed auto
  1.1044 +
  1.1045 +lemma arctan_minus: "arctan (- x) = - arctan x" 
  1.1046 +proof -
  1.1047 +  obtain y where "- (pi / 2) < y" and "y < pi / 2" and "tan y = x" using tan_total by blast
  1.1048 +  thus ?thesis unfolding `tan y = x`[symmetric] tan_minus[symmetric] using arctan_tan[of y] arctan_tan[of "-y"] by auto 
  1.1049 +qed
  1.1050 +
  1.1051 +lemma arctan_inverse: assumes "x \<noteq> 0" shows "arctan (1 / x) = sgn x * pi / 2 - arctan x"
  1.1052 +proof -
  1.1053 +  obtain y where "- (pi / 2) < y" and "y < pi / 2" and "tan y = x" using tan_total by blast
  1.1054 +  hence "y = arctan x" unfolding `tan y = x`[symmetric] using arctan_tan by auto
  1.1055 +
  1.1056 +  { fix y x :: real assume "0 < y" and "y < pi /2" and "y = arctan x" and "tan y = x" hence "- (pi / 2) < y" by auto
  1.1057 +    have "tan y > 0" using tan_monotone'[OF _ _ `- (pi / 2) < y` `y < pi / 2`, of 0] tan_zero `0 < y` by auto
  1.1058 +    hence "x > 0" using `tan y = x` by auto
  1.1059 +
  1.1060 +    have "- (pi / 2) < pi / 2 - y" using `y > 0` `y < pi / 2` by auto
  1.1061 +    moreover have "pi / 2 - y < pi / 2" using `y > 0` `y < pi / 2` by auto
  1.1062 +    ultimately have "arctan (1 / x) = pi / 2 - y" unfolding `tan y = x`[symmetric] tan_inverse using arctan_tan by auto
  1.1063 +    hence "arctan (1 / x) = sgn x * pi / 2 - arctan x" unfolding `y = arctan x` real_sgn_pos[OF `x > 0`] by auto
  1.1064 +  } note pos_y = this
  1.1065 +
  1.1066 +  show ?thesis
  1.1067 +  proof (cases "y > 0")
  1.1068 +    case True from pos_y[OF this `y < pi / 2` `y = arctan x` `tan y = x`] show ?thesis .
  1.1069 +  next
  1.1070 +    case False hence "y \<le> 0" by auto
  1.1071 +    moreover have "y \<noteq> 0" 
  1.1072 +    proof (rule ccontr)
  1.1073 +      assume "\<not> y \<noteq> 0" hence "y = 0" by auto
  1.1074 +      have "x = 0" unfolding `tan y = x`[symmetric] `y = 0` tan_zero ..
  1.1075 +      thus False using `x \<noteq> 0` by auto
  1.1076 +    qed
  1.1077 +    ultimately have "y < 0" by auto
  1.1078 +    hence "0 < - y" and "-y < pi / 2" using `- (pi / 2) < y` by auto
  1.1079 +    moreover have "-y = arctan (-x)" unfolding arctan_minus `y = arctan x` ..
  1.1080 +    moreover have "tan (-y) = -x" unfolding tan_minus `tan y = x` ..
  1.1081 +    ultimately have "arctan (1 / -x) = sgn (-x) * pi / 2 - arctan (-x)" using pos_y by blast
  1.1082 +    hence "arctan (- (1 / x)) = - (sgn x * pi / 2 - arctan x)" unfolding arctan_minus[of x] divide_minus_right sgn_minus by auto
  1.1083 +    thus ?thesis unfolding arctan_minus neg_equal_iff_equal .
  1.1084 +  qed
  1.1085 +qed
  1.1086 +
  1.1087 +theorem pi_series: "pi / 4 = (\<Sum> k. (-1)^k * 1 / real (k*2+1))" (is "_ = ?SUM")
  1.1088 +proof -
  1.1089 +  have "pi / 4 = arctan 1" using arctan1_eq_pi4 by auto
  1.1090 +  also have "\<dots> = ?SUM" using arctan_series[of 1] by auto
  1.1091 +  finally show ?thesis by auto
  1.1092 +qed
  1.1093  
  1.1094  subsection {* Existence of Polar Coordinates *}
  1.1095