src/HOL/Analysis/Harmonic_Numbers.thy
changeset 63627 6ddb43c6b711
parent 63594 bd218a9320b5
child 63721 492bb53c3420
equal deleted inserted replaced
63626:44ce6b524ff3 63627:6ddb43c6b711
       
     1 (*  Title:    HOL/Analysis/Harmonic_Numbers.thy
       
     2     Author:   Manuel Eberl, TU München
       
     3 *)
       
     4 
       
     5 section \<open>Harmonic Numbers\<close>
       
     6 
       
     7 theory Harmonic_Numbers
       
     8 imports
       
     9   Complex_Transcendental
       
    10   Summation_Tests
       
    11   Integral_Test
       
    12 begin
       
    13 
       
    14 text \<open>
       
    15   The definition of the Harmonic Numbers and the Euler-Mascheroni constant.
       
    16   Also provides a reasonably accurate approximation of @{term "ln 2 :: real"}
       
    17   and the Euler-Mascheroni constant.
       
    18 \<close>
       
    19 
       
    20 lemma ln_2_less_1: "ln 2 < (1::real)"
       
    21 proof -
       
    22   have "2 < 5/(2::real)" by simp
       
    23   also have "5/2 \<le> exp (1::real)" using exp_lower_taylor_quadratic[of 1, simplified] by simp
       
    24   finally have "exp (ln 2) < exp (1::real)" by simp
       
    25   thus "ln 2 < (1::real)" by (subst (asm) exp_less_cancel_iff) simp
       
    26 qed
       
    27 
       
    28 lemma setsum_Suc_diff':
       
    29   fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
       
    30   assumes "m \<le> n"
       
    31   shows "(\<Sum>i = m..<n. f (Suc i) - f i) = f n - f m"
       
    32 using assms by (induct n) (auto simp: le_Suc_eq)
       
    33 
       
    34 
       
    35 subsection \<open>The Harmonic numbers\<close>
       
    36 
       
    37 definition harm :: "nat \<Rightarrow> 'a :: real_normed_field" where
       
    38   "harm n = (\<Sum>k=1..n. inverse (of_nat k))"
       
    39 
       
    40 lemma harm_altdef: "harm n = (\<Sum>k<n. inverse (of_nat (Suc k)))"
       
    41   unfolding harm_def by (induction n) simp_all
       
    42 
       
    43 lemma harm_Suc: "harm (Suc n) = harm n + inverse (of_nat (Suc n))"
       
    44   by (simp add: harm_def)
       
    45 
       
    46 lemma harm_nonneg: "harm n \<ge> (0 :: 'a :: {real_normed_field,linordered_field})"
       
    47   unfolding harm_def by (intro setsum_nonneg) simp_all
       
    48 
       
    49 lemma harm_pos: "n > 0 \<Longrightarrow> harm n > (0 :: 'a :: {real_normed_field,linordered_field})"
       
    50   unfolding harm_def by (intro setsum_pos) simp_all
       
    51 
       
    52 lemma of_real_harm: "of_real (harm n) = harm n"
       
    53   unfolding harm_def by simp
       
    54 
       
    55 lemma norm_harm: "norm (harm n) = harm n"
       
    56   by (subst of_real_harm [symmetric]) (simp add: harm_nonneg)
       
    57 
       
    58 lemma harm_expand:
       
    59   "harm 0 = 0"
       
    60   "harm (Suc 0) = 1"
       
    61   "harm (numeral n) = harm (pred_numeral n) + inverse (numeral n)"
       
    62 proof -
       
    63   have "numeral n = Suc (pred_numeral n)" by simp
       
    64   also have "harm \<dots> = harm (pred_numeral n) + inverse (numeral n)"
       
    65     by (subst harm_Suc, subst numeral_eq_Suc[symmetric]) simp
       
    66   finally show "harm (numeral n) = harm (pred_numeral n) + inverse (numeral n)" .
       
    67 qed (simp_all add: harm_def)
       
    68 
       
    69 lemma not_convergent_harm: "\<not>convergent (harm :: nat \<Rightarrow> 'a :: real_normed_field)"
       
    70 proof -
       
    71   have "convergent (\<lambda>n. norm (harm n :: 'a)) \<longleftrightarrow>
       
    72             convergent (harm :: nat \<Rightarrow> real)" by (simp add: norm_harm)
       
    73   also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k=Suc 0..Suc n. inverse (of_nat k) :: real)"
       
    74     unfolding harm_def[abs_def] by (subst convergent_Suc_iff) simp_all
       
    75   also have "... \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k\<le>n. inverse (of_nat (Suc k)) :: real)"
       
    76     by (subst setsum_shift_bounds_cl_Suc_ivl) (simp add: atLeast0AtMost)
       
    77   also have "... \<longleftrightarrow> summable (\<lambda>n. inverse (of_nat n) :: real)"
       
    78     by (subst summable_Suc_iff [symmetric]) (simp add: summable_iff_convergent')
       
    79   also have "\<not>..." by (rule not_summable_harmonic)
       
    80   finally show ?thesis by (blast dest: convergent_norm)
       
    81 qed
       
    82 
       
    83 lemma harm_pos_iff [simp]: "harm n > (0 :: 'a :: {real_normed_field,linordered_field}) \<longleftrightarrow> n > 0"
       
    84   by (rule iffI, cases n, simp add: harm_expand, simp, rule harm_pos)
       
    85 
       
    86 lemma ln_diff_le_inverse:
       
    87   assumes "x \<ge> (1::real)"
       
    88   shows   "ln (x + 1) - ln x < 1 / x"
       
    89 proof -
       
    90   from assms have "\<exists>z>x. z < x + 1 \<and> ln (x + 1) - ln x = (x + 1 - x) * inverse z"
       
    91     by (intro MVT2) (auto intro!: derivative_eq_intros simp: field_simps)
       
    92   then obtain z where z: "z > x" "z < x + 1" "ln (x + 1) - ln x = inverse z" by auto
       
    93   have "ln (x + 1) - ln x = inverse z" by fact
       
    94   also from z(1,2) assms have "\<dots> < 1 / x" by (simp add: field_simps)
       
    95   finally show ?thesis .
       
    96 qed
       
    97 
       
    98 lemma ln_le_harm: "ln (real n + 1) \<le> (harm n :: real)"
       
    99 proof (induction n)
       
   100   fix n assume IH: "ln (real n + 1) \<le> harm n"
       
   101   have "ln (real (Suc n) + 1) = ln (real n + 1) + (ln (real n + 2) - ln (real n + 1))" by simp
       
   102   also have "(ln (real n + 2) - ln (real n + 1)) \<le> 1 / real (Suc n)"
       
   103     using ln_diff_le_inverse[of "real n + 1"] by (simp add: add_ac)
       
   104   also note IH
       
   105   also have "harm n + 1 / real (Suc n) = harm (Suc n)" by (simp add: harm_Suc field_simps)
       
   106   finally show "ln (real (Suc n) + 1) \<le> harm (Suc n)" by - simp
       
   107 qed (simp_all add: harm_def)
       
   108 
       
   109 
       
   110 subsection \<open>The Euler--Mascheroni constant\<close>
       
   111 
       
   112 text \<open>
       
   113   The limit of the difference between the partial harmonic sum and the natural logarithm
       
   114   (approximately 0.577216). This value occurs e.g. in the definition of the Gamma function.
       
   115  \<close>
       
   116 definition euler_mascheroni :: "'a :: real_normed_algebra_1" where
       
   117   "euler_mascheroni = of_real (lim (\<lambda>n. harm n - ln (of_nat n)))"
       
   118 
       
   119 lemma of_real_euler_mascheroni [simp]: "of_real euler_mascheroni = euler_mascheroni"
       
   120   by (simp add: euler_mascheroni_def)
       
   121 
       
   122 interpretation euler_mascheroni: antimono_fun_sum_integral_diff "\<lambda>x. inverse (x + 1)"
       
   123   by unfold_locales (auto intro!: continuous_intros)
       
   124 
       
   125 lemma euler_mascheroni_sum_integral_diff_series:
       
   126   "euler_mascheroni.sum_integral_diff_series n = harm (Suc n) - ln (of_nat (Suc n))"
       
   127 proof -
       
   128   have "harm (Suc n) = (\<Sum>k=0..n. inverse (of_nat k + 1) :: real)" unfolding harm_def
       
   129     unfolding One_nat_def by (subst setsum_shift_bounds_cl_Suc_ivl) (simp add: add_ac)
       
   130   moreover have "((\<lambda>x. inverse (x + 1) :: real) has_integral ln (of_nat n + 1) - ln (0 + 1))
       
   131                    {0..of_nat n}"
       
   132     by (intro fundamental_theorem_of_calculus)
       
   133        (auto intro!: derivative_eq_intros simp: divide_inverse
       
   134            has_field_derivative_iff_has_vector_derivative[symmetric])
       
   135   hence "integral {0..of_nat n} (\<lambda>x. inverse (x + 1) :: real) = ln (of_nat (Suc n))"
       
   136     by (auto dest!: integral_unique)
       
   137   ultimately show ?thesis
       
   138     by (simp add: euler_mascheroni.sum_integral_diff_series_def atLeast0AtMost)
       
   139 qed
       
   140 
       
   141 lemma euler_mascheroni_sequence_decreasing:
       
   142   "m > 0 \<Longrightarrow> m \<le> n \<Longrightarrow> harm n - ln (of_nat n) \<le> harm m - ln (of_nat m :: real)"
       
   143   by (cases m, simp, cases n, simp, hypsubst,
       
   144       subst (1 2) euler_mascheroni_sum_integral_diff_series [symmetric],
       
   145       rule euler_mascheroni.sum_integral_diff_series_antimono, simp)
       
   146 
       
   147 lemma euler_mascheroni_sequence_nonneg:
       
   148   "n > 0 \<Longrightarrow> harm n - ln (of_nat n) \<ge> (0::real)"
       
   149   by (cases n, simp, hypsubst, subst euler_mascheroni_sum_integral_diff_series [symmetric],
       
   150       rule euler_mascheroni.sum_integral_diff_series_nonneg)
       
   151 
       
   152 lemma euler_mascheroni_convergent: "convergent (\<lambda>n. harm n - ln (of_nat n) :: real)"
       
   153 proof -
       
   154   have A: "(\<lambda>n. harm (Suc n) - ln (of_nat (Suc n))) =
       
   155              euler_mascheroni.sum_integral_diff_series"
       
   156     by (subst euler_mascheroni_sum_integral_diff_series [symmetric]) (rule refl)
       
   157   have "convergent (\<lambda>n. harm (Suc n) - ln (of_nat (Suc n) :: real))"
       
   158     by (subst A) (fact euler_mascheroni.sum_integral_diff_series_convergent)
       
   159   thus ?thesis by (subst (asm) convergent_Suc_iff)
       
   160 qed
       
   161 
       
   162 lemma euler_mascheroni_LIMSEQ:
       
   163   "(\<lambda>n. harm n - ln (of_nat n) :: real) \<longlonglongrightarrow> euler_mascheroni"
       
   164   unfolding euler_mascheroni_def
       
   165   by (simp add: convergent_LIMSEQ_iff [symmetric] euler_mascheroni_convergent)
       
   166 
       
   167 lemma euler_mascheroni_LIMSEQ_of_real:
       
   168   "(\<lambda>n. of_real (harm n - ln (of_nat n))) \<longlonglongrightarrow>
       
   169       (euler_mascheroni :: 'a :: {real_normed_algebra_1, topological_space})"
       
   170 proof -
       
   171   have "(\<lambda>n. of_real (harm n - ln (of_nat n))) \<longlonglongrightarrow> (of_real (euler_mascheroni) :: 'a)"
       
   172     by (intro tendsto_of_real euler_mascheroni_LIMSEQ)
       
   173   thus ?thesis by simp
       
   174 qed
       
   175 
       
   176 lemma euler_mascheroni_sum:
       
   177   "(\<lambda>n. inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2)) :: real)
       
   178        sums euler_mascheroni"
       
   179  using sums_add[OF telescope_sums[OF LIMSEQ_Suc[OF euler_mascheroni_LIMSEQ]]
       
   180                    telescope_sums'[OF LIMSEQ_inverse_real_of_nat]]
       
   181   by (simp_all add: harm_def algebra_simps)
       
   182 
       
   183 lemma alternating_harmonic_series_sums: "(\<lambda>k. (-1)^k / real_of_nat (Suc k)) sums ln 2"
       
   184 proof -
       
   185   let ?f = "\<lambda>n. harm n - ln (real_of_nat n)"
       
   186   let ?g = "\<lambda>n. if even n then 0 else (2::real)"
       
   187   let ?em = "\<lambda>n. harm n - ln (real_of_nat n)"
       
   188   have "eventually (\<lambda>n. ?em (2*n) - ?em n + ln 2 = (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))) at_top"
       
   189     using eventually_gt_at_top[of "0::nat"]
       
   190   proof eventually_elim
       
   191     fix n :: nat assume n: "n > 0"
       
   192     have "(\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k)) =
       
   193               (\<Sum>k<2*n. ((-1)^k + ?g k) / of_nat (Suc k)) - (\<Sum>k<2*n. ?g k / of_nat (Suc k))"
       
   194       by (simp add: setsum.distrib algebra_simps divide_inverse)
       
   195     also have "(\<Sum>k<2*n. ((-1)^k + ?g k) / real_of_nat (Suc k)) = harm (2*n)"
       
   196       unfolding harm_altdef by (intro setsum.cong) (auto simp: field_simps)
       
   197     also have "(\<Sum>k<2*n. ?g k / real_of_nat (Suc k)) = (\<Sum>k|k<2*n \<and> odd k. ?g k / of_nat (Suc k))"
       
   198       by (intro setsum.mono_neutral_right) auto
       
   199     also have "\<dots> = (\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k)))"
       
   200       by (intro setsum.cong) auto
       
   201     also have "(\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k))) = harm n"
       
   202       unfolding harm_altdef
       
   203       by (intro setsum.reindex_cong[of "\<lambda>n. 2*n+1"]) (auto simp: inj_on_def field_simps elim!: oddE)
       
   204     also have "harm (2*n) - harm n = ?em (2*n) - ?em n + ln 2" using n
       
   205       by (simp_all add: algebra_simps ln_mult)
       
   206     finally show "?em (2*n) - ?em n + ln 2 = (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))" ..
       
   207   qed
       
   208   moreover have "(\<lambda>n. ?em (2*n) - ?em n + ln (2::real))
       
   209                      \<longlonglongrightarrow> euler_mascheroni - euler_mascheroni + ln 2"
       
   210     by (intro tendsto_intros euler_mascheroni_LIMSEQ filterlim_compose[OF euler_mascheroni_LIMSEQ]
       
   211               filterlim_subseq) (auto simp: subseq_def)
       
   212   hence "(\<lambda>n. ?em (2*n) - ?em n + ln (2::real)) \<longlonglongrightarrow> ln 2" by simp
       
   213   ultimately have "(\<lambda>n. (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))) \<longlonglongrightarrow> ln 2"
       
   214     by (rule Lim_transform_eventually)
       
   215 
       
   216   moreover have "summable (\<lambda>k. (-1)^k * inverse (real_of_nat (Suc k)))"
       
   217     using LIMSEQ_inverse_real_of_nat
       
   218     by (intro summable_Leibniz(1) decseq_imp_monoseq decseq_SucI) simp_all
       
   219   hence A: "(\<lambda>n. \<Sum>k<n. (-1)^k / real_of_nat (Suc k)) \<longlonglongrightarrow> (\<Sum>k. (-1)^k / real_of_nat (Suc k))"
       
   220     by (simp add: summable_sums_iff divide_inverse sums_def)
       
   221   from filterlim_compose[OF this filterlim_subseq[of "op * (2::nat)"]]
       
   222     have "(\<lambda>n. \<Sum>k<2*n. (-1)^k / real_of_nat (Suc k)) \<longlonglongrightarrow> (\<Sum>k. (-1)^k / real_of_nat (Suc k))"
       
   223     by (simp add: subseq_def)
       
   224   ultimately have "(\<Sum>k. (- 1) ^ k / real_of_nat (Suc k)) = ln 2" by (intro LIMSEQ_unique)
       
   225   with A show ?thesis by (simp add: sums_def)
       
   226 qed
       
   227 
       
   228 lemma alternating_harmonic_series_sums':
       
   229   "(\<lambda>k. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2))) sums ln 2"
       
   230 unfolding sums_def
       
   231 proof (rule Lim_transform_eventually)
       
   232   show "(\<lambda>n. \<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) \<longlonglongrightarrow> ln 2"
       
   233     using alternating_harmonic_series_sums unfolding sums_def
       
   234     by (rule filterlim_compose) (rule mult_nat_left_at_top, simp)
       
   235   show "eventually (\<lambda>n. (\<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) =
       
   236             (\<Sum>k<n. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2)))) sequentially"
       
   237   proof (intro always_eventually allI)
       
   238     fix n :: nat
       
   239     show "(\<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) =
       
   240               (\<Sum>k<n. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2)))"
       
   241       by (induction n) (simp_all add: inverse_eq_divide)
       
   242   qed
       
   243 qed
       
   244 
       
   245 
       
   246 subsection \<open>Bounds on the Euler--Mascheroni constant\<close>
       
   247 
       
   248 (* TODO: Move? *)
       
   249 lemma ln_inverse_approx_le:
       
   250   assumes "(x::real) > 0" "a > 0"
       
   251   shows   "ln (x + a) - ln x \<le> a * (inverse x + inverse (x + a))/2" (is "_ \<le> ?A")
       
   252 proof -
       
   253   define f' where "f' = (inverse (x + a) - inverse x)/a"
       
   254   have f'_nonpos: "f' \<le> 0" using assms by (simp add: f'_def divide_simps)
       
   255   let ?f = "\<lambda>t. (t - x) * f' + inverse x"
       
   256   let ?F = "\<lambda>t. (t - x)^2 * f' / 2 + t * inverse x"
       
   257   have diff: "\<forall>t\<in>{x..x+a}. (?F has_vector_derivative ?f t)
       
   258                                (at t within {x..x+a})" using assms
       
   259     by (auto intro!: derivative_eq_intros
       
   260              simp: has_field_derivative_iff_has_vector_derivative[symmetric])
       
   261   from assms have "(?f has_integral (?F (x+a) - ?F x)) {x..x+a}"
       
   262     by (intro fundamental_theorem_of_calculus[OF _ diff])
       
   263        (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] field_simps
       
   264              intro!: derivative_eq_intros)
       
   265   also have "?F (x+a) - ?F x = (a*2 + f'*a\<^sup>2*x) / (2*x)" using assms by (simp add: field_simps)
       
   266   also have "f'*a^2 = - (a^2) / (x*(x + a))" using assms
       
   267     by (simp add: divide_simps f'_def power2_eq_square)
       
   268   also have "(a*2 + - a\<^sup>2/(x*(x+a))*x) / (2*x) = ?A" using assms
       
   269     by (simp add: divide_simps power2_eq_square) (simp add: algebra_simps)
       
   270   finally have int1: "((\<lambda>t. (t - x) * f' + inverse x) has_integral ?A) {x..x + a}" .
       
   271 
       
   272   from assms have int2: "(inverse has_integral (ln (x + a) - ln x)) {x..x+a}"
       
   273     by (intro fundamental_theorem_of_calculus)
       
   274        (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps
       
   275              intro!: derivative_eq_intros)
       
   276   hence "ln (x + a) - ln x = integral {x..x+a} inverse" by (simp add: integral_unique)
       
   277   also have ineq: "\<forall>xa\<in>{x..x + a}. inverse xa \<le> (xa - x) * f' + inverse x"
       
   278   proof
       
   279     fix t assume t': "t \<in> {x..x+a}"
       
   280     with assms have t: "0 \<le> (t - x) / a" "(t - x) / a \<le> 1" by simp_all
       
   281     have "inverse t = inverse ((1 - (t - x) / a) *\<^sub>R x + ((t - x) / a) *\<^sub>R (x + a))" (is "_ = ?A")
       
   282       using assms t' by (simp add: field_simps)
       
   283     also from assms have "convex_on {x..x+a} inverse" by (intro convex_on_inverse) auto
       
   284     from convex_onD_Icc[OF this _ t] assms
       
   285       have "?A \<le> (1 - (t - x) / a) * inverse x + (t - x) / a * inverse (x + a)" by simp
       
   286     also have "\<dots> = (t - x) * f' + inverse x" using assms
       
   287       by (simp add: f'_def divide_simps) (simp add: f'_def field_simps)
       
   288     finally show "inverse t \<le> (t - x) * f' + inverse x" .
       
   289   qed
       
   290   hence "integral {x..x+a} inverse \<le> integral {x..x+a} ?f" using f'_nonpos assms
       
   291     by (intro integral_le has_integral_integrable[OF int1] has_integral_integrable[OF int2] ineq)
       
   292   also have "\<dots> = ?A" using int1 by (rule integral_unique)
       
   293   finally show ?thesis .
       
   294 qed
       
   295 
       
   296 lemma ln_inverse_approx_ge:
       
   297   assumes "(x::real) > 0" "x < y"
       
   298   shows   "ln y - ln x \<ge> 2 * (y - x) / (x + y)" (is "_ \<ge> ?A")
       
   299 proof -
       
   300   define m where "m = (x+y)/2"
       
   301   define f' where "f' = -inverse (m^2)"
       
   302   from assms have m: "m > 0" by (simp add: m_def)
       
   303   let ?F = "\<lambda>t. (t - m)^2 * f' / 2 + t / m"
       
   304   from assms have "((\<lambda>t. (t - m) * f' + inverse m) has_integral (?F y - ?F x)) {x..y}"
       
   305     by (intro fundamental_theorem_of_calculus)
       
   306        (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps
       
   307              intro!: derivative_eq_intros)
       
   308   also from m have "?F y - ?F x = ((y - m)^2 - (x - m)^2) * f' / 2 + (y - x) / m"
       
   309     by (simp add: field_simps)
       
   310   also have "((y - m)^2 - (x - m)^2) = 0" by (simp add: m_def power2_eq_square field_simps)
       
   311   also have "0 * f' / 2 + (y - x) / m = ?A" by (simp add: m_def)
       
   312   finally have int1: "((\<lambda>t. (t - m) * f' + inverse m) has_integral ?A) {x..y}" .
       
   313 
       
   314   from assms have int2: "(inverse has_integral (ln y - ln x)) {x..y}"
       
   315     by (intro fundamental_theorem_of_calculus)
       
   316        (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps
       
   317              intro!: derivative_eq_intros)
       
   318   hence "ln y - ln x = integral {x..y} inverse" by (simp add: integral_unique)
       
   319   also have ineq: "\<forall>xa\<in>{x..y}. inverse xa \<ge> (xa - m) * f' + inverse m"
       
   320   proof
       
   321     fix t assume t: "t \<in> {x..y}"
       
   322     from t assms have "inverse t - inverse m \<ge> f' * (t - m)"
       
   323       by (intro convex_on_imp_above_tangent[of "{0<..}"] convex_on_inverse)
       
   324          (auto simp: m_def interior_open f'_def power2_eq_square intro!: derivative_eq_intros)
       
   325     thus "(t - m) * f' + inverse m \<le> inverse t" by (simp add: algebra_simps)
       
   326   qed
       
   327   hence "integral {x..y} inverse \<ge> integral {x..y} (\<lambda>t. (t - m) * f' + inverse m)"
       
   328     using int1 int2 by (intro integral_le has_integral_integrable)
       
   329   also have "integral {x..y} (\<lambda>t. (t - m) * f' + inverse m) = ?A"
       
   330     using integral_unique[OF int1] by simp
       
   331   finally show ?thesis .
       
   332 qed
       
   333 
       
   334 
       
   335 lemma euler_mascheroni_lower:
       
   336         "euler_mascheroni \<ge> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))"
       
   337   and euler_mascheroni_upper:
       
   338         "euler_mascheroni \<le> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))"
       
   339 proof -
       
   340   define D :: "_ \<Rightarrow> real"
       
   341     where "D n = inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2))" for n
       
   342   let ?g = "\<lambda>n. ln (of_nat (n+2)) - ln (of_nat (n+1)) - inverse (of_nat (n+1)) :: real"
       
   343   define inv where [abs_def]: "inv n = inverse (real_of_nat n)" for n
       
   344   fix n :: nat
       
   345   note summable = sums_summable[OF euler_mascheroni_sum, folded D_def]
       
   346   have sums: "(\<lambda>k. (inv (Suc (k + (n+1))) - inv (Suc (Suc k + (n+1))))/2) sums ((inv (Suc (0 + (n+1))) - 0)/2)"
       
   347     unfolding inv_def
       
   348     by (intro sums_divide telescope_sums' LIMSEQ_ignore_initial_segment LIMSEQ_inverse_real_of_nat)
       
   349   have sums': "(\<lambda>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2) sums ((inv (Suc (0 + n)) - 0)/2)"
       
   350     unfolding inv_def
       
   351     by (intro sums_divide telescope_sums' LIMSEQ_ignore_initial_segment LIMSEQ_inverse_real_of_nat)
       
   352   from euler_mascheroni_sum have "euler_mascheroni = (\<Sum>k. D k)"
       
   353     by (simp add: sums_iff D_def)
       
   354   also have "\<dots> = (\<Sum>k. D (k + Suc n)) + (\<Sum>k\<le>n. D k)"
       
   355     by (subst suminf_split_initial_segment[OF summable, of "Suc n"], subst lessThan_Suc_atMost) simp
       
   356   finally have sum: "(\<Sum>k\<le>n. D k) - euler_mascheroni = -(\<Sum>k. D (k + Suc n))" by simp
       
   357 
       
   358   note sum
       
   359   also have "\<dots> \<le> -(\<Sum>k. (inv (k + Suc n + 1) - inv (k + Suc n + 2)) / 2)"
       
   360   proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable])
       
   361     fix k' :: nat
       
   362     define k where "k = k' + Suc n"
       
   363     hence k: "k > 0" by (simp add: k_def)
       
   364     have "real_of_nat (k+1) > 0" by (simp add: k_def)
       
   365     with ln_inverse_approx_le[OF this zero_less_one]
       
   366       have "ln (of_nat k + 2) - ln (of_nat k + 1) \<le> (inv (k+1) + inv (k+2))/2"
       
   367       by (simp add: inv_def add_ac)
       
   368     hence "(inv (k+1) - inv (k+2))/2 \<le> inv (k+1) + ln (of_nat (k+1)) - ln (of_nat (k+2))"
       
   369       by (simp add: field_simps)
       
   370     also have "\<dots> = D k" unfolding D_def inv_def ..
       
   371     finally show "D (k' + Suc n) \<ge> (inv (k' + Suc n + 1) - inv (k' + Suc n + 2)) / 2"
       
   372       by (simp add: k_def)
       
   373     from sums_summable[OF sums]
       
   374       show "summable (\<lambda>k. (inv (k + Suc n + 1) - inv (k + Suc n + 2))/2)" by simp
       
   375   qed
       
   376   also from sums have "\<dots> = -inv (n+2) / 2" by (simp add: sums_iff)
       
   377   finally have "euler_mascheroni \<ge> (\<Sum>k\<le>n. D k) + 1 / (of_nat (2 * (n+2)))"
       
   378     by (simp add: inv_def field_simps)
       
   379   also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))"
       
   380     unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add:  setsum.distrib setsum_subtractf)
       
   381   also have "(\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1))) = ln (of_nat (n+2))"
       
   382     by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all
       
   383   finally show "euler_mascheroni \<ge> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))"
       
   384     by simp
       
   385 
       
   386   note sum
       
   387   also have "-(\<Sum>k. D (k + Suc n)) \<ge> -(\<Sum>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)"
       
   388   proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable])
       
   389     fix k' :: nat
       
   390     define k where "k = k' + Suc n"
       
   391     hence k: "k > 0" by (simp add: k_def)
       
   392     have "real_of_nat (k+1) > 0" by (simp add: k_def)
       
   393     from ln_inverse_approx_ge[of "of_nat k + 1" "of_nat k + 2"]
       
   394       have "2 / (2 * real_of_nat k + 3) \<le> ln (of_nat (k+2)) - ln (real_of_nat (k+1))"
       
   395       by (simp add: add_ac)
       
   396     hence "D k \<le> 1 / real_of_nat (k+1) - 2 / (2 * real_of_nat k + 3)"
       
   397       by (simp add: D_def inverse_eq_divide inv_def)
       
   398     also have "\<dots> = inv ((k+1)*(2*k+3))" unfolding inv_def by (simp add: field_simps)
       
   399     also have "\<dots> \<le> inv (2*k*(k+1))" unfolding inv_def using k
       
   400       by (intro le_imp_inverse_le)
       
   401          (simp add: algebra_simps, simp del: of_nat_add)
       
   402     also have "\<dots> = (inv k - inv (k+1))/2" unfolding inv_def using k
       
   403       by (simp add: divide_simps del: of_nat_mult) (simp add: algebra_simps)
       
   404     finally show "D k \<le> (inv (Suc (k' + n)) - inv (Suc (Suc k' + n)))/2" unfolding k_def by simp
       
   405   next
       
   406     from sums_summable[OF sums']
       
   407       show "summable (\<lambda>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)" by simp
       
   408   qed
       
   409   also from sums' have "(\<Sum>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2) = inv (n+1)/2"
       
   410     by (simp add: sums_iff)
       
   411   finally have "euler_mascheroni \<le> (\<Sum>k\<le>n. D k) + 1 / of_nat (2 * (n+1))"
       
   412     by (simp add: inv_def field_simps)
       
   413   also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))"
       
   414     unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add:  setsum.distrib setsum_subtractf)
       
   415   also have "(\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1))) = ln (of_nat (n+2))"
       
   416     by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all
       
   417   finally show "euler_mascheroni \<le> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))"
       
   418     by simp
       
   419 qed
       
   420 
       
   421 lemma euler_mascheroni_pos: "euler_mascheroni > (0::real)"
       
   422   using euler_mascheroni_lower[of 0] ln_2_less_1 by (simp add: harm_def)
       
   423 
       
   424 context
       
   425 begin
       
   426 
       
   427 private lemma ln_approx_aux:
       
   428   fixes n :: nat and x :: real
       
   429   defines "y \<equiv> (x-1)/(x+1)"
       
   430   assumes x: "x > 0" "x \<noteq> 1"
       
   431   shows "inverse (2*y^(2*n+1)) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in>
       
   432             {0..(1 / (1 - y^2) / of_nat (2*n+1))}"
       
   433 proof -
       
   434   from x have norm_y: "norm y < 1" unfolding y_def by simp
       
   435   from power_strict_mono[OF this, of 2] have norm_y': "norm y^2 < 1" by simp
       
   436 
       
   437   let ?f = "\<lambda>k. 2 * y ^ (2*k+1) / of_nat (2*k+1)"
       
   438   note sums = ln_series_quadratic[OF x(1)]
       
   439   define c where "c = inverse (2*y^(2*n+1))"
       
   440   let ?d = "c * (ln x - (\<Sum>k<n. ?f k))"
       
   441   have "\<forall>k. y\<^sup>2^k / of_nat (2*(k+n)+1) \<le> y\<^sup>2 ^ k / of_nat (2*n+1)"
       
   442     by (intro allI divide_left_mono mult_right_mono mult_pos_pos zero_le_power[of "y^2"]) simp_all
       
   443   moreover {
       
   444     have "(\<lambda>k. ?f (k + n)) sums (ln x - (\<Sum>k<n. ?f k))"
       
   445       using sums_split_initial_segment[OF sums] by (simp add: y_def)
       
   446     hence "(\<lambda>k. c * ?f (k + n)) sums ?d" by (rule sums_mult)
       
   447     also have "(\<lambda>k. c * (2*y^(2*(k+n)+1) / of_nat (2*(k+n)+1))) =
       
   448                    (\<lambda>k. (c * (2*y^(2*n+1))) * ((y^2)^k / of_nat (2*(k+n)+1)))"
       
   449       by (simp only: ring_distribs power_add power_mult) (simp add: mult_ac)
       
   450     also from x have "c * (2*y^(2*n+1)) = 1" by (simp add: c_def y_def)
       
   451     finally have "(\<lambda>k. (y^2)^k / of_nat (2*(k+n)+1)) sums ?d" by simp
       
   452   } note sums' = this
       
   453   moreover from norm_y' have "(\<lambda>k. (y^2)^k / of_nat (2*n+1)) sums (1 / (1 - y^2) / of_nat (2*n+1))"
       
   454     by (intro sums_divide geometric_sums) (simp_all add: norm_power)
       
   455   ultimately have "?d \<le> (1 / (1 - y^2) / of_nat (2*n+1))" by (rule sums_le)
       
   456   moreover have "c * (ln x - (\<Sum>k<n. 2 * y ^ (2 * k + 1) / real_of_nat (2 * k + 1))) \<ge> 0"
       
   457     by (intro sums_le[OF _ sums_zero sums']) simp_all
       
   458   ultimately show ?thesis unfolding c_def by simp
       
   459 qed
       
   460 
       
   461 lemma
       
   462   fixes n :: nat and x :: real
       
   463   defines "y \<equiv> (x-1)/(x+1)"
       
   464   defines "approx \<equiv> (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))"
       
   465   defines "d \<equiv> y^(2*n+1) / (1 - y^2) / of_nat (2*n+1)"
       
   466   assumes x: "x > 1"
       
   467   shows   ln_approx_bounds: "ln x \<in> {approx..approx + 2*d}"
       
   468   and     ln_approx_abs:    "abs (ln x - (approx + d)) \<le> d"
       
   469 proof -
       
   470   define c where "c = 2*y^(2*n+1)"
       
   471   from x have c_pos: "c > 0" unfolding c_def y_def
       
   472     by (intro mult_pos_pos zero_less_power) simp_all
       
   473   have A: "inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in>
       
   474               {0.. (1 / (1 - y^2) / of_nat (2*n+1))}" using assms unfolding y_def c_def
       
   475     by (intro ln_approx_aux) simp_all
       
   476   hence "inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1)/of_nat (2*k+1))) \<le> (1 / (1-y^2) / of_nat (2*n+1))"
       
   477     by simp
       
   478   hence "(ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) / c \<le> (1 / (1 - y^2) / of_nat (2*n+1))"
       
   479     by (auto simp add: divide_simps)
       
   480   with c_pos have "ln x \<le> c / (1 - y^2) / of_nat (2*n+1) + approx"
       
   481     by (subst (asm) pos_divide_le_eq) (simp_all add: mult_ac approx_def)
       
   482   moreover {
       
   483     from A c_pos have "0 \<le> c * (inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))))"
       
   484       by (intro mult_nonneg_nonneg[of c]) simp_all
       
   485     also have "\<dots> = (c * inverse c) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1)))"
       
   486       by (simp add: mult_ac)
       
   487     also from c_pos have "c * inverse c = 1" by simp
       
   488     finally have "ln x \<ge> approx" by (simp add: approx_def)
       
   489   }
       
   490   ultimately show "ln x \<in> {approx..approx + 2*d}" by (simp add: c_def d_def)
       
   491   thus "abs (ln x - (approx + d)) \<le> d" by auto
       
   492 qed
       
   493 
       
   494 end
       
   495 
       
   496 lemma euler_mascheroni_bounds:
       
   497   fixes n :: nat assumes "n \<ge> 1" defines "t \<equiv> harm n - ln (of_nat (Suc n)) :: real"
       
   498   shows "euler_mascheroni \<in> {t + inverse (of_nat (2*(n+1)))..t + inverse (of_nat (2*n))}"
       
   499   using assms euler_mascheroni_upper[of "n-1"] euler_mascheroni_lower[of "n-1"]
       
   500   unfolding t_def by (cases n) (simp_all add: harm_Suc t_def inverse_eq_divide)
       
   501 
       
   502 lemma euler_mascheroni_bounds':
       
   503   fixes n :: nat assumes "n \<ge> 1" "ln (real_of_nat (Suc n)) \<in> {l<..<u}"
       
   504   shows "euler_mascheroni \<in>
       
   505            {harm n - u + inverse (of_nat (2*(n+1)))<..<harm n - l + inverse (of_nat (2*n))}"
       
   506   using euler_mascheroni_bounds[OF assms(1)] assms(2) by auto
       
   507 
       
   508 
       
   509 text \<open>
       
   510   Approximation of @{term "ln 2"}. The lower bound is accurate to about 0.03; the upper
       
   511   bound is accurate to about 0.0015.
       
   512 \<close>
       
   513 lemma ln2_ge_two_thirds: "2/3 \<le> ln (2::real)"
       
   514   and ln2_le_25_over_36: "ln (2::real) \<le> 25/36"
       
   515   using ln_approx_bounds[of 2 1, simplified, simplified eval_nat_numeral, simplified] by simp_all
       
   516 
       
   517 
       
   518 text \<open>
       
   519   Approximation of the Euler--Mascheroni constant. The lower bound is accurate to about 0.0015;
       
   520   the upper bound is accurate to about 0.015.
       
   521 \<close>
       
   522 lemma euler_mascheroni_gt_19_over_33: "(euler_mascheroni :: real) > 19/33" (is ?th1)
       
   523   and euler_mascheroni_less_13_over_22: "(euler_mascheroni :: real) < 13/22" (is ?th2)
       
   524 proof -
       
   525   have "ln (real (Suc 7)) = 3 * ln 2" by (simp add: ln_powr [symmetric] powr_numeral)
       
   526   also from ln_approx_bounds[of 2 3] have "\<dots> \<in> {3*307/443<..<3*4615/6658}"
       
   527     by (simp add: eval_nat_numeral)
       
   528   finally have "ln (real (Suc 7)) \<in> \<dots>" .
       
   529   from euler_mascheroni_bounds'[OF _ this] have "?th1 \<and> ?th2" by (simp_all add: harm_expand)
       
   530   thus ?th1 ?th2 by blast+
       
   531 qed
       
   532 
       
   533 end